Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/Bicategory/BicategoryOfSpans.thy b/thys/Bicategory/BicategoryOfSpans.thy
--- a/thys/Bicategory/BicategoryOfSpans.thy
+++ b/thys/Bicategory/BicategoryOfSpans.thy
@@ -1,14745 +1,14745 @@
(* Title: BicategoryOfSpans
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Bicategories of Spans"
theory BicategoryOfSpans
imports Category3.ConcreteCategory IsomorphismClass CanonicalIsos EquivalenceOfBicategories
SpanBicategory Tabulation
begin
text \<open>
In this section, we prove CKS Theorem 4, which characterizes up to equivalence the
bicategories of spans in a category with pullbacks.
The characterization consists of three conditions:
BS1: ``Every 1-cell is isomorphic to a composition \<open>g \<star> f\<^sup>*\<close>, where f and g are maps'';
BS2: ``For every span of maps \<open>(f, g)\<close> there is a 2-cell \<open>\<rho>\<close> such that \<open>(f, \<rho>, g)\<close>
is a tabulation''; and
BS3: ``Any two 2-cells between the same pair of maps are equal and invertible.''
One direction of the proof, which is the easier direction once it is established that
BS1 and BS3 are respected by equivalence of bicategories, shows that if a bicategory \<open>B\<close>
is biequivalent to the bicategory of spans in some category \<open>C\<close> with pullbacks,
then it satisfies BS1 -- BS3.
The other direction, which is harder, shows that a bicategory \<open>B\<close> satisfying BS1 -- BS3 is
biequivalent to the bicategory of spans in a certain category with pullbacks that
is constructed from the sub-bicategory of maps of \<open>B\<close>.
\<close>
subsection "Definition"
text \<open>
We define a \emph{bicategory of spans} to be a bicategory that satisfies the conditions
\<open>BS1\<close> -- \<open>BS3\<close> stated informally above.
\<close>
locale bicategory_of_spans =
bicategory + chosen_right_adjoints +
assumes BS1: "\<And>r. ide r \<Longrightarrow> \<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> isomorphic r (g \<star> f\<^sup>*)"
and BS2: "\<And>f g. \<lbrakk> is_left_adjoint f; is_left_adjoint g; src f = src g \<rbrakk>
\<Longrightarrow> \<exists>r \<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
and BS3: "\<And>f f' \<mu> \<mu>'. \<lbrakk> is_left_adjoint f; is_left_adjoint f'; \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright> \<rbrakk>
\<Longrightarrow> iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
text \<open>
Using the already-established fact \<open>equivalence_pseudofunctor.reflects_tabulation\<close>
that tabulations are reflected by equivalence pseudofunctors, it is not difficult to prove
that the notion `bicategory of spans' respects equivalence of bicategories.
\<close>
lemma bicategory_of_spans_respects_equivalence:
assumes "equivalent_bicategories V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
and "bicategory_of_spans V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C"
shows "bicategory_of_spans V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
proof -
interpret C: bicategory_of_spans V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
using assms by simp
interpret C: chosen_right_adjoints V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
interpret D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
using assms equivalent_bicategories_def equivalence_pseudofunctor.axioms(1)
pseudofunctor.axioms(2)
by fast
interpret D: chosen_right_adjoints V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
obtain F \<Phi> where F: "equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>"
using assms equivalent_bicategories_def by blast
interpret F: equivalence_pseudofunctor
V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi>
using F by simp
interpret E: equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C (* 17 sec *)
F \<Phi> F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1
using F.extends_to_equivalence_of_bicategories by simp
interpret E': converse_equivalence_of_bicategories
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F \<Phi> F.right_map F.right_cmp F.unit\<^sub>0 F.unit\<^sub>1 F.counit\<^sub>0 F.counit\<^sub>1
..
interpret G: equivalence_pseudofunctor
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
F.right_map F.right_cmp
using E'.equivalence_pseudofunctor_left by simp
write V\<^sub>C (infixr "\<cdot>\<^sub>C" 55)
write V\<^sub>D (infixr "\<cdot>\<^sub>D" 55)
write H\<^sub>C (infixr "\<star>\<^sub>C" 53)
write H\<^sub>D (infixr "\<star>\<^sub>D" 53)
write \<a>\<^sub>C ("\<a>\<^sub>C[_, _, _]")
write \<a>\<^sub>D ("\<a>\<^sub>D[_, _, _]")
write C.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
write C.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>C _\<guillemotright>")
write D.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
write D.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>D _\<guillemotright>")
write C.isomorphic (infix "\<cong>\<^sub>C" 50)
write D.isomorphic (infix "\<cong>\<^sub>D" 50)
write C.some_right_adjoint ("_\<^sup>*\<^sup>C" [1000] 1000)
write D.some_right_adjoint ("_\<^sup>*\<^sup>D" [1000] 1000)
show "bicategory_of_spans V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D"
proof
show "\<And>r'. D.ide r' \<Longrightarrow>
\<exists>f' g'. D.is_left_adjoint f' \<and> D.is_left_adjoint g' \<and> r' \<cong>\<^sub>D g' \<star>\<^sub>D (f')\<^sup>*\<^sup>D"
proof -
fix r'
assume r': "D.ide r'"
obtain f g
where fg: "C.is_left_adjoint f \<and> C.is_left_adjoint g \<and> F.right_map r' \<cong>\<^sub>C g \<star>\<^sub>C f\<^sup>*\<^sup>C"
using r' C.BS1 [of "F.right_map r'"] by auto
have trg_g: "trg\<^sub>C g = E.G.map\<^sub>0 (trg\<^sub>D r')"
proof -
have "trg\<^sub>C g = trg\<^sub>C (g \<star>\<^sub>C f\<^sup>*\<^sup>C)"
using fg C.isomorphic_implies_ide by auto
also have "... = trg\<^sub>C (F.right_map r')"
using fg C.isomorphic_implies_hpar by auto
also have "... = E.G.map\<^sub>0 (trg\<^sub>D r')"
using r' by simp
finally show ?thesis by simp
qed
have trg_f: "trg\<^sub>C f = E.G.map\<^sub>0 (src\<^sub>D r')"
proof -
have "trg\<^sub>C f = src\<^sub>C f\<^sup>*\<^sup>C"
using fg by simp
also have "... = src\<^sub>C (g \<star>\<^sub>C f\<^sup>*\<^sup>C)"
using fg C.isomorphic_implies_ide by force
also have "... = src\<^sub>C (F.right_map r')"
using fg C.isomorphic_implies_hpar by auto
also have "... = E.G.map\<^sub>0 (src\<^sub>D r')"
using r' by simp
finally show ?thesis by simp
qed
define d_src where "d_src \<equiv> F.counit\<^sub>0 (src\<^sub>D r')"
define e_src where "e_src \<equiv> (F.counit\<^sub>0 (src\<^sub>D r'))\<^sup>~\<^sup>D"
have d_src: "\<guillemotleft>d_src : F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r')) \<rightarrow>\<^sub>D src\<^sub>D r'\<guillemotright> \<and>
D.equivalence_map d_src"
using d_src_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
have e_src: "\<guillemotleft>e_src : src\<^sub>D r' \<rightarrow>\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (src\<^sub>D r'))\<guillemotright> \<and>
D.equivalence_map e_src"
using e_src_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
obtain \<eta>_src \<epsilon>_src
where eq_src: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \<eta>_src \<epsilon>_src"
using d_src_def e_src_def d_src e_src D.quasi_inverses_some_quasi_inverse
D.quasi_inverses_def
by blast
interpret eq_src: equivalence_in_bicategory
V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_src e_src \<eta>_src \<epsilon>_src
using eq_src by simp
define d_trg where "d_trg \<equiv> F.counit\<^sub>0 (trg\<^sub>D r')"
define e_trg where "e_trg \<equiv> (F.counit\<^sub>0 (trg\<^sub>D r'))\<^sup>~\<^sup>D"
have d_trg: "\<guillemotleft>d_trg : F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r')) \<rightarrow>\<^sub>D trg\<^sub>D r'\<guillemotright> \<and>
D.equivalence_map d_trg"
using d_trg_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
have e_trg: "\<guillemotleft>e_trg : trg\<^sub>D r' \<rightarrow>\<^sub>D F.map\<^sub>0 (E.G.map\<^sub>0 (trg\<^sub>D r'))\<guillemotright> \<and>
D.equivalence_map e_trg"
using e_trg_def r' E.\<epsilon>.map\<^sub>0_in_hhom E.\<epsilon>.components_are_equivalences by simp
obtain \<eta>_trg \<epsilon>_trg
where eq_trg: "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \<eta>_trg \<epsilon>_trg"
using d_trg_def e_trg_def d_trg e_trg D.quasi_inverses_some_quasi_inverse
D.quasi_inverses_def
by blast
interpret eq_trg: equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D d_trg e_trg \<eta>_trg \<epsilon>_trg
using eq_trg by simp
interpret eqs: two_equivalences_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
d_src e_src \<eta>_src \<epsilon>_src d_trg e_trg \<eta>_trg \<epsilon>_trg
..
interpret hom: subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : trg\<^sub>D d_src \<rightarrow>\<^sub>D trg\<^sub>D d_trg\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret hom': subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : src\<^sub>D d_src \<rightarrow>\<^sub>D src\<^sub>D d_trg\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
interpret e: equivalence_of_categories hom.comp hom'.comp eqs.F eqs.G eqs.\<phi> eqs.\<psi>
using eqs.induces_equivalence_of_hom_categories by simp
have r'_in_hhom: "D.in_hhom r' (src\<^sub>D e_src) (src\<^sub>D e_trg)"
using r' e_src e_trg by (simp add: D.in_hhom_def)
define g'
where "g' = d_trg \<star>\<^sub>D F g"
have g': "D.is_left_adjoint g'"
proof -
have "D.is_left_adjoint d_trg"
using r' d_trg D.equivalence_is_adjoint by simp
moreover have "src\<^sub>D d_trg = trg\<^sub>D (F g)"
using fg d_trg trg_g C.left_adjoint_is_ide by auto
ultimately show ?thesis
unfolding g'_def
using fg r' D.left_adjoints_compose F.preserves_left_adjoint by blast
qed
have 1: "D.is_right_adjoint (F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src)"
proof -
have "D.is_right_adjoint e_src"
using r' e_src D.equivalence_is_adjoint by simp
moreover have "D.is_right_adjoint (F f\<^sup>*\<^sup>C)"
using fg C.left_adjoint_extends_to_adjoint_pair F.preserves_adjoint_pair by blast
moreover have "src\<^sub>D (F f\<^sup>*\<^sup>C) = trg\<^sub>D e_src"
using fg r' trg_f C.right_adjoint_is_ide e_src by auto
ultimately show ?thesis
using fg r' D.right_adjoints_compose F.preserves_right_adjoint by blast
qed
obtain f' where f': "D.adjoint_pair f' (F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src)"
using 1 by auto
have f': "D.is_left_adjoint f' \<and> F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src \<cong>\<^sub>D (f')\<^sup>*\<^sup>D"
using f' D.left_adjoint_determines_right_up_to_iso D.left_adjoint_extends_to_adjoint_pair
by blast
have "r' \<cong>\<^sub>D d_trg \<star>\<^sub>D (e_trg \<star>\<^sub>D r' \<star>\<^sub>D d_src) \<star>\<^sub>D e_src"
using r' r'_in_hhom D.isomorphic_def eqs.\<psi>_in_hom eqs.\<psi>_components_are_iso
D.isomorphic_symmetric D.ide_char eq_src.antipar(2) eq_trg.antipar(2)
by metis
also have 1: "... \<cong>\<^sub>D d_trg \<star>\<^sub>D F (F.right_map r') \<star>\<^sub>D e_src"
proof -
have "e_trg \<star>\<^sub>D r' \<star>\<^sub>D d_src \<cong>\<^sub>D F (F.right_map r')"
proof -
have "D.in_hom (F.counit\<^sub>1 r')
(r' \<star>\<^sub>D d_src) (F.counit\<^sub>0 (trg\<^sub>D r') \<star>\<^sub>D F (F.right_map r'))"
unfolding d_src_def
using r' E.\<epsilon>.map\<^sub>1_in_hom(2) [of r'] by simp
hence "r' \<star>\<^sub>D d_src \<cong>\<^sub>D F.counit\<^sub>0 (trg\<^sub>D r') \<star>\<^sub>D F (F.right_map r')"
using r' D.isomorphic_def E.\<epsilon>.iso_map\<^sub>1_ide by auto
thus ?thesis
using r' e_trg_def E.\<epsilon>.components_are_equivalences D.isomorphic_symmetric
D.quasi_inverse_transpose(2)
[of "F.counit\<^sub>0 (trg\<^sub>D r')" "F (F.right_map r')" "r' \<star>\<^sub>D d_src"]
by (metis D.ide_hcomp D.in_hhomE E.G.preserves_ide F.preserves_ide d_trg_def
eq_src.antipar(2) eq_src.ide_left eq_trg.ide_left r'_in_hhom)
qed
thus ?thesis
using r' d_src e_trg D.hcomp_isomorphic_ide D.hcomp_ide_isomorphic
by (metis D.hseqI' D.ideD(1) D.in_hhomE D.src_hcomp D.trg_hcomp eq_src.antipar(1)
eq_src.ide_right eq_trg.antipar(1) eq_trg.ide_left)
qed
also have 2: "... \<cong>\<^sub>D d_trg \<star>\<^sub>D (F g \<star>\<^sub>D F f\<^sup>*\<^sup>C) \<star>\<^sub>D e_src"
proof -
have "F (F.right_map r') \<cong>\<^sub>D F g \<star>\<^sub>D F f\<^sup>*\<^sup>C"
proof -
have "F (F.right_map r') \<cong>\<^sub>D F (g \<star>\<^sub>C f\<^sup>*\<^sup>C)"
using fg F.preserves_iso C.isomorphic_def D.isomorphic_def by auto
also have "F (g \<star>\<^sub>C f\<^sup>*\<^sup>C) \<cong>\<^sub>D F g \<star>\<^sub>D F f\<^sup>*\<^sup>C"
using fg F.weakly_preserves_hcomp [of g "f\<^sup>*\<^sup>C"] D.isomorphic_symmetric
C.left_adjoint_is_ide C.right_adjoint_is_ide
by (meson C.hseq_char C.ideD(1) C.isomorphic_implies_ide(2) C.right_adjoint_simps(1))
finally show ?thesis by simp
qed
thus ?thesis
using D.hcomp_ide_isomorphic D.hcomp_isomorphic_ide
by (metis 1 D.hseqE D.ideD(1) D.isomorphic_implies_hpar(2)
eq_src.ide_right eq_trg.ide_left)
qed
also have 3: "... \<cong>\<^sub>D (d_trg \<star>\<^sub>D F g) \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src"
proof -
have 4: "... \<cong>\<^sub>D d_trg \<star>\<^sub>D F g \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src"
using fg C.left_adjoint_is_ide C.right_adjoint_is_ide D.hcomp_assoc_isomorphic
D.hcomp_ide_isomorphic [of "d_trg"]
by (metis 2 C.right_adjoint_simps(1) D.hseqE D.ideD(1) D.isomorphic_implies_hpar(2)
D.src_hcomp F.preserves_ide eq_src.ide_right eq_trg.ide_left)
also have "... \<cong>\<^sub>D (d_trg \<star>\<^sub>D F g) \<star>\<^sub>D F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src"
using fg C.left_adjoint_is_ide C.right_adjoint_is_ide D.isomorphic_symmetric
D.hcomp_assoc_isomorphic [of "d_trg" "F g" "F f\<^sup>*\<^sup>C \<star>\<^sub>D e_src"]
by (metis 4 D.adjoint_pair_antipar(1) D.hseqE D.ideD(1) D.isomorphic_implies_hpar(1)
F.preserves_ide eq_trg.ide_left f' g' g'_def)
finally show ?thesis by blast
qed
also have "... \<cong>\<^sub>D g' \<star>\<^sub>D f'\<^sup>*\<^sup>D"
using g'_def f'
by (metis 3 D.adjoint_pair_antipar(1) D.hcomp_ide_isomorphic D.hseq_char D.ideD(1)
D.isomorphic_implies_ide(2) g')
finally have "D.isomorphic r' (g' \<star>\<^sub>D f'\<^sup>*\<^sup>D)"
by simp
thus "\<exists>f' g'. D.is_left_adjoint f' \<and> D.is_left_adjoint g' \<and> r' \<cong>\<^sub>D g' \<star>\<^sub>D f'\<^sup>*\<^sup>D"
using f' g' by auto
qed
show "\<And>f f' \<mu> \<mu>'. \<lbrakk> D.is_left_adjoint f; D.is_left_adjoint f';
\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>D f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>D f'\<guillemotright> \<rbrakk> \<Longrightarrow> D.iso \<mu> \<and> D.iso \<mu>' \<and> \<mu> = \<mu>'"
proof -
fix f f' \<mu> \<mu>'
assume f: "D.is_left_adjoint f"
and f': "D.is_left_adjoint f'"
and \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>D f'\<guillemotright>"
and \<mu>': "\<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>D f'\<guillemotright>"
have "C.is_left_adjoint (F.right_map f) \<and> C.is_left_adjoint (F.right_map f')"
using f f' E.G.preserves_left_adjoint by blast
moreover have "\<guillemotleft>F.right_map \<mu> : F.right_map f \<Rightarrow>\<^sub>C F.right_map f'\<guillemotright> \<and>
\<guillemotleft>F.right_map \<mu>' : F.right_map f \<Rightarrow>\<^sub>C F.right_map f'\<guillemotright>"
using \<mu> \<mu>' E.G.preserves_hom by simp
ultimately have "C.iso (F.right_map \<mu>) \<and> C.iso (F.right_map \<mu>') \<and>
F.right_map \<mu> = F.right_map \<mu>'"
using C.BS3 by blast
thus "D.iso \<mu> \<and> D.iso \<mu>' \<and> \<mu> = \<mu>'"
using \<mu> \<mu>' G.reflects_iso G.is_faithful by blast
qed
show "\<And>f g. \<lbrakk> D.is_left_adjoint f; D.is_left_adjoint g; src\<^sub>D f = src\<^sub>D g \<rbrakk>
\<Longrightarrow> \<exists>r \<rho>. tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r \<rho> f g"
proof -
fix f g
assume f: "D.is_left_adjoint f"
assume g: "D.is_left_adjoint g"
assume fg: "src\<^sub>D f = src\<^sub>D g"
have "C.is_left_adjoint (F.right_map f)"
using f E.G.preserves_left_adjoint by blast
moreover have "C.is_left_adjoint (F.right_map g)"
using g E.G.preserves_left_adjoint by blast
moreover have "src\<^sub>C (F.right_map f) = src\<^sub>C (F.right_map g)"
using f g D.left_adjoint_is_ide fg by simp
ultimately have
1: "\<exists>r \<rho>. tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> (F.right_map f) (F.right_map g)"
using C.BS2 by simp
obtain r \<rho> where
\<rho>: "tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> (F.right_map f) (F.right_map g)"
using 1 by auto
interpret \<rho>: tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C r \<rho> \<open>F.right_map f\<close> \<open>F.right_map g\<close>
using \<rho> by simp
obtain r' where
r': "D.ide r' \<and> D.in_hhom r' (trg\<^sub>D f) (trg\<^sub>D g) \<and> C.isomorphic (F.right_map r') r"
using f g \<rho>.ide_base \<rho>.tab_in_hom G.locally_essentially_surjective
by (metis D.obj_trg E.G.preserves_reflects_arr E.G.preserves_trg \<rho>.leg0_simps(2-3)
\<rho>.leg1_simps(2,4) \<rho>.base_in_hom(1))
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : r \<Rightarrow>\<^sub>C F.right_map r'\<guillemotright> \<and> C.iso \<phi>"
using r' C.isomorphic_symmetric by blast
have \<sigma>: "tabulation V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C
(F.right_map r') ((\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>) (F.right_map f) (F.right_map g)"
using \<phi> \<rho>.is_preserved_by_base_iso by simp
have 1: "\<exists>\<rho>'. \<guillemotleft>\<rho>' : g \<Rightarrow>\<^sub>D H\<^sub>D r' f\<guillemotright> \<and>
F.right_map \<rho>' = V\<^sub>C (F.right_cmp (r', f)) (V\<^sub>C (H\<^sub>C \<phi> (F.right_map f)) \<rho>)"
proof -
have "D.ide g"
by (simp add: D.left_adjoint_is_ide g)
moreover have "D.ide (H\<^sub>D r' f)"
using f r' D.left_adjoint_is_ide by auto
moreover have "src\<^sub>D g = src\<^sub>D (H\<^sub>D r' f)"
using fg by (simp add: calculation(2))
moreover have "trg\<^sub>D g = trg\<^sub>D (H\<^sub>D r' f)"
using calculation(2) r' by auto
moreover have "\<guillemotleft>F.right_cmp (r', f) \<cdot>\<^sub>C (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho> :
F.right_map g \<Rightarrow>\<^sub>C F.right_map (r' \<star>\<^sub>D f)\<guillemotright>"
using f g r' \<phi> D.left_adjoint_is_ide \<rho>.ide_base
by (intro C.comp_in_homI, auto)
ultimately show ?thesis
using G.locally_full by simp
qed
obtain \<rho>' where \<rho>': "\<guillemotleft>\<rho>' : g \<Rightarrow>\<^sub>D H\<^sub>D r' f\<guillemotright> \<and>
F.right_map \<rho>' = F.right_cmp (r', f) \<cdot>\<^sub>C (\<phi> \<star>\<^sub>C F.right_map f) \<cdot>\<^sub>C \<rho>"
using 1 by auto
have "tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r' \<rho>' f g"
proof -
have "V\<^sub>C (C.inv (F.right_cmp (r', f))) (F.right_map \<rho>') =
V\<^sub>C (H\<^sub>C \<phi> (F.right_map f)) \<rho>"
using r' f \<rho>' C.comp_assoc C.comp_cod_arr
C.invert_side_of_triangle(1)
[of "F.right_map \<rho>'" "F.right_cmp (r', f)" "V\<^sub>C (H\<^sub>C \<phi> (F.right_map f)) \<rho>"]
by (metis D.adjoint_pair_antipar(1) D.arrI D.in_hhomE E.G.cmp_components_are_iso
E.G.preserves_arr)
thus ?thesis
using \<sigma> \<rho>' G.reflects_tabulation
by (simp add: D.left_adjoint_is_ide f r')
qed
thus "\<exists>r' \<rho>'. tabulation V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D r' \<rho>' f g"
by auto
qed
qed
qed
subsection "Span(C) is a Bicategory of Spans"
text \<open>
We first consider an arbitrary 1-cell \<open>r\<close> in \<open>Span(C)\<close>, and show that it has a tabulation
as a span of maps. This is CKS Proposition 3 (stated more strongly to assert that
the ``output leg'' can also be taken to be a map, which the proof shows already).
\<close>
locale identity_arrow_in_span_bicategory = (* 20 sec *)
span_bicategory C prj0 prj1 +
r: identity_arrow_of_spans C r
for C :: "'a comp" (infixr "\<cdot>" 55)
and prj0 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>0[_, _]")
and prj1 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<p>\<^sub>1[_, _]")
and r :: "'a arrow_of_spans_data"
begin
text \<open>
CKS say: ``Suppose \<open>r = (r\<^sub>0, R, r\<^sub>1): A \<rightarrow> B\<close> and put \<open>f = (1, R, r\<^sub>0)\<close>, \<open>g = (1, R, r\<^sub>1)\<close>.
Let \<open>k\<^sub>0, k\<^sub>1\<close> form a kernel pair for \<open>r\<^sub>0\<close> and define \<open>\<rho>\<close> by \<open>k\<^sub>0\<rho> = k\<^sub>1\<rho> = 1\<^sub>R\<close>.''
\<close>
abbreviation ra where "ra \<equiv> r.dom.apex"
abbreviation r0 where "r0 \<equiv> r.dom.leg0"
abbreviation r1 where "r1 \<equiv> r.dom.leg1"
abbreviation f where "f \<equiv> mkIde ra r0"
abbreviation g where "g \<equiv> mkIde ra r1"
abbreviation k0 where "k0 \<equiv> \<p>\<^sub>0[r0, r0]"
abbreviation k1 where "k1 \<equiv> \<p>\<^sub>1[r0, r0]"
text \<open>
Here \<open>ra\<close> is the apex \<open>R\<close> of the span \<open>(r\<^sub>0, R, r\<^sub>1)\<close>, and the spans \<open>f\<close> and \<open>g\<close> also have
that same 0-cell as their apex. The tabulation 2-cell \<open>\<rho>\<close> has to be an arrow of spans
from \<open>g = (1, R, r\<^sub>1)\<close> to \<open>r \<star> f\<close>, which is the span \<open>(k\<^sub>0, r\<^sub>1 \<cdot> k\<^sub>1)\<close>.
\<close>
abbreviation \<rho> where "\<rho> \<equiv> \<lparr>Chn = \<langle>ra \<lbrakk>r0, r0\<rbrakk> ra\<rangle>,
Dom = \<lparr>Leg0 = ra, Leg1 = r1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>"
lemma has_tabulation:
shows "tabulation vcomp hcomp assoc unit src trg r \<rho> f g"
and "is_left_adjoint f" and "is_left_adjoint g"
proof -
have ide_f: "ide f"
using ide_mkIde r.dom.leg_in_hom(1) C.arr_dom C.dom_dom r.dom.apex_def r.dom.leg_simps(1)
by presburger
interpret f: identity_arrow_of_spans C f
using ide_f ide_char' by auto
have ide_g: "ide g"
using ide_mkIde r.dom.leg_in_hom
by (metis C.arr_dom C.dom_dom r.dom.leg_simps(3) r.dom.leg_simps(4))
interpret g: identity_arrow_of_spans C g
using ide_g ide_char' by auto
show "is_left_adjoint f"
using is_left_adjoint_char [of f] ide_f by simp
show "is_left_adjoint g"
using is_left_adjoint_char [of g] ide_g by simp
have ide_r: "ide r"
using ide_char' r.identity_arrow_of_spans_axioms by auto
have src_r: "src r = mkObj (C.cod r0)"
by (simp add: ide_r src_def)
have trg_r: "trg r = mkObj (C.cod r1)"
by (simp add: ide_r trg_def)
have src_f: "src f = mkObj ra"
using ide_f src_def by auto
have trg_f: "trg f = mkObj (C.cod r0)"
using ide_f trg_def by auto
have src_g: "src g = mkObj ra"
using ide_g src_def by auto
have trg_g: "trg g = mkObj (C.cod r1)"
using ide_g trg_def by auto
have hseq_rf: "hseq r f"
using ide_r ide_f src_r trg_f by simp
interpret rf: two_composable_arrows_of_spans C prj0 prj1 r f
using hseq_rf hseq_char by unfold_locales auto
interpret rf: two_composable_identity_arrows_of_spans C prj0 prj1 r f ..
interpret rf: identity_arrow_of_spans C \<open>r \<star> f\<close>
using rf.ide_composite ide_char' by auto
let ?rf = "r \<star> f"
(* TODO: Put this expansion into two_composable_identity_arrows_of_spans. *)
have rf: "?rf = \<lparr>Chn = r0 \<down>\<down> r0,
Dom = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>"
unfolding hcomp_def chine_hcomp_def
using hseq_rf C.comp_cod_arr by auto
interpret Cod_rf: span_in_category C \<open>\<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<close>
using ide_r ide_f rf C.comp_cod_arr
by unfold_locales auto
have Dom_g: "Dom g = \<lparr>Leg0 = ra, Leg1 = r1\<rparr>" by simp
interpret Dom_g: span_in_category C \<open>\<lparr>Leg0 = ra, Leg1 = r1\<rparr>\<close>
using Dom_g g.dom.span_in_category_axioms by simp
interpret Dom_\<rho>: span_in_category C \<open>Dom \<rho>\<close>
using Dom_g g.dom.span_in_category_axioms by simp
interpret Cod_\<rho>: span_in_category C \<open>Cod \<rho>\<close>
using rf Cod_rf.span_in_category_axioms by simp
interpret \<rho>: arrow_of_spans C \<rho>
using Dom_\<rho>.apex_def Cod_\<rho>.apex_def C.comp_assoc C.comp_arr_dom
by unfold_locales auto
have \<rho>: "\<guillemotleft>\<rho> : g \<Rightarrow> r \<star> f\<guillemotright>"
proof
show 1: "arr \<rho>"
using arr_char \<rho>.arrow_of_spans_axioms by simp
show "dom \<rho> = g"
using 1 dom_char ideD(2) ide_g by fastforce
show "cod \<rho> = r \<star> f"
using 1 cod_char rf Cod_rf.apex_def by simp
qed
show "tabulation vcomp hcomp assoc unit src trg r \<rho> f g"
proof -
interpret T: tabulation_data vcomp hcomp assoc unit src trg r \<rho> f g
using ide_f \<rho> by unfold_locales auto
show ?thesis
proof
show T1: "\<And>u \<omega>.
\<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "ide u"
assume \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
show "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof -
interpret u: identity_arrow_of_spans C u
using u ide_char' by auto
have v: "ide (dom \<omega>)"
using \<omega> by auto
interpret v: identity_arrow_of_spans C \<open>dom \<omega>\<close>
using v ide_char' by auto
interpret \<omega>: arrow_of_spans C \<omega>
using \<omega> arr_char by auto
have hseq_ru: "hseq r u"
using u \<omega> ide_cod by fastforce
interpret ru: two_composable_arrows_of_spans C prj0 prj1 r u
using hseq_ru hseq_char by unfold_locales auto
interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u ..
text \<open>
CKS say:
``We must show that \<open>(f, \<rho>, g)\<close> is a wide tabulation of \<open>r\<close>.
Take \<open>u = (u\<^sub>0, U, u\<^sub>1): X \<rightarrow> A\<close>, \<open>v = (v\<^sub>0, V, v\<^sub>1): X \<rightarrow> B\<close>,
\<open>\<omega>: v \<Rightarrow> ru\<close> as in \<open>T1\<close>. Let \<open>P\<close> be the pullback of \<open>u\<^sub>1, r\<^sub>0\<close>.
Let \<open>w = (v\<^sub>0, V, p\<^sub>1\<omega>): X \<rightarrow> R\<close>, \<open>\<theta> = p\<^sub>0\<omega>: fw \<Rightarrow> u\<close>,
\<open>\<nu> = 1: v \<Rightarrow> gw\<close>; so \<open>\<omega> = (r\<theta>)(\<rho>w)\<nu>\<close> as required.''
\<close>
let ?R = "r.apex"
let ?A = "C.cod r0"
let ?B = "C.cod r1"
let ?U = "u.apex"
let ?u0 = "u.leg0"
let ?u1 = "u.leg1"
let ?X = "C.cod ?u0"
let ?V = "v.apex"
let ?v0 = "v.leg0"
let ?v1 = "v.leg1"
let ?\<omega> = "\<omega>.chine"
let ?P = "r0 \<down>\<down> ?u1"
let ?p0 = "\<p>\<^sub>0[r0, ?u1]"
let ?p1 = "\<p>\<^sub>1[r0, ?u1]"
let ?w1 = "?p1 \<cdot> ?\<omega>"
define w where "w = mkIde ?v0 ?w1"
let ?Q = "?R \<down>\<down> ?w1"
let ?q1 = "\<p>\<^sub>1[?R, ?w1]"
let ?\<rho> = "\<langle>?R \<lbrakk>r0, r0\<rbrakk> ?R\<rangle>"
have P: "?P = ru.apex"
using ru.apex_composite by auto
have Chn_\<omega>: "\<guillemotleft>?\<omega> : ?V \<rightarrow>\<^sub>C ?P\<guillemotright>"
using P Chn_in_hom \<omega> by force
have p0\<omega>: "\<guillemotleft>?p0 \<cdot> ?\<omega> : ?V \<rightarrow>\<^sub>C ?U\<guillemotright>"
using Chn_\<omega> ru.legs_form_cospan by auto
have w1: "\<guillemotleft>?w1 : ?V \<rightarrow>\<^sub>C ?R\<guillemotright>"
using Chn_\<omega> ru.legs_form_cospan r.dom.apex_def by blast
have r1w1: "r1 \<cdot> ?w1 = ?v1"
by (metis C.comp_assoc T.base_simps(3) \<omega> \<omega>.leg1_commutes
arrow_of_spans_data.select_convs(3) cod_char dom_char ideD(1) ideD(2)
in_homE ru.composite_in_hom ru.leg1_composite u v)
have w: "ide w"
unfolding w_def
using P \<omega> w1 by (intro ide_mkIde, auto)
interpret w: identity_arrow_of_spans C w
using w_def w ide_char' by auto
have hseq_fw: "hseq f w"
using w_def ide_f w src_def trg_def w1 by auto
interpret fw: two_composable_arrows_of_spans C prj0 prj1 f w
using w_def hseq_fw hseq_char by unfold_locales auto
interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w ..
have hseq_gw: "hseq g w"
using w_def ide_g w src_def trg_def w1 by auto
interpret gw: two_composable_arrows_of_spans C prj0 prj1 g w
using hseq_gw hseq_char by unfold_locales auto
interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w ..
interpret rfw: three_composable_arrows_of_spans C prj0 prj1 r f w ..
interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w ..
have arfw: "\<guillemotleft>\<a>[r, f, w] : (r \<star> f) \<star> w \<Rightarrow> r \<star> f \<star> w\<guillemotright>"
using fw.composable ide_f ide_r w rf.composable by auto
have fw_apex_eq: "fw.apex = ?R \<down>\<down> ?w1"
using w_def fw.dom.apex_def hcomp_def hseq_fw hseq_char \<omega> C.arr_dom_iff_arr
C.pbdom_def fw.chine_eq_apex fw.chine_simps(1)
by auto
have gw_apex_eq: "gw.apex = ?R \<down>\<down> ?w1"
using w_def \<omega> w1 gw.dom.apex_def hcomp_def hseq_gw hseq_char by auto
text \<open>
Well, CKS say take \<open>\<theta> = p\<^sub>0\<omega>\<close>, but taking this literally and trying to define
\<open>\<theta>\<close> so that \<open>Chn \<theta> = ?p\<^sub>0 \<cdot> ?\<omega>\<close>, does not yield the required \<open>dom \<theta> = ?R \<down>\<down> ?w1\<close>.
We need \<open>\<guillemotleft>Chn \<theta> : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?U\<guillemotright>\<close>, so we have to compose with a
projection, which leads to: \<open>Chn \<theta> = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]\<close>.
What CKS say is only correct if the projection \<open>\<p>\<^sub>0[?R, ?w1]\<close> is an identity,
which can't be guaranteed for an arbitrary choice of pullbacks.
\<close>
define \<theta>
where
"\<theta> = \<lparr>Chn = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1], Dom = Dom (f \<star> w), Cod = Cod u\<rparr>"
interpret Dom_\<theta>: span_in_category C \<open>Dom \<theta>\<close>
using \<theta>_def fw.dom.span_in_category_axioms by simp
interpret Cod_\<theta>: span_in_category C \<open>Cod \<theta>\<close>
using \<theta>_def u.cod.span_in_category_axioms by simp
have Dom_\<theta>_leg0_eq: "Dom_\<theta>.leg0 = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Dom_\<theta>_leg1_eq: "Dom_\<theta>.leg1 = r0 \<cdot> ?q1"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Cod_\<theta>_leg0_eq: "Cod_\<theta>.leg0 = ?u0"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Cod_\<theta>_leg1_eq: "Cod_\<theta>.leg1 = ?u1"
using w_def \<theta>_def hcomp_def hseq_fw hseq_char by simp
have Chn_\<theta>_eq: "Chn \<theta> = ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using \<theta>_def by simp
interpret \<theta>: arrow_of_spans C \<theta>
proof
show 1: "\<guillemotleft>Chn \<theta> : Dom_\<theta>.apex \<rightarrow>\<^sub>C Cod_\<theta>.apex\<guillemotright>"
using \<theta>_def Chn_\<omega> ru.legs_form_cospan fw_apex_eq
by (intro C.in_homI, auto)
show "Cod_\<theta>.leg0 \<cdot> Chn \<theta> = Dom_\<theta>.leg0"
proof -
have "Cod_\<theta>.leg0 \<cdot> Chn \<theta> = ?u0 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using Cod_\<theta>_leg0_eq Chn_\<theta>_eq by simp
also have "... = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
proof -
have "?u0 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1] = (?u0 \<cdot> ?p0 \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by simp
also have "... = ?v0 \<cdot> \<p>\<^sub>0[?R, ?w1]"
proof -
have "?u0 \<cdot> ?p0 \<cdot> ?\<omega> = (?u0 \<cdot> ?p0) \<cdot> ?\<omega>"
using C.comp_assoc by simp
also have "... = \<omega>.cod.leg0 \<cdot> ?\<omega>"
proof -
have "\<omega>.cod.leg0 = ru.leg0"
using \<omega> cod_char hcomp_def hseq_ru by auto
also have "... = ?u0 \<cdot> ?p0"
using hcomp_def hseq_ru by auto
finally show ?thesis by simp
qed
also have "... = \<omega>.dom.leg0"
using \<omega>.leg0_commutes by simp
also have "... = ?v0"
using \<omega> dom_char by auto
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
also have "... = Dom_\<theta>.leg0"
using Dom_\<theta>_leg0_eq by simp
finally show ?thesis by blast
qed
show "Cod_\<theta>.leg1 \<cdot> Chn \<theta> = Dom_\<theta>.leg1"
proof -
have "Cod_\<theta>.leg1 \<cdot> Chn \<theta> = ?u1 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1]"
using Cod_\<theta>_leg1_eq Chn_\<theta>_eq by simp
also have "... = r0 \<cdot> ?q1"
proof -
have "?u1 \<cdot> ?p0 \<cdot> ?\<omega> \<cdot> \<p>\<^sub>0[?R, ?w1] = ((?u1 \<cdot> ?p0) \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by fastforce
also have "... = ((r0 \<cdot> ?p1) \<cdot> ?\<omega>) \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.pullback_commutes' ru.legs_form_cospan by simp
also have "... = r0 \<cdot> ?w1 \<cdot> \<p>\<^sub>0[?R, ?w1]"
using C.comp_assoc by fastforce
also have "... = r0 \<cdot> ?R \<cdot> ?q1"
using \<omega> C.in_homE C.pullback_commutes' w1 by auto
also have "... = r0 \<cdot> ?q1"
using \<omega> w1 C.comp_cod_arr by auto
finally show ?thesis by simp
qed
also have "... = Dom_\<theta>.leg1"
using Dom_\<theta>_leg1_eq by simp
finally show ?thesis by blast
qed
qed
text \<open>
Similarly, CKS say to take \<open>\<nu> = 1: v \<Rightarrow> gw\<close>, but obviously this can't be
interpreted literally, either, because \<open>v.apex\<close> and \<open>gw.apex\<close> are not equal.
Instead, we have to define \<open>\<nu>\<close> so that \<open>Chn \<nu> = C.inv \<p>\<^sub>0[?R, ?w1]\<close>,
noting that \<open>\<p>\<^sub>0[?R, ?w1]\<close> is the pullback of an identity,
hence is an isomorphism. Then \<open>?v0 \<cdot> \<p>\<^sub>0[?R, ?w1] \<cdot> Chn \<nu> = ?v0\<close> and
\<open>?v1 \<cdot> \<p>\<^sub>1[?R, ?w1] \<cdot> Chn \<nu> = ?v1 \<cdot> ?w1\<close>, showing that \<open>\<nu>\<close> is an arrow
of spans.
\<close>
let ?\<nu>' = "\<p>\<^sub>0[?R, ?w1]"
define \<nu>
where
"\<nu> = \<lparr>Chn = C.inv ?\<nu>', Dom = Dom (dom \<omega>), Cod = Cod (g \<star> w)\<rparr>"
have iso_\<nu>: "C.inverse_arrows ?\<nu>' (Chn \<nu>)"
using \<nu>_def \<omega> w1 C.iso_pullback_ide
by (metis C.inv_is_inverse C.seqE arrow_of_spans_data.select_convs(1)
r.chine_eq_apex r.chine_simps(1) r.chine_simps(3) r.cod_simps(1)
r.dom.apex_def r.dom.ide_apex r.dom.is_span r1w1 v.dom.leg_simps(3))
text \<open>
$$
\xymatrix{
&& X \\
&& V \ar[u]_{v_0} \ar[d]_{\omega} \ar@/ul50pt/[ddddll]_{v_1} \ar@/l30pt/[dd]_<>(0.7){w_1}\\
& R\downarrow\downarrow w_1 \ar[ur]^{\nu'} \ar[dd]_{q_1}
& r_0\downarrow\downarrow u_1 \ar[d]_{p_1} \ar@/dl10pt/[drr]_<>(0.4){p_0}
& R\downarrow\downarrow w_1 \ar[ul]_{\nu'} \ar[dd]^<>(0.7){q_1} \ar@ {.>}[dr]_{\theta}\\
&& R && U \ar@/ur20pt/[uuull]_{u_0} \ar[dd]^{u_1} \\
& R \ar[dl]_{r_1} \ar@ {<->}[ur]_{R} \ar@ {.>}[dr]^{\rho} \ar@/dl5pt/[ddr]_<>(0.4){R}
&& R \ar@ {<->}[ul]^{R} \ar[dr]^{r_0} \ar[ur]_{r_1}\\
B && r_0\downarrow\downarrow r_0 \ar[uu]_{k_0} \ar[d]^{k_1} \ar[uu] \ar[ur]_{k_0} && A \\
& & R \ar[ull]^{r_1} \ar[urr]_{r_0} \\
}
$$
\<close>
have w1_eq: "?w1 = ?q1 \<cdot> C.inv ?\<nu>'"
proof -
have "?R \<cdot> ?q1 = ?w1 \<cdot> ?\<nu>'"
using iso_\<nu> \<omega> w1 C.pullback_commutes [of ?R ?w1] by blast
hence "?q1 = ?w1 \<cdot> ?\<nu>'"
using \<omega> w1 C.comp_cod_arr by auto
thus ?thesis
using iso_\<nu> \<omega> w1 r.dom.apex_def r.cod.apex_def r.chine_eq_apex
C.invert_side_of_triangle(2)
by (metis C.isoI C.prj1_simps(1) arrow_of_spans_data.select_convs(3)
fw.legs_form_cospan(2) span_data.simps(1-2) w_def)
qed
interpret Dom_\<nu>: span_in_category C \<open>Dom \<nu>\<close>
using \<nu>_def v.dom.span_in_category_axioms by simp
interpret Cod_\<nu>: span_in_category C \<open>Cod \<nu>\<close>
using \<nu>_def gw.cod.span_in_category_axioms by simp
interpret \<nu>: arrow_of_spans C \<nu>
proof
show 1: "\<guillemotleft>Chn \<nu> : Dom_\<nu>.apex \<rightarrow>\<^sub>C Cod_\<nu>.apex\<guillemotright>"
proof
show "C.arr (Chn \<nu>)"
using iso_\<nu> by auto
show "C.dom (Chn \<nu>) = Dom_\<nu>.apex"
using \<nu>_def iso_\<nu> w1 gw_apex_eq by fastforce
show "C.cod (Chn \<nu>) = Cod_\<nu>.apex"
using \<nu>_def iso_\<nu> gw_apex_eq C.comp_inv_arr C.pbdom_def
by (metis C.cod_comp arrow_of_spans_data.select_convs(3)
gw.apex_composite gw.chine_composite gw.chine_simps(1) gw.chine_simps(3))
qed
show "Cod_\<nu>.leg0 \<cdot> Chn \<nu> = Dom_\<nu>.leg0"
using w_def \<nu>_def 1 iso_\<nu> hcomp_def hseq_gw C.comp_arr_inv
C.comp_assoc v.leg0_commutes
by auto
show "Cod_\<nu>.leg1 \<cdot> Chn \<nu> = Dom_\<nu>.leg1"
using w_def \<nu>_def hcomp_def hseq_gw C.comp_assoc w1_eq r1w1
by auto
qed
text \<open>
Now we can proceed to establishing the conclusions of \<open>T1\<close>.
\<close>
have "ide w \<and> \<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> dom \<rho> \<star> w\<guillemotright> \<and> iso \<nu> \<and>
T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof (intro conjI)
show ide_w: "ide w"
using w by blast
show \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
using \<theta>_def \<theta>.arrow_of_spans_axioms arr_char dom_char cod_char hseq_fw hseq_char
hcomp_def fw.chine_eq_apex
by auto
show \<nu>: "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> dom \<rho> \<star> w\<guillemotright>"
proof -
have "\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> g \<star> w\<guillemotright>"
using \<nu>_def \<omega> \<nu>.arrow_of_spans_axioms arr_char dom_char cod_char hseq_gw
hseq_char hcomp_def gw.chine_eq_apex
apply (intro in_homI) by auto
thus ?thesis
using T.tab_in_hom by simp
qed
show "iso \<nu>"
using iso_\<nu> iso_char arr_char \<nu>.arrow_of_spans_axioms by auto
show "T.composite_cell w \<theta> \<bullet> \<nu> = \<omega>"
proof (intro arr_eqI)
have \<rho>w: "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using w_def \<rho> ide_w hseq_rf hseq_fw hseq_gw by auto
have r\<theta>: "\<guillemotleft>r \<star> \<theta> : r \<star> f \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using arfw ide_r \<theta> fw.composite_simps(2) rf.composable by auto
have 1: "\<guillemotleft>T.composite_cell w \<theta> \<bullet> \<nu> : dom \<omega> \<Rightarrow> r \<star> u\<guillemotright>"
using \<nu> \<rho>w arfw r\<theta> by auto
show "par (T.composite_cell w \<theta> \<bullet> \<nu>) \<omega>"
using 1 \<omega> by (elim in_homE, auto)
show "Chn (T.composite_cell w \<theta> \<bullet> \<nu>) = ?\<omega>"
proof -
have 2: "Chn (T.composite_cell w \<theta> \<bullet> \<nu>) =
Chn (r \<star> \<theta>) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
proof -
have "Chn (T.composite_cell w \<theta> \<bullet> \<nu>) =
Chn (T.composite_cell w \<theta>) \<cdot> Chn \<nu>"
using 1 Chn_vcomp by blast
also have "... = (Chn (r \<star> \<theta>) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w)) \<cdot> Chn \<nu>"
proof -
have "seq (r \<star> \<theta>) (\<a>[r, f, w] \<bullet> (\<rho> \<star> w)) \<and> seq \<a>[r, f, w] (\<rho> \<star> w)"
using 1 by blast
thus ?thesis
using 1 Chn_vcomp by presburger
qed
finally show ?thesis
using C.comp_assoc by auto
qed
also have "... = ?\<omega>"
proof -
let ?LHS = "Chn (r \<star> \<theta>) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
have Chn_r\<theta>: "Chn (r \<star> \<theta>) = \<langle>r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?q1]
\<lbrakk>r0, ?u1\<rbrakk>
\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]\<rangle>"
using r\<theta> hcomp_def \<theta>_def chine_hcomp_def Dom_\<theta>_leg1_eq
by (metis arrI arrow_of_spans_data.select_convs(1,3)
hseq_char r.cod_simps(2) u.cod_simps(3))
have Chn_arfw: "Chn \<a>[r, f, w] = rfw.chine_assoc"
using \<alpha>_ide ide_f rf.composable fw.composable w by auto
have Chn_\<rho>w: "Chn (\<rho> \<star> w) = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
proof -
have "Chn (\<rho> \<star> w) =
chine_hcomp
\<lparr>Chn = ?\<rho>,
Dom = \<lparr>Leg0 = ?R, Leg1 = r1\<rparr>,
Cod = \<lparr>Leg0 = k0, Leg1 = r1 \<cdot> k1\<rparr>\<rparr>
\<lparr>Chn = v.apex,
Dom = \<lparr>Leg0 = ?v0, Leg1 = ?w1\<rparr>,
Cod = \<lparr>Leg0 = ?v0, Leg1 = ?w1\<rparr>\<rparr>"
using \<rho> ide_w hseq_rf hseq_char hcomp_def src_def trg_def
by (metis (no_types, lifting) \<rho>w arrI arrow_of_spans_data.select_convs(1)
v.dom.apex_def w_def)
also have "... = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?V \<cdot> ?\<nu>'\<rangle>"
unfolding chine_hcomp_def by simp
also have "... = \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
proof -
have "?V \<cdot> ?\<nu>' = ?\<nu>'"
using C.comp_ide_arr v.dom.ide_apex \<rho> w1 by auto
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
have 3: "C.seq ?p1 ?\<omega>"
using w1 by blast
moreover have 4: "C.seq ?p1 ?LHS"
proof
show "\<guillemotleft>?LHS : v.apex \<rightarrow>\<^sub>C ru.apex\<guillemotright>"
by (metis (no_types, lifting) 1 2 Chn_in_hom ru.chine_eq_apex
v.chine_eq_apex)
show "\<guillemotleft>?p1 : ru.apex \<rightarrow>\<^sub>C ?R\<guillemotright>"
using P C.prj1_in_hom ru.legs_form_cospan by fastforce
qed
moreover have "?p0 \<cdot> ?LHS = ?p0 \<cdot> ?\<omega>"
proof -
have "?p0 \<cdot> ?LHS =
(?p0 \<cdot> Chn (r \<star> \<theta>)) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using C.comp_assoc by simp
also have "... = (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]) \<cdot>
Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
proof -
have "?p0 \<cdot> Chn (r \<star> \<theta>) = \<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?q1]"
by (metis C.prj_tuple(1) Chn_r\<theta> \<theta>_def arrI Dom_\<theta>_leg1_eq
arrow_of_spans_data.select_convs(3) chine_hcomp_props(2)
hseq_char r.cod_simps(2) r\<theta> u.cod_simps(3))
thus ?thesis by argo
qed
also have
"... = ?p0 \<cdot> ?\<omega> \<cdot> (rfw.Prj\<^sub>0\<^sub>0 \<cdot> Chn \<a>[r, f, w]) \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using w_def \<theta>_def C.comp_assoc by simp
also have "... = ?p0 \<cdot> ?\<omega> \<cdot> (rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w)) \<cdot> Chn \<nu>"
using Chn_arfw rfw.prj_chine_assoc C.comp_assoc by simp
also have "... = ?p0 \<cdot> ?\<omega> \<cdot> ?\<nu>' \<cdot> Chn \<nu>"
proof -
have "rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w) = \<p>\<^sub>0[k0, ?w1] \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using w_def Chn_\<rho>w C.comp_cod_arr by simp
also have "... = ?\<nu>'"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(1) C.seqE
C.tuple_is_extensional Chn_\<rho>w 4)
finally have "rfw.Prj\<^sub>0 \<cdot> Chn (\<rho> \<star> w) = ?\<nu>'"
by blast
thus ?thesis by simp
qed
also have "... = ?p0 \<cdot> ?\<omega>"
using iso_\<nu> C.comp_arr_dom
by (metis (no_types, lifting) C.comp_arr_inv C.dom_comp \<nu>_def
\<omega>.chine_simps(1) 3 arrow_of_spans_data.simps(1) w1_eq)
finally show ?thesis by blast
qed
moreover have "?p1 \<cdot> ?LHS = ?w1"
proof -
have "?p1 \<cdot> ?LHS =
(?p1 \<cdot> Chn (r \<star> \<theta>)) \<cdot> Chn \<a>[r, f, w] \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using C.comp_assoc by simp
also have "... = (r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?q1]) \<cdot> Chn \<a>[r, f, w] \<cdot>
Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2) C.seqE
C.tuple_is_extensional Chn_r\<theta> 4)
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1 \<cdot> Chn \<a>[r, f, w]) \<cdot> Chn (\<rho> \<star> w) \<cdot> Chn \<nu>"
using w_def Dom_\<theta>_leg1_eq C.comp_assoc by simp
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w)) \<cdot> Chn \<nu>"
using Chn_arfw rfw.prj_chine_assoc(1) C.comp_assoc by simp
also have "... = r.chine \<cdot> ?q1 \<cdot> Chn \<nu>"
proof -
have "rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w) =
(k1 \<cdot> \<p>\<^sub>1[k0, ?w1]) \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using w_def Chn_\<rho>w C.comp_cod_arr by simp
also have "... = k1 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<langle>?\<rho> \<cdot> ?q1 \<lbrakk>k0, ?w1\<rbrakk> ?\<nu>'\<rangle>"
using C.comp_assoc by simp
also have "... = k1 \<cdot> ?\<rho> \<cdot> ?q1"
by (metis (no_types, lifting) C.not_arr_null C.prj_tuple(2)
C.seqE C.tuple_is_extensional Chn_\<rho>w 4)
also have "... = (k1 \<cdot> ?\<rho>) \<cdot> ?q1"
using C.comp_assoc by presburger
also have "... = ?R \<cdot> ?q1"
by simp
also have "... = ?q1"
by (metis Dom_\<theta>_leg1_eq C.comp_ide_arr C.prj1_simps(3)
C.prj1_simps_arr C.seqE C.seqI Dom_\<theta>.leg_simps(3)
r.dom.ide_apex)
finally have "rfw.Prj\<^sub>1\<^sub>1 \<cdot> Chn (\<rho> \<star> w) = ?q1"
by blast
thus ?thesis by simp
qed
also have "... = (r.chine \<cdot> ?p1) \<cdot> ?\<omega>"
using \<nu>_def w1_eq C.comp_assoc by simp
also have "... = ?w1"
using C.comp_cod_arr r.chine_eq_apex ru.prj_simps(1) by auto
finally show ?thesis by blast
qed
ultimately show ?thesis
using ru.legs_form_cospan C.prj_joint_monic by blast
qed
finally show ?thesis by argo
qed
qed
qed
thus ?thesis
using w_def by auto
qed
qed
show T2: "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w';
\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>;
T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<bullet> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume ide_w: "ide w"
assume ide_w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : f \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : f \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
assume E: "T.composite_cell w \<theta> = T.composite_cell w' \<theta>' \<bullet> \<beta>"
interpret T: uw\<theta>w'\<theta>'\<beta> vcomp hcomp assoc unit src trg r \<rho> f g u w \<theta> w' \<theta>' \<beta>
using ide_w ide_w' \<theta> \<theta>' \<beta> E comp_assoc
by unfold_locales auto
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)"
proof
interpret u: identity_arrow_of_spans C u
using T.uw\<theta>.u_simps(1) ide_char' by auto
interpret w: identity_arrow_of_spans C w
using ide_w ide_char' by auto
interpret w': identity_arrow_of_spans C w'
using ide_w' ide_char' by auto
let ?u0 = u.leg0
let ?u1 = u.leg1
let ?w0 = w.leg0
let ?w1 = w.leg1
let ?wa = "w.apex"
let ?w0' = w'.leg0
let ?w1' = w'.leg1
let ?wa' = "w'.apex"
let ?R = ra
let ?p0 = "\<p>\<^sub>0[?R, ?w1]"
let ?p0' = "\<p>\<^sub>0[?R, ?w1']"
let ?p1 = "\<p>\<^sub>1[?R, ?w1]"
let ?p1' = "\<p>\<^sub>1[?R, ?w1']"
interpret fw: two_composable_identity_arrows_of_spans C prj0 prj1 f w
using hseq_char by unfold_locales auto
interpret fw': two_composable_identity_arrows_of_spans C prj0 prj1 f w'
using hseq_char by unfold_locales auto
have hseq_gw: "hseq g w"
using T.leg1_in_hom by auto
interpret gw: two_composable_identity_arrows_of_spans C prj0 prj1 g w
using hseq_gw hseq_char by unfold_locales auto
have hseq_gw': "hseq g w'"
using T.leg1_in_hom by auto
interpret gw': two_composable_identity_arrows_of_spans C prj0 prj1 g w'
using hseq_gw' hseq_char by unfold_locales auto
interpret rfw: three_composable_identity_arrows_of_spans C prj0 prj1 r f w ..
interpret rfw: identity_arrow_of_spans C \<open>r \<star> f \<star> w\<close>
using rfw.composites_are_identities ide_char' by auto
interpret rfw': three_composable_arrows_of_spans C prj0 prj1 r f w' ..
interpret rfw': three_composable_identity_arrows_of_spans C prj0 prj1 r f w' ..
interpret rfw': identity_arrow_of_spans C \<open>r \<star> f \<star> w'\<close>
using rfw'.composites_are_identities ide_char' by auto
have \<rho>w: "\<guillemotleft>\<rho> \<star> w : g \<star> w \<Rightarrow> (r \<star> f) \<star> w\<guillemotright>"
using \<rho> hseq_gw by blast
interpret \<rho>w: two_composable_arrows_of_spans C prj0 prj1 \<rho> w
using \<rho>w by unfold_locales auto
have \<rho>w': "\<guillemotleft>\<rho> \<star> w' : g \<star> w' \<Rightarrow> (r \<star> f) \<star> w'\<guillemotright>"
using \<rho> hseq_gw' by blast
interpret \<rho>w': two_composable_arrows_of_spans C prj0 prj1 \<rho> w'
using \<rho>w' by unfold_locales auto
have arfw: "\<guillemotleft>\<a>[r, f, w] : (r \<star> f) \<star> w \<Rightarrow> r \<star> f \<star> w\<guillemotright>"
using fw.composable ide_f ide_r ide_w rf.composable by auto
have arfw': "\<guillemotleft>\<a>[r, f, w'] : (r \<star> f) \<star> w' \<Rightarrow> r \<star> f \<star> w'\<guillemotright>"
using fw'.composable ide_f ide_r ide_w' rf.composable by auto
have r\<theta>: "\<guillemotleft>r \<star> \<theta> : r \<star> f \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
by fastforce
interpret Dom_\<theta>: span_in_category C \<open>Dom \<theta>\<close>
using fw.dom.span_in_category_axioms
by (metis \<theta> arrow_of_spans_data.select_convs(2) in_homE dom_char)
interpret Cod_\<theta>: span_in_category C \<open>Cod \<theta>\<close>
using \<theta> u.cod.span_in_category_axioms cod_char by auto
interpret \<theta>: arrow_of_spans C \<theta>
using arr_char T.uw\<theta>.\<theta>_simps(1) by auto
interpret r\<theta>: two_composable_arrows_of_spans C prj0 prj1 r \<theta>
using r\<theta> by unfold_locales auto
have r\<theta>': "\<guillemotleft>r \<star> \<theta>' : r \<star> f \<star> w' \<Rightarrow> r \<star> u\<guillemotright>"
by fastforce
interpret Dom_\<theta>': span_in_category C \<open>Dom \<theta>'\<close>
using fw'.dom.span_in_category_axioms
by (metis \<theta>' arrow_of_spans_data.select_convs(2) in_homE dom_char)
interpret Cod_\<theta>': span_in_category C \<open>Cod \<theta>'\<close>
using \<theta>' u.cod.span_in_category_axioms cod_char by auto
interpret \<theta>': arrow_of_spans C \<theta>'
using arr_char T.uw'\<theta>'.\<theta>_simps(1) by auto
interpret r\<theta>': two_composable_arrows_of_spans C prj0 prj1 r \<theta>'
using r\<theta>' by unfold_locales auto
have 7: "\<guillemotleft>T.composite_cell w' \<theta>' \<bullet> \<beta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<beta> \<rho>w' arfw' r\<theta>' by auto
have 8: "\<guillemotleft>T.composite_cell w \<theta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<rho>w arfw r\<theta> by auto
interpret ru: two_composable_identity_arrows_of_spans C prj0 prj1 r u
using hseq_char by unfold_locales auto
interpret Dom_\<beta>: span_in_category C \<open>Dom \<beta>\<close>
using \<beta> fw.dom.span_in_category_axioms arr_char
by (metis comp_arr_dom in_homE gw.cod.span_in_category_axioms seq_char)
interpret Cod_\<beta>: span_in_category C \<open>Cod \<beta>\<close>
using \<beta> fw.cod.span_in_category_axioms arr_char
by (metis (no_types, lifting) comp_arr_dom ideD(2) in_homI
gw'.cod.span_in_category_axioms gw'.chine_is_identity hseq_gw' seqI'
seq_char ide_char)
interpret \<beta>: arrow_of_spans C \<beta>
using \<beta> arr_char by auto
text \<open>
CKS say: ``Take \<open>u\<close>, \<open>w\<close>, \<open>w'\<close>, \<open>\<theta>\<close>, \<open>\<theta>'\<close> as in \<open>T2\<close> and note that
\<open>fw = (w\<^sub>0, W, r\<^sub>0 w\<^sub>1)\<close>, \<open>gw = (w\<^sub>0, W, r\<^sub>1 w\<^sub>1)\<close>, \emph{etc}.
So \<open>\<beta>: W \<rightarrow> W'\<close> satisfies \<open>w\<^sub>0 = w\<^sub>0' \<beta>\<close>, \<open>r\<^sub>1 w\<^sub>1 = r\<^sub>1 w\<^sub>1' \<beta>\<close>.
But the equation \<open>(r\<theta>)(\<rho>w) = (r\<theta>')(\<rho>w')\<beta>\<close> gives \<open>w\<^sub>1 = w\<^sub>1'\<close>.
So \<open>\<gamma> = \<beta> : w \<Rightarrow> w'\<close> is unique with \<open>\<beta> = g \<gamma>, \<theta> = \<theta>' (f \<gamma>).\<close>''
Once again, there is substantial punning in the proof sketch given by CKS.
We can express \<open>fw\<close> and \<open>gw\<close> almost in the form they indicate, but projections
are required.
\<close>
have cospan: "C.cospan ?R ?w1"
using hseq_char [of \<rho> w] src_def trg_def by auto
have cospan': "C.cospan ?R ?w1'"
using hseq_char [of \<rho> w'] src_def trg_def by auto
have fw: "f \<star> w = \<lparr>Chn = ?R \<down>\<down> ?w1,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r0 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r0 \<cdot> ?p1\<rparr>\<rparr>"
using ide_f hseq_char hcomp_def chine_hcomp_def fw.dom.apex_def cospan
fw.chine_eq_apex
by auto
have gw: "g \<star> w = \<lparr>Chn = ?R \<down>\<down> ?w1,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>\<rparr>"
using hseq_gw hseq_char hcomp_def chine_hcomp_def gw.dom.apex_def cospan
gw.chine_eq_apex
by auto
have fw': "f \<star> w' = \<lparr>Chn = ?R \<down>\<down> ?w1',
Dom = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r0 \<cdot> ?p1'\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r0 \<cdot> ?p1'\<rparr>\<rparr>"
using ide_f hseq_char hcomp_def chine_hcomp_def fw'.dom.apex_def cospan'
fw'.chine_eq_apex
by auto
have gw': "g \<star> w' = \<lparr>Chn = ?R \<down>\<down> ?w1',
Dom = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>\<rparr>"
using hseq_gw' hseq_char hcomp_def chine_hcomp_def gw'.dom.apex_def cospan'
gw'.chine_eq_apex
by auto
text \<open>
Note that \<open>?p0\<close> and \<open>?p0'\<close> are only isomorphisms, not identities,
and we have \<open>?p1\<close> (which equals \<open>?w1 \<cdot> ?p0\<close>) and \<open>?p1'\<close> (which equals \<open>?w1' \<cdot> ?p0'\<close>)
in place of \<open>?w1\<close> and \<open>?w1'\<close>.
\<close>
text \<open>
The following diagram summarizes the
various given and defined arrows involved in the proof.
We have deviated slightly here from the nomenclature used in in CKS.
We prefer to use \<open>W\<close> and \<open>W'\<close> to denote the apexes of
\<open>w\<close> and \<open>w'\<close>, respectively.
We already have the expressions \<open>?R \<down>\<down> ?w1\<close> and \<open>?R \<down>\<down> ?w1'\<close> for the
apexes of \<open>fw\<close> and \<open>fw'\<close> (which are the same as the apexes of
\<open>gw\<close> and \<open>gw'\<close>, respectively) and we will not use any abbreviation for them.
\<close>
text \<open>
$$
\xymatrix{
&&& X \\
&& W \ar[ur]^{w_0} \ar[dr]_{w_1} \ar@ {.>}[rr]^{\gamma}
&& W' \ar[ul]_{w_0'} \ar[dl]^{w_1'} && U \ar@/r10pt/[dddl]^{u_1} \ar@/u7pt/[ulll]_{u_0}\\
& R\downarrow\downarrow w_1 \ar[ur]_{p_0} \ar[dr]^{p_1} \ar@/d15pt/[rrrr]_{\beta}
\ar@/u100pt/[urrrrr]^{\theta}
&& R && R \downarrow\downarrow w_1' \ar[ul]^{p_0'} \ar[dl]^{p_1'} \ar[ur]_{\theta'} \\
&& R \ar@ {.>}[dr]_{\rho} \ar@/dl7pt/[ddr]_{R} \ar[ur]_{R} \ar[dl]_{r_1} \ar@ {<->}[rr]_{R}
&& R \ar[ul]^{R} \ar[dr]_{r_0} \\
& B && r_0 \downarrow\downarrow r_0 \ar[d]^{k_1} \ar[ur]_{k_0} && A \\
&&& R \ar@/dr10pt/[urr]_{r_0} \ar@/dl5pt/[ull]^{r_1}
}
$$
\<close>
have Chn_\<beta>: "\<guillemotleft>\<beta>.chine: ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?R \<down>\<down> ?w1'\<guillemotright>"
using gw gw' Chn_in_hom \<beta> gw'.chine_eq_apex gw.chine_eq_apex by force
have \<beta>_eq: "\<beta> = \<lparr>Chn = \<beta>.chine,
Dom = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>,
Cod = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have Dom_\<beta>_eq: "Dom \<beta> = \<lparr>Leg0 = ?w0 \<cdot> ?p0, Leg1 = r1 \<cdot> ?p1\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have Cod_\<beta>_eq: "Cod \<beta> = \<lparr>Leg0 = ?w0' \<cdot> ?p0', Leg1 = r1 \<cdot> ?p1'\<rparr>"
using \<beta> gw gw' dom_char cod_char by auto
have \<beta>0: "?w0 \<cdot> ?p0 = ?w0' \<cdot> ?p0' \<cdot> \<beta>.chine"
using Dom_\<beta>_eq Cod_\<beta>_eq \<beta>.leg0_commutes C.comp_assoc by simp
have \<beta>1: "r1 \<cdot> ?p1 = r1 \<cdot> ?p1' \<cdot> \<beta>.chine"
using Dom_\<beta>_eq Cod_\<beta>_eq \<beta>.leg1_commutes C.comp_assoc by simp
have Dom_\<theta>_0: "Dom_\<theta>.leg0 = ?w0 \<cdot> ?p0"
using arrI dom_char fw T.uw\<theta>.\<theta>_simps(4) by auto
have Cod_\<theta>_0: "Cod_\<theta>.leg0 = ?u0"
using \<theta> cod_char by auto
have Dom_\<theta>_1: "Dom_\<theta>.leg1 = r0 \<cdot> ?p1"
using arrI dom_char fw T.uw\<theta>.\<theta>_simps(4) by auto
have Cod_\<theta>_1: "Cod_\<theta>.leg1 = ?u1"
using T.uw\<theta>.\<theta>_simps(5) cod_char by auto
have Dom_\<theta>'_0: "Dom_\<theta>'.leg0 = ?w0' \<cdot> ?p0'"
using dom_char fw' T.uw'\<theta>'.\<theta>_simps(4) by auto
have Cod_\<theta>'_0: "Cod_\<theta>'.leg0 = ?u0"
using T.uw'\<theta>'.\<theta>_simps(5) cod_char by auto
have Dom_\<theta>'_1: "Dom_\<theta>'.leg1 = r0 \<cdot> ?p1'"
using dom_char fw' T.uw'\<theta>'.\<theta>_simps(4) by auto
have Cod_\<theta>'_1: "Cod_\<theta>'.leg1 = ?u1"
using T.uw'\<theta>'.\<theta>_simps(5) cod_char by auto
have Dom_\<rho>_0: "Dom_\<rho>.leg0 = ?R"
by simp
have Dom_\<rho>_1: "Dom_\<rho>.leg1 = r1"
by simp
have Cod_\<rho>_0: "Cod_\<rho>.leg0 = k0"
by simp
have Cod_\<rho>_1: "Cod_\<rho>.leg1 = r1 \<cdot> k1"
by simp
have Chn_r\<theta>: "\<guillemotleft>r\<theta>.chine : rfw.chine \<rightarrow>\<^sub>C ru.chine\<guillemotright>"
using r\<theta>.chine_composite_in_hom ru.chine_composite rfw.chine_composite
Cod_\<theta>_1 Dom_\<theta>_1 fw.leg1_composite
by auto
have Chn_r\<theta>_eq: "r\<theta>.chine = \<langle>\<p>\<^sub>1[r0, r0 \<cdot> ?p1] \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1]\<rangle>"
using r\<theta>.chine_composite Cod_\<theta>_1 Dom_\<theta>_1 fw.leg1_composite C.comp_cod_arr
by (metis arrow_of_spans_data.simps(2) fw r.chine_eq_apex r.cod_simps(2)
rfw.prj_simps(10) rfw.prj_simps(16) span_data.simps(2))
have r\<theta>_cod_apex_eq: "r\<theta>.cod.apex = r0 \<down>\<down> ?u1"
using Cod_\<theta>_1 r\<theta>.chine_composite_in_hom by auto
hence r\<theta>'_cod_apex_eq: "r\<theta>'.cod.apex = r0 \<down>\<down> ?u1"
using Cod_\<theta>'_1 r\<theta>'.chine_composite_in_hom by auto
have Chn_r\<theta>': "\<guillemotleft>r\<theta>'.chine : rfw'.chine \<rightarrow>\<^sub>C ru.chine\<guillemotright>"
using r\<theta>'.chine_composite_in_hom ru.chine_composite rfw'.chine_composite
Cod_\<theta>'_1 Dom_\<theta>'_1 fw'.leg1_composite
by auto
have Chn_r\<theta>'_eq: "r\<theta>'.chine =
\<langle>\<p>\<^sub>1[r0, r0 \<cdot> ?p1'] \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1']\<rangle>"
using r\<theta>'.chine_composite Cod_\<theta>'_1 Dom_\<theta>'_1 fw'.leg1_composite C.comp_cod_arr
by (metis arrow_of_spans_data.simps(2) fw' r.chine_eq_apex r.cod_simps(2)
rfw'.prj_simps(10) rfw'.prj_simps(16) span_data.simps(2))
have Chn_\<rho>w: "\<guillemotleft>\<rho>w.chine : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C k0 \<down>\<down> ?w1\<guillemotright>"
using \<rho>w.chine_composite_in_hom by simp
have Chn_\<rho>w_eq: "\<rho>w.chine = \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using \<rho>w.chine_composite C.comp_cod_arr ide_w
by (simp add: chine_hcomp_arr_ide hcomp_def)
have Chn_\<rho>w': "\<guillemotleft>\<rho>w'.chine : ?R \<down>\<down> ?w1' \<rightarrow>\<^sub>C k0 \<down>\<down> ?w1'\<guillemotright>"
using \<rho>w'.chine_composite_in_hom by simp
have Chn_\<rho>w'_eq: "\<rho>w'.chine = \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> ?p0'\<rangle>"
using \<rho>w'.chine_composite C.comp_cod_arr ide_w' Dom_\<rho>_0 Cod_\<rho>_0
by (metis \<rho>w'.composite_is_arrow chine_hcomp_arr_ide chine_hcomp_def hseq_char
w'.cod_simps(3))
text \<open>
The following are some collected commutativity properties that are used
subsequently.
\<close>
have "C.commutative_square r0 ?u1 ?p1 \<theta>.chine"
using ru.legs_form_cospan(1) Dom_\<theta>.is_span Dom_\<theta>_1 Cod_\<theta>_1 \<theta>.leg1_commutes
apply (intro C.commutative_squareI) by auto
have "C.commutative_square r0 ?u1 (?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)"
proof
have 1: "r0 \<cdot> ?p1' = ?u1 \<cdot> \<theta>'.chine"
using \<theta>'.leg1_commutes Cod_\<theta>'_1 Dom_\<theta>'_1 fw'.leg1_composite by simp
show "C.cospan r0 ?u1"
using ru.legs_form_cospan(1) by blast
show "C.span (?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)"
using \<beta>.chine_in_hom \<theta>'.chine_in_hom
by (metis "1" C.dom_comp C.in_homE C.prj1_simps(1) C.prj1_simps(2)
C.seqI Cod_\<theta>'_1 Dom_\<theta>'.leg_simps(3) Chn_\<beta> \<theta>'.leg1_commutes cospan')
show "C.dom r0 = C.cod (?p1' \<cdot> \<beta>.chine)"
using \<beta>.chine_in_hom
by (metis C.cod_comp C.prj1_simps(3)
\<open>C.span (?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)\<close>
cospan' r.dom.apex_def r.chine_eq_apex r.chine_simps(2))
show "r0 \<cdot> ?p1' \<cdot> \<beta>.chine = ?u1 \<cdot> \<theta>'.chine \<cdot> \<beta>.chine"
using 1 \<beta>.chine_in_hom C.comp_assoc by metis
qed
have "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1] (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])"
using ru.legs_form_cospan(1) Dom_\<theta>.is_span Dom_\<theta>_1
C.comp_assoc C.pullback_commutes' r\<theta>.legs_form_cospan(1)
apply (intro C.commutative_squareI)
apply auto
by (metis C.comp_assoc Cod_\<theta>_1 \<theta>.leg1_commutes)
hence "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1] (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])"
using fw.leg1_composite by auto
have "C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1'] (\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])"
using C.tuple_is_extensional Chn_r\<theta>'_eq r\<theta>'.chine_simps(1) fw' by force
have "C.commutative_square ra ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0"
using C.pullback_commutes' gw.legs_form_cospan(1) rfw.prj_simps(2) C.comp_assoc
C.comp_cod_arr
apply (intro C.commutative_squareI) by auto
have "C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0"
using cospan'
apply (intro C.commutative_squareI)
apply simp_all
by (metis C.comp_assoc C.prj0_simps_arr C.pullback_commutes'
arrow_of_spans_data.select_convs(2) rfw'.prj_simps(3)
span_data.select_convs(1-2))
have "C.commutative_square r0 (r0 \<cdot> ?p1) rfw.Prj\<^sub>1\<^sub>1 \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>"
proof -
have "C.arr rfw.chine_assoc"
by (metis C.seqE rfw.prj_chine_assoc(1) rfw.prj_simps(1))
thus ?thesis
using C.tuple_is_extensional rfw.chine_assoc_def by fastforce
qed
have "C.commutative_square r0 (r0 \<cdot> ?p1') rfw'.Prj\<^sub>1\<^sub>1 \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>"
by (metis (no_types, lifting) C.not_arr_null C.seqE C.tuple_is_extensional
arrow_of_spans_data.select_convs(2) rfw'.chine_assoc_def
rfw'.prj_chine_assoc(1) rfw'.prj_simps(1) span_data.select_convs(1-2))
have "C.commutative_square k0 ?w1 (\<rho>.chine \<cdot> ?p1) ?p0"
using C.tuple_is_extensional Chn_\<rho>w_eq \<rho>w.chine_simps(1) by fastforce
have "C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')"
using C.tuple_is_extensional \<rho>w'.chine_composite \<rho>w'.chine_simps(1) by force
have "C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') ?p0'"
using C.tuple_is_extensional Chn_\<rho>w'_eq \<rho>w'.chine_simps(1) by force
text \<open>
Now, derive the consequences of the equation:
\[
\<open>(r \<star> \<theta>) \<bullet> \<a>[r, ?f, w] \<bullet> (?\<rho> \<star> w) = (r \<star> \<theta>') \<bullet> \<a>[r, ?f, w'] \<bullet> (?\<rho> \<star> w') \<bullet> \<beta>\<close>
\]
The strategy is to expand and simplify the left and right hand side to tuple form,
then compose with projections and equate corresponding components.
We first work on the right-hand side.
\<close>
have R: "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) =
\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
proof -
have "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) =
r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
proof -
have 1: "\<guillemotleft>T.composite_cell w' \<theta>' \<bullet> \<beta> : g \<star> w \<Rightarrow> r \<star> u\<guillemotright>"
using \<beta> \<rho>w' arfw' r\<theta>' by auto
have "Chn (T.composite_cell w' \<theta>' \<bullet> \<beta>) = Chn (T.composite_cell w' \<theta>') \<cdot> \<beta>.chine"
using 1 Chn_vcomp by blast
also have "... = (r\<theta>'.chine \<cdot> Chn (\<a>[r, f, w'] \<bullet> (\<rho> \<star> w'))) \<cdot> \<beta>.chine"
proof -
have "seq (r \<star> \<theta>') (\<a>[r, f, w'] \<bullet> (\<rho> \<star> w'))"
using 1 by blast
thus ?thesis
using 1 Chn_vcomp by presburger
qed
also have "... = (r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
proof -
have "seq \<a>[r, f, w'] (\<rho> \<star> w')"
using 1 by blast
thus ?thesis
using 1 Chn_vcomp by presburger
qed
finally show ?thesis
using C.comp_assoc by auto
qed
also have "... = \<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
proof -
let ?LHS = "r\<theta>'.chine \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
let ?RHS = "\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle>"
have LHS: "\<guillemotleft>?LHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
proof (intro C.comp_in_homI)
show "\<guillemotleft>\<beta>.chine : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C ?R \<down>\<down> ?w1'\<guillemotright>"
using Chn_\<beta> by simp
show "\<guillemotleft>\<rho>w'.chine : ?R \<down>\<down> ?w1' \<rightarrow>\<^sub>C Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1\<guillemotright>"
using Chn_\<rho>w' by simp
show "\<guillemotleft>Chn \<a>[r, f, w'] : Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1 \<rightarrow>\<^sub>C rfw'.chine\<guillemotright>"
using arfw'
by (metis (no_types, lifting) Chn_in_hom Cod_\<rho>_0
arrow_of_spans_data.simps(2) rf rf.leg0_composite rfw'.chine_composite(1)
span_data.select_convs(1) w'.cod_simps(3))
show "\<guillemotleft>r\<theta>'.chine : rfw'.chine \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
using Chn_r\<theta>' by auto
qed
have 2: "C.commutative_square r0 ?u1
(?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)"
by fact
have RHS: "\<guillemotleft>?RHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r\<theta>'.cod.apex\<guillemotright>"
using 2 Chn_\<beta> r\<theta>'_cod_apex_eq
C.tuple_in_hom [of r0 ?u1 "?p1' \<cdot> \<beta>.chine" "\<theta>'.chine \<cdot> \<beta>.chine"]
by fastforce
show ?thesis
proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS])
show "C.cospan r0 ?u1"
using ru.legs_form_cospan(1) by blast
show "C.seq ru.prj\<^sub>1 ?LHS"
using LHS r\<theta>'_cod_apex_eq by auto
show "C.seq ru.prj\<^sub>1 ?RHS"
using RHS r\<theta>'_cod_apex_eq by auto
show "ru.prj\<^sub>0 \<cdot> ?LHS = ru.prj\<^sub>0 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>0 \<cdot> ?LHS =
(ru.prj\<^sub>0 \<cdot> r\<theta>'.chine) \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = ((\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1']) \<cdot> Chn \<a>[r, f, w']) \<cdot>
\<rho>w'.chine \<cdot> \<beta>.chine"
using Chn_r\<theta>'_eq C.comp_assoc fw'
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1']
(\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])\<close>
by simp
also have "... = \<theta>'.chine \<cdot> (\<p>\<^sub>0[r0, r0 \<cdot> ?p1'] \<cdot> Chn \<a>[r, f, w']) \<cdot>
\<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = \<theta>'.chine \<cdot> (\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine) \<cdot>
\<beta>.chine"
using ide_f hseq_rf hseq_char \<alpha>_ide C.comp_assoc
rfw'.chine_assoc_def fw'.leg1_composite C.prj_tuple(1)
\<open>C.commutative_square r0 (r0 \<cdot> ?p1')
rfw'.Prj\<^sub>1\<^sub>1 \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>\<close>
by simp
also have "... = \<theta>'.chine \<cdot> \<beta>.chine"
proof -
have "\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine = gw'.apex"
proof (intro C.prj_joint_monic
[of ?R ?w1' "\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine"
gw'.apex])
show "C.cospan ?R ?w1'"
using fw'.legs_form_cospan(1) by simp
show "C.seq ?p1' (\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine)"
proof (intro C.seqI' C.comp_in_homI)
show "\<guillemotleft>\<rho>w'.chine : Dom_\<rho>.leg0 \<down>\<down> w'.leg1 \<rightarrow>\<^sub>C Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1\<guillemotright>"
using \<rho>w'.chine_composite_in_hom by simp
show "\<guillemotleft>\<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, w'.leg1\<rbrakk> rfw'.Prj\<^sub>0\<rangle> :
Cod_\<rho>.leg0 \<down>\<down> w'.cod.leg1 \<rightarrow>\<^sub>C ?R \<down>\<down> w'.leg1\<guillemotright>"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close>
C.tuple_in_hom [of ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0]
rf rf.leg0_composite
by auto
show "\<guillemotleft>?p1' : ?R \<down>\<down> w'.leg1 \<rightarrow>\<^sub>C f.apex\<guillemotright>"
using fw'.prj_in_hom(1) by auto
qed
show "C.seq ?p1' gw'.apex"
using gw'.dom.apex_def gw'.leg0_composite fw'.prj_in_hom by auto
show "?p0' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
?p0' \<cdot> gw'.apex"
proof -
have "?p0' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
(?p0' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w'.chine"
using C.comp_assoc by simp
also have "... = rfw'.Prj\<^sub>0 \<cdot> \<rho>w'.chine"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close> by auto
also have
"... = \<p>\<^sub>0[k0, ?w1'] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> w'.chine \<cdot> ?p0'\<rangle>"
using \<rho>w'.chine_composite Dom_\<rho>_0 Cod_\<rho>_0 C.comp_cod_arr by simp
also have "... = w'.chine \<cdot> ?p0'"
using \<open>C.commutative_square k0 ?w1'
(\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')\<close>
by simp
also have "... = ?p0' \<cdot> gw'.apex"
using cospan C.comp_cod_arr C.comp_arr_dom gw'.chine_is_identity
gw'.chine_eq_apex gw'.chine_composite fw'.prj_in_hom
by auto
finally show ?thesis by simp
qed
show "?p1' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
?p1' \<cdot> gw'.apex"
proof -
have "?p1' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle> \<cdot> \<rho>w'.chine =
(?p1' \<cdot> \<langle>rfw'.Prj\<^sub>0\<^sub>1 \<lbrakk>ra, ?w1'\<rbrakk> rfw'.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w'.chine"
using C.comp_assoc by simp
also have "... = rfw'.Prj\<^sub>0\<^sub>1 \<cdot> \<rho>w'.chine"
using \<open>C.commutative_square ?R ?w1' rfw'.Prj\<^sub>0\<^sub>1 rfw'.Prj\<^sub>0\<close> by simp
also have
"... = k0 \<cdot> \<p>\<^sub>1[k0, ?w1'] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1' \<lbrakk>k0, ?w1'\<rbrakk> w'.chine \<cdot> ?p0'\<rangle>"
using \<rho>w'.chine_composite Cod_\<rho>_0 C.comp_assoc C.comp_cod_arr
by simp
also have "... = k0 \<cdot> \<rho>.chine \<cdot> ?p1'"
using \<open>C.commutative_square k0 ?w1'
(\<rho>.chine \<cdot> ?p1') (w'.chine \<cdot> ?p0')\<close>
by simp
also have "... = (k0 \<cdot> \<rho>.chine) \<cdot> ?p1'"
using C.comp_assoc by metis
also have "... = ?p1'"
using \<rho>.leg0_commutes C.comp_cod_arr cospan' by simp
also have "... = ?p1' \<cdot> gw'.apex"
using C.comp_arr_dom cospan' gw'.chine_eq_apex gw'.chine_composite
by simp
finally show ?thesis by simp
qed
qed
thus ?thesis
using Chn_\<beta> C.comp_cod_arr gw'.apex_composite by auto
qed
also have "... = \<p>\<^sub>0[r0, ?u1] \<cdot> ?RHS"
using RHS 2 C.prj_tuple [of r0 ?u1] by simp
finally show ?thesis by simp
qed
show "ru.prj\<^sub>1 \<cdot> ?LHS = ru.prj\<^sub>1 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>1 \<cdot> ?LHS =
(ru.prj\<^sub>1 \<cdot> r\<theta>'.chine) \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = \<p>\<^sub>1[r0, fw'.leg1] \<cdot> Chn \<a>[r, f, w'] \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using Chn_r\<theta>' Chn_r\<theta>'_eq fw'
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1']
(\<theta>'.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1'])\<close>
by simp
also have "... = (rfw'.Prj\<^sub>1 \<cdot> rfw'.chine_assoc) \<cdot> \<rho>w'.chine \<cdot> \<beta>.chine"
using ide_f ide_w' hseq_rf hseq_char \<alpha>_ide fw'.leg1_composite C.comp_assoc
by auto
also have "... = (rfw'.Prj\<^sub>1\<^sub>1 \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using rfw'.prj_chine_assoc C.comp_assoc by simp
also have "... = ((k1 \<cdot> \<p>\<^sub>1[k0, ?w1']) \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using C.comp_cod_arr by simp
also have "... = (k1 \<cdot> \<p>\<^sub>1[k0, ?w1'] \<cdot> \<rho>w'.chine) \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = (k1 \<cdot> \<rho>.chine \<cdot> ?p1') \<cdot> \<beta>.chine"
using Chn_\<rho>w'_eq Dom_\<rho>_0 Cod_\<rho>_0
\<open>C.commutative_square k0 ?w1' (\<rho>.chine \<cdot> ?p1') ?p0'\<close>
by simp
also have "... = (k1 \<cdot> \<rho>.chine) \<cdot> ?p1' \<cdot> \<beta>.chine"
using C.comp_assoc by metis
also have "... = (?R \<cdot> ?p1') \<cdot> \<beta>.chine"
using C.comp_assoc by simp
also have "... = ?p1' \<cdot> \<beta>.chine"
using C.comp_cod_arr C.prj1_in_hom [of ?R ?w1'] cospan' by simp
also have "... = ru.prj\<^sub>1 \<cdot> ?RHS"
using RHS 2 by simp
finally show ?thesis by simp
qed
qed
qed
finally show ?thesis by simp
qed
text \<open>
Now we work on the left-hand side.
\<close>
have L: "Chn (T.composite_cell w \<theta>) = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
proof -
have "Chn (T.composite_cell w \<theta>) = r\<theta>.chine \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using Chn_vcomp arfw C.comp_assoc by auto
moreover have "... = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
proof -
let ?LHS = "r\<theta>.chine \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
let ?RHS = "\<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
have 2: "C.commutative_square r0 ?u1 ?p1 \<theta>.chine" by fact
have LHS: "\<guillemotleft>?LHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r0 \<down>\<down> ?u1\<guillemotright>"
using Chn_r\<theta> Chn_\<rho>w rfw.chine_assoc_in_hom
by (metis (no_types, lifting) "8" Chn_in_hom Dom_\<rho>_0
arrow_of_spans_data.simps(2) calculation gw.chine_composite
r\<theta>_cod_apex_eq ru.chine_composite)
have RHS: "\<guillemotleft>?RHS : ?R \<down>\<down> ?w1 \<rightarrow>\<^sub>C r0 \<down>\<down> ?u1\<guillemotright>"
using 2 C.tuple_in_hom [of r0 ?u1 "?p1" \<theta>.chine] cospan r\<theta>_cod_apex_eq
by simp
show ?thesis
proof (intro C.prj_joint_monic [of r0 ?u1 ?LHS ?RHS])
show "C.cospan r0 ?u1"
using ru.legs_form_cospan(1) by blast
show "C.seq ru.prj\<^sub>1 ?LHS"
using LHS r\<theta>_cod_apex_eq by auto
show "C.seq ru.prj\<^sub>1 ?RHS"
using RHS r\<theta>_cod_apex_eq by auto
show "ru.prj\<^sub>0 \<cdot> ?LHS = ru.prj\<^sub>0 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>0 \<cdot> ?LHS = (ru.prj\<^sub>0 \<cdot> r\<theta>.chine) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = (\<theta>.chine \<cdot> \<p>\<^sub>0[r0, f.leg1 \<cdot> fw.prj\<^sub>1]) \<cdot>
Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using Chn_r\<theta>_eq Dom_\<theta>_1 Cod_\<theta>_1 fw.leg1_composite
\<open>C.commutative_square r0 ?u1 \<p>\<^sub>1[r0, r0 \<cdot> ?p1]
(\<theta>.chine \<cdot> \<p>\<^sub>0[r0, r0 \<cdot> ?p1])\<close>
by simp
also have "... = \<theta>.chine \<cdot> (\<p>\<^sub>0[r0, r0 \<cdot> ?p1] \<cdot> Chn \<a>[r, f, w]) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = \<theta>.chine \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine"
proof -
have "Chn \<a>[r, f, w] = rfw.chine_assoc"
using ide_f ide_w hseq_rf hseq_char \<alpha>_ide by auto
moreover have "\<p>\<^sub>0[r0, r0 \<cdot> ?p1] \<cdot> rfw.chine_assoc =
\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>"
using rfw.chine_assoc_def
\<open>C.commutative_square r0 (r0 \<cdot> ?p1) rfw.Prj\<^sub>1\<^sub>1
\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>\<close>
by simp
ultimately show ?thesis by simp
qed
also have "... = \<theta>.chine \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine = ?R \<down>\<down> ?w1"
proof (intro C.prj_joint_monic
[of ?R ?w1 "\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine"
"?R \<down>\<down> ?w1"])
show "C.cospan ?R ?w1" by fact
show "C.seq ?p1 (\<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine)"
proof -
have "C.seq rfw.Prj\<^sub>0\<^sub>1 \<rho>w.chine"
by (meson C.seqI' Chn_in_hom \<rho>w rfw.prj_in_hom(2)
\<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close>)
thus ?thesis
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close>
by (metis (no_types) C.comp_assoc C.prj_tuple(2))
qed
show "C.seq ?p1 (?R \<down>\<down> ?w1)"
using gw.dom.apex_def gw.leg0_composite gw.prj_in_hom by auto
show "?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
?p0 \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
(?p0 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = rfw.Prj\<^sub>0 \<cdot> \<rho>w.chine"
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close> by simp
also have "... = \<p>\<^sub>0[k0, ?w1] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using Chn_\<rho>w_eq C.comp_cod_arr by simp
also have "... = ?p0"
using \<open>C.commutative_square k0 ?w1 (\<rho>.chine \<cdot> ?p1) ?p0\<close>
C.prj_tuple(1)
by blast
also have "... = ?p0 \<cdot> (?R \<down>\<down> ?w1)"
using C.comp_arr_dom gw.chine_eq_apex gw.chine_is_identity
by (metis C.arr_dom_iff_arr C.pbdom_def Dom_g gw.chine_composite
gw.chine_simps(1) span_data.select_convs(1))
finally show ?thesis by simp
qed
show "?p1 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
?p1 \<cdot> (?R \<down>\<down> ?w1)"
proof -
have "?p1 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle> \<cdot> \<rho>w.chine =
(?p1 \<cdot> \<langle>rfw.Prj\<^sub>0\<^sub>1 \<lbrakk>?R, ?w1\<rbrakk> rfw.Prj\<^sub>0\<rangle>) \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = rfw.Prj\<^sub>0\<^sub>1 \<cdot> \<rho>w.chine"
using \<open>C.commutative_square ?R ?w1 rfw.Prj\<^sub>0\<^sub>1 rfw.Prj\<^sub>0\<close> by simp
also have "... = (k0 \<cdot> \<p>\<^sub>1[k0, ?w1]) \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using Chn_\<rho>w_eq C.comp_cod_arr by simp
also have "... = k0 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<langle>\<rho>.chine \<cdot> ?p1 \<lbrakk>k0, ?w1\<rbrakk> ?p0\<rangle>"
using C.comp_assoc by simp
also have "... = k0 \<cdot> \<rho>.chine \<cdot> ?p1"
using \<open>C.commutative_square k0 ?w1 (\<rho>.chine \<cdot> ?p1) ?p0\<close> by simp
also have "... = (k0 \<cdot> \<rho>.chine) \<cdot> ?p1"
using C.comp_assoc by metis
also have "... = ?p1 \<cdot> (?R \<down>\<down> ?w1)"
using C.comp_arr_dom C.comp_cod_arr cospan by simp
finally show ?thesis by blast
qed
qed
thus ?thesis by simp
qed
also have "... = \<theta>.chine"
using C.comp_arr_dom \<theta>.chine_in_hom gw.chine_eq_apex gw.chine_is_identity
Dom_\<theta>_0 Cod_\<theta>_0 Dom_\<theta>.apex_def Cod_\<theta>.apex_def
by (metis Dom_g \<theta>.chine_simps(1) \<theta>.chine_simps(2) gw.chine_composite
gw.dom.apex_def gw.leg0_composite span_data.select_convs(1))
also have "... = ru.prj\<^sub>0 \<cdot> ?RHS"
using 2 by simp
finally show ?thesis by blast
qed
show "ru.prj\<^sub>1 \<cdot> ?LHS = ru.prj\<^sub>1 \<cdot> ?RHS"
proof -
have "ru.prj\<^sub>1 \<cdot> ?LHS = (ru.prj\<^sub>1 \<cdot> r\<theta>.chine) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
using C.comp_assoc by simp
also have "... = (r.chine \<cdot> \<p>\<^sub>1[r0, r0 \<cdot> ?p1]) \<cdot> Chn \<a>[r, f, w] \<cdot> \<rho>w.chine"
proof -
have "r\<theta>.chine \<noteq> C.null \<Longrightarrow>
\<p>\<^sub>1[r.cod.leg0, Cod_\<theta>.leg1] \<cdot> r\<theta>.chine =
r.chine \<cdot> \<p>\<^sub>1[r0, Dom_\<theta>.leg1]"
by (metis (lifting) C.prj_tuple(2) C.tuple_is_extensional r.cod_simps(2)
r\<theta>.chine_composite)
thus ?thesis
using Cod_\<theta>_1 Dom_\<theta>_1 r\<theta>.chine_simps(1) fw by fastforce
qed
also have "... = r.chine \<cdot> (rfw.Prj\<^sub>1 \<cdot> Chn \<a>[r, f, w]) \<cdot> \<rho>w.chine"
using C.comp_assoc fw.leg1_composite by simp
also have "... = r.chine \<cdot> rfw.Prj\<^sub>1\<^sub>1 \<cdot> \<rho>w.chine"
using ide_f ide_w hseq_rf hseq_char \<alpha>_ide
rfw.prj_chine_assoc(1)
by auto
also have "... = r.chine \<cdot> k1 \<cdot> \<p>\<^sub>1[k0, ?w1] \<cdot> \<rho>w.chine"
using C.comp_cod_arr C.comp_assoc by simp
also have "... = r.chine \<cdot> k1 \<cdot> \<rho>.chine \<cdot> \<p>\<^sub>1[Dom_\<rho>.leg0, ?w1]"
using Chn_\<rho>w_eq
\<open>C.commutative_square k0 ?w1
(\<rho>.chine \<cdot> \<p>\<^sub>1[ra, w.leg1]) \<p>\<^sub>0[ra, w.leg1]\<close>
by auto
also have "... = r.chine \<cdot> (k1 \<cdot> \<rho>.chine) \<cdot> ?p1"
using C.comp_assoc Dom_\<rho>_0 by metis
also have "... = r.chine \<cdot> ra \<cdot> ?p1"
by simp
also have "... = r.chine \<cdot> ?p1"
using C.comp_cod_arr
by (metis C.comp_assoc r.cod_simps(1) r.chine_eq_apex r.chine_simps(1)
r.chine_simps(3))
also have "... = ?p1"
using C.comp_cod_arr r.chine_eq_apex r.chine_is_identity
by (metis 2 C.commutative_squareE r.dom.apex_def)
also have "... = ru.prj\<^sub>1 \<cdot> ?RHS"
using 2 by simp
finally show ?thesis by simp
qed
qed
qed
ultimately show ?thesis
by simp
qed
text \<open>
This is the main point: the equation E boils down to the following:
\[
\<open>?p1' \<cdot> \<beta>.chine = ?p1 \<and> \<theta>'.chine \<cdot> \<beta>.chine = \<theta>.chine\<close>
\]
The first equation gets us close to what we need, but we still need
\<open>?p1 \<cdot> C.inv ?p0 = ?w1\<close>, which follows from the fact that ?p0 is the
pullback of ?R.
\<close>
have *: "\<langle>?p1' \<cdot> \<beta>.chine \<lbrakk>r0, ?u1\<rbrakk> \<theta>'.chine \<cdot> \<beta>.chine\<rangle> = \<langle>?p1 \<lbrakk>r0, ?u1\<rbrakk> \<theta>.chine\<rangle>"
using L R E by simp
have **: "?p1' \<cdot> \<beta>.chine = ?p1"
by (metis "*" C.in_homE C.not_arr_null C.prj_tuple(2) C.tuple_in_hom
C.tuple_is_extensional
\<open>C.commutative_square r0 u.leg1
(\<p>\<^sub>1[ra, w'.leg1] \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)\<close>)
have ***: "\<theta>'.chine \<cdot> \<beta>.chine = \<theta>.chine"
by (metis "*" C.prj_tuple(1) \<open>C.commutative_square r0 ?u1
(?p1' \<cdot> \<beta>.chine) (\<theta>'.chine \<cdot> \<beta>.chine)\<close>
\<open>C.commutative_square r0 ?u1 ?p1 \<theta>.chine\<close>)
text \<open>
CKS say to take \<open>\<gamma> = \<beta>\<close>, but obviously this cannot work as
literally described, because \<open>\<guillemotleft>\<beta> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>\<close>, whereas we must have
\<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>\<close>. Instead, we have to define \<open>\<gamma>\<close> by transporting \<open>\<beta>\<close> along the
projections from \<open>?R \<down>\<down> ?w1\<close> to \<open>?W\<close> and \<open>?R \<down>\<down> ?w1'\<close> to \<open>?W'\<close>.
These are isomorphisms by virtue of their being pullbacks of identities,
but they are not themselves necessarily identities.
Specifically, we take \<open>Chn \<gamma> = ?p0' \<cdot> Chn \<beta> \<cdot> C.inv ?p0\<close>.
\<close>
let ?\<gamma> = "\<lparr>Chn = ?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0, Dom = Dom w, Cod = Cod w'\<rparr>"
interpret Dom_\<gamma>: span_in_category C \<open>Dom ?\<gamma>\<close>
using w.dom.span_in_category_axioms by simp
interpret Cod_\<gamma>: span_in_category C \<open>Cod ?\<gamma>\<close>
using w'.cod.span_in_category_axioms by simp
text \<open>
It has to be shown that \<open>\<gamma>\<close> is an arrow of spans.
\<close>
interpret \<gamma>: arrow_of_spans C ?\<gamma>
proof
show "\<guillemotleft>Chn ?\<gamma> : Dom_\<gamma>.apex \<rightarrow>\<^sub>C Cod_\<gamma>.apex\<guillemotright>"
proof -
have "\<guillemotleft>Chn \<beta>: gw.apex \<rightarrow>\<^sub>C gw'.apex\<guillemotright>"
using Chn_in_hom \<beta> gw'.chine_eq_apex gw.chine_eq_apex by force
moreover have "\<guillemotleft>?p0' : gw'.apex \<rightarrow>\<^sub>C w'.apex\<guillemotright>"
using cospan' hseq_gw' hseq_char hcomp_def gw'.dom.apex_def w'.dom.apex_def
by auto
moreover have "\<guillemotleft>C.inv ?p0 : w.apex \<rightarrow>\<^sub>C gw.apex\<guillemotright>"
using cospan hseq_gw hseq_char hcomp_def gw.dom.apex_def w.dom.apex_def
C.iso_pullback_ide
by auto
ultimately show ?thesis
using Dom_\<gamma>.apex_def Cod_\<gamma>.apex_def by auto
qed
text \<open>
The commutativity property for the ``input leg'' follows directly from that
for \<open>\<beta>\<close>.
\<close>
show "Cod_\<gamma>.leg0 \<cdot> Chn ?\<gamma> = Dom_\<gamma>.leg0"
using C.comp_assoc C.comp_arr_dom cospan C.iso_pullback_ide C.comp_arr_inv'
by (metis C.invert_side_of_triangle(2) Dom_\<beta>.leg_simps(1) Dom_\<beta>_eq \<beta>0
arrow_of_spans_data.select_convs(1,3) arrow_of_spans_data.simps(2)
r.dom.ide_apex span_data.select_convs(1) w'.cod_simps(2))
text \<open>
The commutativity property for the ``output leg'' is a bit more subtle.
\<close>
show "Cod_\<gamma>.leg1 \<cdot> Chn ?\<gamma> = Dom_\<gamma>.leg1"
proof -
have "Cod_\<gamma>.leg1 \<cdot> Chn ?\<gamma> = ((?w1' \<cdot> ?p0') \<cdot> \<beta>.chine) \<cdot> C.inv ?p0"
using C.comp_assoc by simp
also have "... = ((?R \<cdot> ?p1') \<cdot> Chn \<beta>) \<cdot> C.inv ?p0"
using cospan' C.pullback_commutes [of ?R ?w1'] by auto
also have "... = (?p1' \<cdot> \<beta>.chine) \<cdot> C.inv ?p0"
using cospan' C.comp_cod_arr by simp
also have "... = ?p1 \<cdot> C.inv ?p0"
using ** by simp
also have "... = ?w1"
text \<open>
Sledgehammer found this at a time when I was still struggling to
understand what was going on.
\<close>
by (metis C.comp_cod_arr C.invert_side_of_triangle(2) C.iso_pullback_ide
C.prj1_simps(1,3) C.pullback_commutes' cospan r.dom.ide_apex
r.chine_eq_apex r.chine_simps(2))
also have "... = Dom_\<gamma>.leg1" by auto
finally show ?thesis by simp
qed
qed
text \<open>
What remains to be shown is that \<open>\<gamma>\<close> is unique with the properties asserted
by \<open>T2\<close>; \emph{i.e.} \<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>)\<close>.
CKS' assertion that the equation \<open>(r\<theta>)(\<rho>w) = (r\<theta>')(\<rho>w')\<beta>\<close> gives \<open>w\<^sub>1 = w\<^sub>1'\<close>
does not really seem to be true. The reason \<open>\<gamma>\<close> is unique is because it is
obtained by transporting \<open>\<beta>\<close> along isomorphisms.
\<close>
have \<gamma>: "\<guillemotleft>?\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using \<gamma>.arrow_of_spans_axioms arr_char dom_char cod_char by auto
have hseq_f\<gamma>: "hseq f ?\<gamma>"
using \<gamma> src_def trg_def arrI fw.composable rf.are_arrows(2) by auto
have hseq_g\<gamma>: "hseq g ?\<gamma>"
using \<gamma> src_def trg_def fw.composable gw.are_arrows(1) src_f by auto
interpret f\<gamma>: two_composable_arrows_of_spans C prj0 prj1 f ?\<gamma>
using hseq_f\<gamma> hseq_char by (unfold_locales, simp)
interpret f\<gamma>: arrow_of_spans C \<open>f \<star> ?\<gamma>\<close>
using f\<gamma>.composite_is_arrow arr_char by simp
interpret g\<gamma>: two_composable_arrows_of_spans C prj0 prj1 g ?\<gamma>
using hseq_g\<gamma> hseq_char by (unfold_locales, simp)
interpret g\<gamma>: arrow_of_spans C \<open>g \<star> ?\<gamma>\<close>
using g\<gamma>.composite_is_arrow arr_char by simp
have Chn_g\<gamma>: "Chn (g \<star> ?\<gamma>) = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> ?p0' \<cdot> \<beta>.chine\<rangle>"
proof -
have "Chn (g \<star> ?\<gamma>) = \<langle>?R \<cdot> ?p1 \<lbrakk>?R, ?w1'\<rbrakk> (?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0\<rangle>"
using g\<gamma>.chine_composite by simp
also have "... = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> (?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0\<rangle>"
using C.comp_cod_arr cospan by simp
also have "... = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> ?p0' \<cdot> \<beta>.chine\<rangle>"
proof -
have "(?p0' \<cdot> \<beta>.chine \<cdot> C.inv ?p0) \<cdot> ?p0 = ?p0' \<cdot> \<beta>.chine"
using C.comp_assoc C.iso_pullback_ide [of ?R ?w1] C.comp_inv_arr
C.comp_arr_dom Chn_\<beta>
by (metis C.comp_inv_arr' C.in_homE C.pbdom_def cospan r.dom.ide_apex)
thus ?thesis by simp
qed
ultimately show ?thesis by simp
qed
have Chn_\<beta>_eq: "\<beta>.chine = Chn (g \<star> ?\<gamma>)"
proof -
have "Chn (g \<star> ?\<gamma>) = \<langle>?p1 \<lbrakk>?R, ?w1'\<rbrakk> ?p0' \<cdot> Chn \<beta>\<rangle>"
using Chn_g\<gamma> by simp
also have "... = \<beta>.chine"
text \<open>Here was another score by sledgehammer while I was still trying
to understand it.\<close>
using ** C.prj_joint_monic
by (metis C.prj1_simps(1) C.tuple_prj cospan cospan')
finally show ?thesis by simp
qed
have \<beta>_eq_g\<gamma>: "\<beta> = g \<star> ?\<gamma>"
proof (intro arr_eqI)
show "par \<beta> (g \<star> ?\<gamma>)"
proof -
have "\<guillemotleft>g \<star> ?\<gamma> : g \<star> w \<Rightarrow> g \<star> w'\<guillemotright>"
using ide_g \<gamma> T.leg1_simps(3)
by (intro hcomp_in_vhom, auto)
thus ?thesis
using \<beta> by (elim in_homE, auto)
qed
show "\<beta>.chine = Chn (g \<star> ?\<gamma>)"
using Chn_\<beta>_eq by simp
qed
moreover have "\<theta> = \<theta>' \<bullet> (f \<star> ?\<gamma>)"
proof (intro arr_eqI)
have f\<gamma>: "\<guillemotleft>f \<star> ?\<gamma> : f \<star> w \<Rightarrow> f \<star> w'\<guillemotright>"
using \<gamma> ide_f by auto
show par: "par \<theta> (\<theta>' \<bullet> (f \<star> ?\<gamma>))"
using \<theta> \<theta>' f\<gamma> by (elim in_homE, auto)
show "\<theta>.chine = Chn (\<theta>' \<bullet> (f \<star> ?\<gamma>))"
using par "***" Chn_vcomp calculation f\<gamma>.chine_composite g\<gamma>.chine_composite
by auto
qed
ultimately show 2: "\<guillemotleft>?\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> ?\<gamma> \<and> \<theta> = \<theta>' \<bullet> (f \<star> ?\<gamma>)"
using \<gamma> by simp
show "\<And>\<gamma>'. \<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>') \<Longrightarrow> \<gamma>' = ?\<gamma>"
proof -
fix \<gamma>'
assume 1: "\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = g \<star> \<gamma>' \<and> \<theta> = \<theta>' \<bullet> (f \<star> \<gamma>')"
interpret \<gamma>': arrow_of_spans C \<gamma>'
using 1 arr_char by auto
have hseq_g\<gamma>': \<open>hseq g \<gamma>'\<close>
using 1 \<beta> by auto
interpret g\<gamma>': two_composable_arrows_of_spans C prj0 prj1 g \<gamma>'
using hseq_g\<gamma>' hseq_char by unfold_locales auto
interpret g\<gamma>': arrow_of_spans C \<open>g \<star> \<gamma>'\<close>
using g\<gamma>'.composite_is_arrow arr_char by simp
show "\<gamma>' = ?\<gamma>"
proof (intro arr_eqI)
show par: "par \<gamma>' ?\<gamma>"
using 1 \<gamma> by fastforce
show "\<gamma>'.chine = \<gamma>.chine"
proof -
have "C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)"
proof
show "C.cospan ?R ?w1'" by fact
show 3: "C.span (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)"
proof (intro conjI)
show "C.seq g.chine ?p1"
using cospan by auto
show "C.seq \<gamma>'.chine ?p0"
using cospan 2 par arrow_of_spans_data.simps(1)
dom_char in_homE w.chine_eq_apex
by auto
thus "C.dom (g.chine \<cdot> ?p1) = C.dom (\<gamma>'.chine \<cdot> ?p0)"
using g.chine_eq_apex cospan by simp
qed
show "C.dom ra = C.cod (g.chine \<cdot> ?p1)"
using cospan by auto
show "?R \<cdot> g.chine \<cdot> ?p1 = ?w1' \<cdot> \<gamma>'.chine \<cdot> ?p0"
proof -
have "?w1' \<cdot> \<gamma>'.chine \<cdot> ?p0 = (?w1' \<cdot> \<gamma>'.chine) \<cdot> ?p0"
using C.comp_assoc by simp
moreover have "... = ?w1 \<cdot> ?p0"
using 1 \<gamma>'.leg1_commutes dom_char cod_char by auto
also have "... = ?R \<cdot> ?p1"
using cospan C.pullback_commutes [of ra ?w1] by auto
also have "... = ?R \<cdot> g.chine \<cdot> ?p1"
using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def
by auto
finally show ?thesis by auto
qed
qed
have "C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)"
proof
show "C.cospan ?R ?w1'" by fact
show 3: "C.span (g.chine \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)"
using cospan \<gamma>.chine_in_hom by auto
show "C.dom ?R = C.cod (g.chine \<cdot> ?p1)"
using cospan by auto
show "?R \<cdot> g.chine \<cdot> ?p1 = ?w1' \<cdot> \<gamma>.chine \<cdot> ?p0"
proof -
have "?w1' \<cdot> \<gamma>.chine \<cdot> ?p0 = (?w1' \<cdot> \<gamma>.chine) \<cdot> ?p0"
using C.comp_assoc by simp
moreover have "... = ?w1 \<cdot> ?p0"
using 1 \<gamma>.leg1_commutes dom_char cod_char by auto
also have "... = ?R \<cdot> ?p1"
using cospan C.pullback_commutes [of ra ?w1] by auto
also have "... = ?R \<cdot> g.chine \<cdot> ?p1"
using 3 C.comp_cod_arr g.chine_is_identity g.chine_eq_apex g.dom.apex_def
by auto
finally show ?thesis by auto
qed
qed
have "\<gamma>'.chine \<cdot> ?p0 = \<gamma>.chine \<cdot> ?p0"
proof -
have "\<gamma>'.chine \<cdot> ?p0 = ?p0' \<cdot> g\<gamma>'.chine"
using 1 dom_char cod_char g\<gamma>'.chine_composite
\<open>C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>'.chine \<cdot> ?p0)\<close>
by auto
also have "... = ?p0' \<cdot> \<beta>.chine"
using 1 by simp
also have "... = ?p0' \<cdot> g\<gamma>.chine"
using Chn_\<beta>_eq by simp
also have "... = \<gamma>.chine \<cdot> ?p0"
using g\<gamma>.chine_composite
\<open>C.commutative_square ?R ?w1' (g.chine \<cdot> ?p1) (\<gamma>.chine \<cdot> ?p0)\<close>
by simp
finally show ?thesis by simp
qed
thus ?thesis
using C.iso_pullback_ide C.iso_is_retraction C.retraction_is_epi
C.epiE [of "?p0" \<gamma>'.chine \<gamma>.chine] cospan \<gamma>.chine_in_hom
\<gamma>'.chine_in_hom
by auto
qed
qed
qed
qed
qed
qed
qed
qed
end
context span_bicategory
begin
interpretation chosen_right_adjoints vcomp hcomp assoc unit src trg ..
notation some_right_adjoint ("_\<^sup>*" [1000] 1000) (* TODO: Why is this needed? *)
notation isomorphic (infix "\<cong>" 50)
text \<open>
\<open>Span(C)\<close> is a bicategory of spans.
\<close>
lemma is_bicategory_of_spans:
shows "bicategory_of_spans vcomp hcomp assoc unit src trg"
proof
text \<open>
Every 1-cell \<open>r\<close> is isomorphic to the composition of a map and the right adjoint
of a map. The proof is to obtain a tabulation of \<open>r\<close> as a span of maps \<open>(f, g)\<close>
and then observe that \<open>r\<close> is isomorphic to \<open>g \<star> f\<^sup>*\<close>.
\<close>
show "\<And>r. ide r \<Longrightarrow> \<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> r \<cong> g \<star> f\<^sup>*"
proof -
fix r
assume r: "ide r"
interpret r: identity_arrow_of_spans C r
using r ide_char' by auto
interpret r: identity_arrow_in_span_bicategory C prj0 prj1 r ..
have \<rho>: "tabulation (\<bullet>) (\<star>) assoc unit src trg r r.\<rho> r.f r.g \<and>
is_left_adjoint r.f \<and> is_left_adjoint r.g"
using r r.has_tabulation by blast
interpret \<rho>: tabulation vcomp hcomp assoc unit src trg r r.\<rho> r.f r.g
using \<rho> by fast
have 1: "r \<cong> r.g \<star> r.f\<^sup>*"
using \<rho> \<rho>.yields_isomorphic_representation' \<rho>.T0.is_map
left_adjoint_extends_to_adjoint_pair
isomorphic_def [of "r.g \<star> r.f\<^sup>*" r] isomorphic_symmetric
by auto
thus "\<exists>f g. is_left_adjoint f \<and> is_left_adjoint g \<and> r \<cong> g \<star> f\<^sup>*"
using \<rho> by blast
qed
text \<open>
Every span of maps extends to a tabulation.
\<close>
show "\<And>f g. \<lbrakk> is_left_adjoint f; is_left_adjoint g; src f = src g \<rbrakk> \<Longrightarrow>
\<exists>r \<rho>. tabulation (\<bullet>) (\<star>) assoc unit src trg r \<rho> f g"
proof -
text \<open>
The proof idea is as follows: Let maps \<open>f = (f\<^sub>1, f\<^sub>0)\<close> and \<open>g = (g\<^sub>1, g\<^sub>0)\<close> be given.
Let \<open>f' = (f\<^sub>1 \<cdot> C.inv f\<^sub>0, C.cod f\<^sub>0)\<close> and \<open>g' = (g\<^sub>1 \<cdot> C.inv g\<^sub>0, C.cod g\<^sub>0)\<close>;
then \<open>f'\<close> and \<open>g'\<close> are maps isomorphic to \<open>f\<close> and \<open>g\<close>, respectively.
By a previous result, \<open>f'\<close> and \<open>g'\<close> extend to a tabulation \<open>(f', \<tau>, g')\<close> of
\<open>r = (f\<^sub>1 \<cdot> C.inv f\<^sub>0, g\<^sub>1 \<cdot> C.inv g\<^sub>0)\<close>.
Compose with isomorphisms \<open>\<guillemotleft>\<phi> : f' \<Rightarrow> f\<guillemotright>\<close> and \<open>\<guillemotleft>\<psi> : g \<Rightarrow> g'\<guillemotright>\<close> to obtain
\<open>(f, (r \<star> \<phi>) \<cdot> \<tau> \<cdot> \<psi>, g)\<close> and show it must also be a tabulation.
\<close>
fix f g
assume f: "is_left_adjoint f"
assume g: "is_left_adjoint g"
assume fg: "src f = src g"
show "\<exists>r \<rho>. tabulation (\<bullet>) (\<star>) assoc unit src trg r \<rho> f g"
proof -
text \<open>We have to unpack the hypotheses to get information about f and g.\<close>
obtain f\<^sub>a \<eta>\<^sub>f \<epsilon>\<^sub>f
where ff\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta>\<^sub>f \<epsilon>\<^sub>f"
using f adjoint_pair_def by auto
interpret ff\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta>\<^sub>f \<epsilon>\<^sub>f
using ff\<^sub>a by simp
interpret f: arrow_of_spans C f
using ide_char [of f] by simp
interpret f: identity_arrow_of_spans C f
using ide_char [of f] by unfold_locales auto
obtain g\<^sub>a \<eta>\<^sub>g \<epsilon>\<^sub>g
where G: "adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \<eta>\<^sub>g \<epsilon>\<^sub>g"
using g adjoint_pair_def by auto
interpret gg\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg g g\<^sub>a \<eta>\<^sub>g \<epsilon>\<^sub>g
using G by simp
interpret g: arrow_of_spans C g
using ide_char [of g] by simp
interpret g: identity_arrow_of_spans C g
using ide_char [of g] by unfold_locales auto
let ?f' = "mkIde (C.cod f.leg0) (f.dom.leg1 \<cdot> C.inv f.leg0)"
have f': "ide ?f'"
proof -
have "C.span (C.cod f.leg0) (f.leg1 \<cdot> C.inv f.leg0)"
using f is_left_adjoint_char by auto
thus ?thesis
using ide_mkIde by blast
qed
interpret f': arrow_of_spans C ?f'
using f' ide_char by blast
interpret f': identity_arrow_of_spans C ?f'
using f' ide_char by unfold_locales auto
let ?g' = "mkIde (C.cod g.leg0) (g.dom.leg1 \<cdot> C.inv g.leg0)"
have g': "ide ?g'"
proof -
have "C.span (C.cod g.leg0) (g.leg1 \<cdot> C.inv g.leg0)"
using g is_left_adjoint_char by auto
thus ?thesis
using ide_mkIde by blast
qed
interpret g': arrow_of_spans C ?g'
using g' ide_char by blast
interpret g': identity_arrow_of_spans C ?g'
using g' ide_char by unfold_locales auto
let ?r = "mkIde (f'.leg1) (g'.leg1)"
have r: "ide ?r"
proof -
have "C.span (f'.leg1) (g'.leg1)"
using f g fg src_def is_left_adjoint_char by simp
thus ?thesis
using ide_mkIde by blast
qed
interpret r: arrow_of_spans C ?r
using r ide_char by blast
interpret r: identity_arrow_of_spans C ?r
using r ide_char by unfold_locales auto
interpret r: identity_arrow_in_span_bicategory C prj0 prj1 ?r ..
have "r.f = ?f'"
using f r.chine_eq_apex is_left_adjoint_char by auto
have "r.g = ?g'"
using f r.chine_eq_apex fg src_def is_left_adjoint_char by simp
interpret \<rho>: tabulation \<open>(\<bullet>)\<close> \<open>(\<star>)\<close> assoc unit src trg ?r r.\<rho> r.f r.g
using r.has_tabulation by simp
have \<rho>_eq: "r.\<rho> = \<lparr>Chn = \<langle>C.cod f.leg0 \<lbrakk>f'.leg1, f'.leg1\<rbrakk> C.cod f.leg0\<rangle>,
Dom = \<lparr>Leg0 = C.cod f.leg0, Leg1 = g'.leg1\<rparr>,
Cod = \<lparr>Leg0 = \<p>\<^sub>0[f'.leg1, f'.leg1],
Leg1 = g'.leg1 \<cdot> \<p>\<^sub>1[f'.leg1, f'.leg1]\<rparr>\<rparr>"
using \<open>r.f = ?f'\<close> by auto
text \<open>Obtain the isomorphism from \<open>f'\<close> to \<open>f\<close>.\<close>
let ?\<phi> = "\<lparr>Chn = C.inv f.leg0, Dom = Dom ?f', Cod = Dom f\<rparr>"
interpret Dom_\<phi>: span_in_category C
\<open>Dom \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr>\<close>
using f'.dom.span_in_category_axioms by simp
interpret Cod_\<phi>: span_in_category C
\<open>Cod \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr>\<close>
using f.dom.span_in_category_axioms by simp
interpret \<phi>: arrow_of_spans C ?\<phi>
proof
show "\<guillemotleft>Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> : Dom_\<phi>.apex \<rightarrow>\<^sub>C Cod_\<phi>.apex\<guillemotright>"
using f f.dom.apex_def f'.dom.apex_def is_left_adjoint_char by auto
show "Cod_\<phi>.leg0 \<cdot> Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> =
Dom_\<phi>.leg0"
using f f.dom.apex_def is_left_adjoint_char C.comp_arr_inv C.inv_is_inverse
by simp
show "Cod_\<phi>.leg1 \<cdot> Chn \<lparr>Chn = C.inv f.leg0,
Dom = Dom (mkIde f.dsrc (f.leg1 \<cdot> C.inv f.leg0)),
Cod = Dom f\<rparr> =
Dom_\<phi>.leg1"
by simp
qed
have \<phi>: "\<guillemotleft>?\<phi> : ?f' \<Rightarrow> f\<guillemotright> \<and> iso ?\<phi>"
using f is_left_adjoint_char iso_char arr_char dom_char cod_char
\<phi>.arrow_of_spans_axioms f'.dom.apex_def f.dom.apex_def
by auto
text \<open>
Obtain the isomorphism from \<open>g\<close> to \<open>g'\<close>.
Recall: \<open>g' = mkIde (C.cod g.leg0) (g.dom.leg1 \<cdot> C.inv g.leg0)\<close>.
The isomorphism is given by \<open>g.leg0\<close>.
\<close>
let ?\<psi> = "\<lparr>Chn = g.leg0, Dom = Dom g, Cod = Dom ?g'\<rparr>"
interpret Dom_\<psi>: span_in_category C
\<open>Dom \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr>\<close>
using g.dom.span_in_category_axioms by simp
interpret Cod_\<psi>: span_in_category C
\<open>Cod \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr>\<close>
using g'.dom.span_in_category_axioms by simp
interpret \<psi>: arrow_of_spans C ?\<psi>
proof
show "\<guillemotleft>Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> :
Dom_\<psi>.apex \<rightarrow>\<^sub>C Cod_\<psi>.apex\<guillemotright>"
using g g.dom.apex_def g'.dom.apex_def is_left_adjoint_char by auto
show "Cod_\<psi>.leg0 \<cdot> Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> =
Dom_\<psi>.leg0"
using C.comp_cod_arr by simp
show "Cod_\<psi>.leg1 \<cdot> Chn \<lparr>Chn = g.leg0,
Dom = Dom g,
Cod = Dom (mkIde g.dsrc (g.leg1 \<cdot> C.inv g.leg0))\<rparr> =
Dom_\<psi>.leg1"
using g g.dom.apex_def is_left_adjoint_char C.comp_inv_arr C.inv_is_inverse
C.comp_assoc C.comp_arr_dom
by simp
qed
have \<psi>: "\<guillemotleft>?\<psi> : g \<Rightarrow> ?g'\<guillemotright> \<and> iso ?\<psi>"
using g is_left_adjoint_char iso_char arr_char dom_char cod_char
\<psi>.arrow_of_spans_axioms g.dom.apex_def g'.dom.apex_def
by auto
have \<rho>\<psi>: "tabulation (\<bullet>) (\<star>) assoc unit src trg ?r (r.\<rho> \<bullet> ?\<psi>) r.f g"
- using \<psi> `r.g = ?g'` r.has_tabulation \<rho>.preserved_by_output_iso by simp
+ using \<psi> \<open>r.g = ?g'\<close> r.has_tabulation \<rho>.preserved_by_output_iso by simp
interpret \<tau>\<psi>: tabulation vcomp hcomp assoc unit src trg ?r \<open>r.\<rho> \<bullet> ?\<psi>\<close> r.f g
using \<rho>\<psi> by auto
have "tabulation (\<bullet>) (\<star>) assoc unit src trg ?r ((?r \<star> ?\<phi>) \<bullet> r.\<rho> \<bullet> ?\<psi>) f g"
- using \<phi> `r.f = ?f'` \<tau>\<psi>.preserved_by_input_iso [of ?\<phi> f] by argo
+ using \<phi> \<open>r.f = ?f'\<close> \<tau>\<psi>.preserved_by_input_iso [of ?\<phi> f] by argo
thus ?thesis by auto
qed
qed
text \<open>The sub-bicategory of maps is locally essentially discrete.\<close>
show "\<And>f f' \<mu> \<mu>'. \<lbrakk> is_left_adjoint f; is_left_adjoint f'; \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>; \<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright> \<rbrakk>
\<Longrightarrow> iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
proof -
fix f f' \<mu> \<mu>'
assume f: "is_left_adjoint f" and f': "is_left_adjoint f'"
assume \<mu>: "\<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright>" and \<mu>': "\<guillemotleft>\<mu>' : f \<Rightarrow> f'\<guillemotright>"
obtain f\<^sub>a \<eta> \<epsilon>
where f\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta> \<epsilon>"
using f adjoint_pair_def by auto
obtain f'\<^sub>a \<eta>' \<epsilon>'
where f'\<^sub>a: "adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \<eta>' \<epsilon>'"
using f' adjoint_pair_def adjunction_def by auto
interpret f\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f f\<^sub>a \<eta> \<epsilon>
using f\<^sub>a by simp
interpret f'\<^sub>a: adjunction_in_bicategory vcomp hcomp assoc unit src trg f' f'\<^sub>a \<eta>' \<epsilon>'
using f'\<^sub>a by simp
interpret f: identity_arrow_of_spans C f
using ide_char' [of f] by simp
interpret f': identity_arrow_of_spans C f'
using ide_char' [of f'] by simp
interpret \<mu>: arrow_of_spans C \<mu> using \<mu> arr_char by auto
interpret \<mu>': arrow_of_spans C \<mu>' using \<mu>' arr_char by auto
have 1: "C.iso f.leg0 \<and> C.iso f'.leg0"
using f f' is_left_adjoint_char by simp
have 2: "\<mu>.chine = C.inv f'.leg0 \<cdot> f.leg0"
using \<mu> 1 dom_char cod_char \<mu>.leg0_commutes C.invert_side_of_triangle by auto
moreover have "\<mu>'.chine = C.inv f'.leg0 \<cdot> f.leg0"
using \<mu>' 1 dom_char cod_char \<mu>'.leg0_commutes C.invert_side_of_triangle by auto
ultimately have 3: "\<mu>.chine = \<mu>'.chine" by simp
have "iso \<mu>"
using 1 2 \<mu> C.isos_compose dom_char cod_char iso_char arr_char by auto
hence "iso \<mu>'"
using 3 iso_char arr_char \<mu>'.arrow_of_spans_axioms by simp
moreover have "\<mu> = \<mu>'"
using 3 \<mu> \<mu>' dom_char cod_char by fastforce
ultimately show "iso \<mu> \<and> iso \<mu>' \<and> \<mu> = \<mu>'"
by simp
qed
qed
text \<open>
We can now prove the easier half of the main result (CKS Theorem 4):
If \<open>B\<close> is biequivalent to \<open>Span(C)\<close>, where \<open>C\<close> is a category with pullbacks,
then \<open>B\<close> is a bicategory of spans.
(Well, it is easier given that we have already done the work to show that the notion
``bicategory of spans'' is respected by equivalence of bicategories.)
\<close>
theorem equivalent_implies_bicategory_of_spans:
assumes "equivalent_bicategories vcomp hcomp assoc unit src trg V\<^sub>1 H\<^sub>1 \<a>\<^sub>1 \<i>\<^sub>1 src\<^sub>1 trg\<^sub>1"
shows "bicategory_of_spans V\<^sub>1 H\<^sub>1 \<a>\<^sub>1 \<i>\<^sub>1 src\<^sub>1 trg\<^sub>1"
using assms is_bicategory_of_spans bicategory_of_spans_respects_equivalence by blast
end
subsection "Properties of Bicategories of Spans"
text \<open>
We now develop consequences of the axioms for a bicategory of spans, in preparation for
proving the other half of the main result.
\<close>
context bicategory_of_spans
begin
notation isomorphic (infix "\<cong>" 50)
text \<open>
The following is a convenience version of \<open>BS2\<close> that gives us what we generally want:
given specified \<open>f, g\<close> obtain \<open>\<rho>\<close> that makes \<open>(f, \<rho>, g)\<close> a tabulation of \<open>g \<star> f\<^sup>*\<close>,
not a tabulation of some \<open>r\<close> isomorphic to \<open>g \<star> f\<^sup>*\<close>.
\<close>
lemma BS2':
assumes "is_left_adjoint f" and "is_left_adjoint g" and "src f = src g"
and "isomorphic (g \<star> f\<^sup>*) r"
shows "\<exists>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
proof -
have 1: "is_left_adjoint f \<and> is_left_adjoint g \<and> g \<star> f\<^sup>* \<cong> r"
using assms BS1 by simp
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : g \<star> f\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
using 1 isomorphic_def by blast
obtain r' \<rho>' where \<rho>': "tabulation V H \<a> \<i> src trg r' \<rho>' f g"
using assms 1 BS2 by blast
interpret \<rho>': tabulation V H \<a> \<i> src trg r' \<rho>' f g
using \<rho>' by simp
let ?\<psi> = "\<rho>'.T0.trnr\<^sub>\<epsilon> r' \<rho>'"
have \<psi>: "\<guillemotleft>?\<psi> : g \<star> f\<^sup>* \<Rightarrow> r'\<guillemotright> \<and> iso ?\<psi>"
using \<rho>'.yields_isomorphic_representation by blast
have "\<guillemotleft>\<phi> \<cdot> inv ?\<psi> : r' \<Rightarrow> r\<guillemotright> \<and> iso (\<phi> \<cdot> inv ?\<psi>)"
using \<phi> \<psi> isos_compose by blast
hence 3: "tabulation V H \<a> \<i> src trg r ((\<phi> \<cdot> inv ?\<psi> \<star> f) \<cdot> \<rho>') f g"
using \<rho>'.is_preserved_by_base_iso by blast
hence "\<exists>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"
by blast
thus ?thesis
using someI_ex [of "\<lambda>\<rho>. tabulation V H \<a> \<i> src trg r \<rho> f g"] by simp
qed
text \<open>
The following observation is made by CKS near the beginning of the proof of Theorem 4:
If \<open>w\<close> is an arbitrary 1-cell, and \<open>g\<close> and \<open>g \<star> w\<close> are maps, then \<open>w\<close> is in fact a map.
It is applied frequently.
\<close>
lemma BS4:
assumes "is_left_adjoint g" and "ide w" and "is_left_adjoint (g \<star> w)"
shows "is_left_adjoint w"
proof -
text \<open>
CKS say: ``by (i) there are maps \<open>m, n\<close> with \<open>w \<cong> nm\<^sup>*\<close>, so, by (ii), we have two
tabulations \<open>(1, \<rho>, gw)\<close>, \<open>(m, \<sigma>, gn)\<close> of \<open>gw\<close>; since tabulations are unique
up to equivalence, \<open>m\<close> is invertible and \<open>w \<cong> nm\<^sup>*\<close> is a map.''
\<close>
have ex_\<rho>: "\<exists>\<rho>. tabulation V H \<a> \<i> src trg (g \<star> w) \<rho> (src w) (g \<star> w)"
proof -
have "(g \<star> w) \<star> src w \<cong> g \<star> w"
by (metis assms(3) iso_runit ideD(1) isomorphic_def left_adjoint_is_ide
runit_in_hom(2) src_hcomp)
moreover have "isomorphic ((g \<star> w) \<star> (src w)\<^sup>*) (g \<star> w)"
proof -
have "(g \<star> w) \<star> src (g \<star> w) \<cong> g \<star> w"
using calculation isomorphic_implies_ide(2) by auto
moreover have "isomorphic (src (g \<star> w)) (src w)\<^sup>*"
proof -
interpret src_w: map_in_bicategory V H \<a> \<i> src trg \<open>src w\<close>
using assms obj_is_self_adjoint by unfold_locales auto
interpret src_w: adjunction_in_bicategory V H \<a> \<i> src trg
\<open>src w\<close> \<open>(src w)\<^sup>*\<close> src_w.\<eta> src_w.\<epsilon>
using src_w.is_map left_adjoint_extends_to_adjunction by simp
have "adjoint_pair (src w) (src w)"
using assms obj_is_self_adjoint by simp
moreover have "adjoint_pair (src w) (src w)\<^sup>*"
using adjoint_pair_def src_w.adjunction_in_bicategory_axioms by auto
ultimately have "src w \<cong> (src w)\<^sup>*"
using left_adjoint_determines_right_up_to_iso by simp
moreover have "src w = src (g \<star> w)"
using assms isomorphic_def hcomp_simps(1) left_adjoint_is_ide by simp
ultimately show ?thesis by simp
qed
moreover have "src (g \<star> w) = trg (src (g \<star> w))"
using assms left_adjoint_is_ide by simp
ultimately show ?thesis
using assms left_adjoint_is_ide isomorphic_transitive isomorphic_symmetric
hcomp_ide_isomorphic
by blast
qed
ultimately show ?thesis
using assms obj_is_self_adjoint
left_adjoint_is_ide BS2' [of "src w" "g \<star> w" "g \<star> w"]
by auto
qed
obtain \<rho> where \<rho>: "tabulation V H \<a> \<i> src trg (g \<star> w) \<rho> (src w) (g \<star> w)"
using ex_\<rho> by auto
obtain m n where mn: "is_left_adjoint m \<and> is_left_adjoint n \<and> isomorphic w (n \<star> m\<^sup>*)"
using assms BS1 [of w] by auto
have m\<^sub>a: "adjoint_pair m m\<^sup>* \<and> isomorphic w (n \<star> m\<^sup>*)"
using mn adjoint_pair_def left_adjoint_extends_to_adjoint_pair by blast
have ex_\<sigma>: "\<exists>\<sigma>. tabulation V H \<a> \<i> src trg (g \<star> w) \<sigma> m (g \<star> n)"
proof -
have "hseq n m\<^sup>*"
using mn isomorphic_implies_ide by auto
have "trg (n \<star> m\<^sup>*) = trg w"
using mn m\<^sub>a isomorphic_def
by (metis (no_types, lifting) arr_inv dom_inv in_homE trg_dom trg_inv)
hence "trg n = trg w"
using mn by (metis assms(2) ideD(1) trg.preserves_reflects_arr trg_hcomp)
hence "hseq g n"
using assms mn left_adjoint_is_ide ideD(1)
by (metis hseq_char)
have "hseq g w"
using assms left_adjoint_is_ide by simp
have "src m = src n"
- using mn m\<^sub>a `hseq n m\<^sup>*` adjoint_pair_antipar [of m "m\<^sup>*"] by fastforce
+ using mn m\<^sub>a \<open>hseq n m\<^sup>*\<close> adjoint_pair_antipar [of m "m\<^sup>*"] by fastforce
have "is_left_adjoint (g \<star> n)"
- using assms mn left_adjoints_compose `hseq g n` by blast
+ using assms mn left_adjoints_compose \<open>hseq g n\<close> by blast
moreover have "src m = src (g \<star> n)"
- using assms mn `hseq g n` `src m = src n` by simp
+ using assms mn \<open>hseq g n\<close> \<open>src m = src n\<close> by simp
moreover have "(g \<star> n) \<star> m\<^sup>* \<cong> g \<star> w"
proof -
have 1: "src g = trg (n \<star> m\<^sup>*)"
- using assms `trg (n \<star> m\<^sup>*) = trg w` `hseq g w` by fastforce
+ using assms \<open>trg (n \<star> m\<^sup>*) = trg w\<close> \<open>hseq g w\<close> by fastforce
hence "(g \<star> n) \<star> m\<^sup>* \<cong> g \<star> n \<star> m\<^sup>*"
- using assms mn m\<^sub>a assoc_in_hom iso_assoc `hseq g n` `hseq n m\<^sup>*`
+ using assms mn m\<^sub>a assoc_in_hom iso_assoc \<open>hseq g n\<close> \<open>hseq n m\<^sup>*\<close>
isomorphic_def left_adjoint_is_ide right_adjoint_is_ide
by (metis hseqE ideD(2) ideD(3))
also have "... \<cong> g \<star> w"
using assms 1 mn m\<^sub>a isomorphic_symmetric hcomp_ide_isomorphic left_adjoint_is_ide
by simp
finally show ?thesis
using isomorphic_transitive by blast
qed
ultimately show ?thesis
using assms mn m\<^sub>a BS2' by blast
qed
obtain \<sigma> where \<sigma>: "tabulation V H \<a> \<i> src trg (g \<star> w) \<sigma> m (g \<star> n)"
using ex_\<sigma> by auto
interpret \<rho>: tabulation V H \<a> \<i> src trg \<open>g \<star> w\<close> \<rho> \<open>src w\<close> \<open>g \<star> w\<close>
using \<rho> by auto
interpret \<sigma>: tabulation V H \<a> \<i> src trg \<open>g \<star> w\<close> \<sigma> m \<open>g \<star> n\<close>
using \<sigma> by auto
text \<open>
As usual, the sketch given by CKS seems more suggestive than it is a precise recipe.
We can obtain an equivalence map \<open>\<guillemotleft>e : src w \<rightarrow> src m\<guillemotright>\<close> and \<open>\<theta>\<close> such that
\<open>\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright>\<close>.
We can also obtain an equivalence map \<open>\<guillemotleft>e' : src m \<rightarrow> src w\<guillemotright>\<close> and \<open>\<theta>'\<close> such that
\<open>\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright>\<close>. If \<open>\<theta>'\<close> can be taken to be an isomorphism; then we have
\<open>e' \<cong> src w \<star> e' \<cong> m\<close>. Since \<open>e'\<close> is an equivalence, this shows \<open>m\<close> is an equivalence,
hence its right adjoint \<open>m\<^sup>*\<close> is also an equivalence and therefore a map.
But \<open>w = n \<star> m\<^sub>a\<close>, so this shows that \<open>w\<close> is a map.
Now, we may assume without loss of generality that \<open>e\<close> and \<open>e'\<close> are part of an
adjoint equivalence.
We have \<open>\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright>\<close>.
We may take the transpose of \<open>\<theta>\<close> to obtain \<open>\<guillemotleft>\<zeta> : m \<Rightarrow> src w \<star> e'\<guillemotright>\<close>;
then \<open>\<guillemotleft>\<theta>' \<cdot> \<zeta> : m \<Rightarrow> m\<guillemotright>\<close> and \<open>\<guillemotleft>\<zeta> \<cdot> \<theta>' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>\<close>.
Since \<open>m\<close> and \<open>src w \<star> e'\<close> are maps, by \<open>BS3\<close> it must be that \<open>\<zeta>\<close> and \<open>\<theta>'\<close> are inverses.
\<close>
text \<open>
{\bf Note:} CKS don't cite \<open>BS3\<close> here. I am not sure whether this result can be proved
without \<open>BS3\<close>. For example, I am interested in knowing whether it can still be
proved under the the assumption that 2-cells between maps are unique, but not
necessarily invertible, or maybe even in a more general situation. It looks like
the invertibility part of \<open>BS3\<close> is not used in the proof below.
\<close>
have 2: "\<exists>e e' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'.
equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi> \<and>
\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<star> n \<Rightarrow> (g \<star> w) \<star> e'\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = \<rho>.composite_cell e' \<theta>' \<cdot> \<nu> \<and>
\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright> \<and> \<guillemotleft>\<nu>' : g \<star> w \<Rightarrow> (g \<star> n) \<star> e\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = ((g \<star> w) \<star> \<theta>) \<cdot> \<a>[g \<star> w, m, e] \<cdot> (\<sigma> \<star> e) \<cdot> \<nu>'"
using \<rho> \<sigma>.apex_unique_up_to_equivalence [of \<rho> "src w" "g \<star> w"] comp_assoc
by metis
obtain e e' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where *: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi> \<and>
\<guillemotleft>\<theta>' : src w \<star> e' \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>\<nu> : g \<star> n \<Rightarrow> (g \<star> w) \<star> e'\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = \<rho>.composite_cell e' \<theta>' \<cdot> \<nu> \<and>
\<guillemotleft>\<theta> : m \<star> e \<Rightarrow> src w\<guillemotright> \<and> \<guillemotleft>\<nu>' : g \<star> w \<Rightarrow> (g \<star> n) \<star> e\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = \<sigma>.composite_cell e \<theta> \<cdot> \<nu>'"
using 2 comp_assoc by auto
interpret ee': equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e e' \<phi> \<psi>
using * by simp
have equiv_e: "equivalence_map e"
using ee'.equivalence_in_bicategory_axioms equivalence_map_def by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg e e' \<phi> \<psi>'"
using equivalence_refines_to_adjoint_equivalence [of e e' \<phi>]
ee'.unit_in_hom(2) ee'.unit_is_iso ee'.antipar equiv_e
by auto
interpret ee': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e e' \<phi> \<psi>'
using \<psi>' by simp
interpret e'e: adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg e' e \<open>inv \<psi>'\<close> \<open>inv \<phi>\<close>
using * ee'.dual_adjoint_equivalence by simp
have equiv_e': "equivalence_map e'"
using e'e.equivalence_in_bicategory_axioms equivalence_map_def by auto
have "hseq m e"
using * ide_dom [of \<theta>]
apply (elim conjE in_homE) by simp
have "hseq (src w) e'"
using * ide_dom [of \<theta>']
apply (elim conjE in_homE) by simp
have "e'e.trnr\<^sub>\<eta> m \<theta> \<in> hom m (src w \<star> e')"
proof -
have "src m = trg e"
- using `hseq m e` by auto
+ using \<open>hseq m e\<close> by auto
moreover have "src (src w) = trg e'"
- using `hseq (src w) e'` by auto
+ using \<open>hseq (src w) e'\<close> by auto
moreover have "ide m"
using mn left_adjoint_is_ide by simp
moreover have "ide (src w)"
using assms by simp
ultimately show ?thesis
using * e'e.adjoint_transpose_right(1) by blast
qed
hence 3: "\<guillemotleft>e'e.trnr\<^sub>\<eta> m \<theta> : m \<Rightarrow> src w \<star> e'\<guillemotright>"
by simp
hence "\<guillemotleft>\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta> : m \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>"
using * by auto
moreover have "\<guillemotleft>m : m \<Rightarrow> m\<guillemotright> \<and> \<guillemotleft>src w \<star> e' : src w \<star> e' \<Rightarrow> src w \<star> e'\<guillemotright>"
using mn 3 ide_cod [of "e'e.trnr\<^sub>\<eta> m \<theta>"] left_adjoint_is_ide by fastforce
moreover have 4: "is_left_adjoint (src w \<star> e')"
proof -
have "is_left_adjoint (src w)"
using assms obj_is_self_adjoint by simp
moreover have "is_left_adjoint e'"
using e'e.adjunction_in_bicategory_axioms adjoint_pair_def by auto
ultimately show ?thesis
- using left_adjoints_compose `hseq (src w) e'` by auto
+ using left_adjoints_compose \<open>hseq (src w) e'\<close> by auto
qed
ultimately have "\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta> = m \<and> e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>' = src w \<star> e'"
using mn BS3 [of m m "\<theta>' \<cdot> e'e.trnr\<^sub>\<eta> m \<theta>" m]
BS3 [of "src w \<star> e'" "src w \<star> e'" "e'e.trnr\<^sub>\<eta> m \<theta> \<cdot> \<theta>'" "src w \<star> e'"]
by auto
hence "inverse_arrows \<theta>' (e'e.trnr\<^sub>\<eta> m \<theta>)"
using mn 4 left_adjoint_is_ide inverse_arrows_def by simp
hence 5: "iso \<theta>'"
by auto
have "equivalence_map (src w \<star> e')"
- using assms obj_is_equivalence_map equiv_e' `hseq (src w) e'` equivalence_maps_compose
+ using assms obj_is_equivalence_map equiv_e' \<open>hseq (src w) e'\<close> equivalence_maps_compose
by auto
hence "equivalence_map m"
using * 5 equivalence_map_preserved_by_iso isomorphic_def by auto
hence "equivalence_map m\<^sup>*"
using mn m\<^sub>a right_adjoint_to_equivalence_is_equivalence by simp
hence "is_left_adjoint m\<^sup>*"
using equivalence_is_left_adjoint by simp
moreover have "hseq n m\<^sup>*"
using mn isomorphic_implies_ide by auto
ultimately have "is_left_adjoint (n \<star> m\<^sup>*)"
using mn left_adjoints_compose by blast
thus ?thesis
using mn left_adjoint_preserved_by_iso isomorphic_def isomorphic_symmetric
by metis
qed
end
subsection "Choosing Tabulations"
context bicategory_of_spans
begin
notation isomorphic (infix "\<cong>" 50)
notation iso_class ("\<lbrakk>_\<rbrakk>")
text \<open>
We will ultimately need to have chosen a specific tabulation for each 1-cell.
This has to be done carefully, to avoid unnecessary choices.
We start out by using \<open>BS1\<close> to choose a specific factorization of the form
\<open>r \<cong> tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<close> for each 1-cell \<open>r\<close>. This has to be done in such a way
that all elements of an isomorphism class are assigned the same factorization.
\<close>
abbreviation isomorphic_rep
where "isomorphic_rep r f g \<equiv> is_left_adjoint f \<and> is_left_adjoint g \<and> g \<star> f\<^sup>* \<cong> r"
definition tab\<^sub>0
where "tab\<^sub>0 r \<equiv> SOME f. \<exists>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"
definition tab\<^sub>1
where "tab\<^sub>1 r \<equiv> SOME g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) g"
definition rep
where "rep r \<equiv> SOME \<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
lemma rep_props:
assumes "ide r"
shows "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>" and "iso (rep r)"
and "r \<cong> iso_class_rep \<lbrakk>r\<rbrakk>"
and "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)"
and "tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<cong> r"
proof -
have 1: "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)"
proof -
have "\<exists>f g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"
using assms BS1 isomorphic_symmetric rep_iso_class isomorphic_transitive
by blast
hence "isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) (tab\<^sub>1 r)"
using assms tab\<^sub>0_def tab\<^sub>1_def
someI_ex [of "\<lambda>f. \<exists>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) f g"]
someI_ex [of "\<lambda>g. isomorphic_rep (iso_class_rep \<lbrakk>r\<rbrakk>) (tab\<^sub>0 r) g"]
by simp
thus ?thesis
using assms isomorphic_symmetric isomorphic_transitive rep_iso_class by blast
qed
hence "\<exists>\<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"
using isomorphic_def by blast
hence 2: "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso (rep r)"
using someI_ex [of "\<lambda>\<phi>. \<guillemotleft>\<phi> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright> \<and> iso \<phi>"] rep_def by auto
show "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
using 2 by simp
show "iso (rep r)"
using 2 by simp
show "r \<cong> iso_class_rep \<lbrakk>r\<rbrakk>"
using assms rep_iso_class isomorphic_symmetric by simp
thus "isomorphic_rep r (tab\<^sub>0 r) (tab\<^sub>1 r)"
using 1 isomorphic_transitive by blast
thus "tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<cong> r"
by simp
qed
lemma tab\<^sub>0_in_hom [intro]:
assumes "ide r"
shows "\<guillemotleft>tab\<^sub>0 r : src (tab\<^sub>0 r) \<rightarrow> src r\<guillemotright>"
and "\<guillemotleft>tab\<^sub>0 r : tab\<^sub>0 r \<Rightarrow> tab\<^sub>0 r\<guillemotright>"
proof -
show "\<guillemotleft>tab\<^sub>0 r : tab\<^sub>0 r \<Rightarrow> tab\<^sub>0 r\<guillemotright>"
using assms rep_props left_adjoint_is_ide by auto
have "trg (tab\<^sub>0 r) = src r"
using assms rep_props
by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(3)
right_adjoint_simps(2) src_hcomp)
thus "\<guillemotleft>tab\<^sub>0 r : src (tab\<^sub>0 r) \<rightarrow> src r\<guillemotright>"
using assms rep_props left_adjoint_is_ide
by (intro in_hhomI, auto)
qed
lemma tab\<^sub>0_simps [simp]:
assumes "ide r"
shows "ide (tab\<^sub>0 r)"
and "is_left_adjoint (tab\<^sub>0 r)"
and "trg (tab\<^sub>0 r) = src r"
and "dom (tab\<^sub>0 r) = tab\<^sub>0 r" and "cod (tab\<^sub>0 r) = tab\<^sub>0 r"
using assms tab\<^sub>0_in_hom rep_props ide_dom left_adjoint_is_ide by auto
lemma tab\<^sub>1_in_hom [intro]:
assumes "ide r"
shows "\<guillemotleft>tab\<^sub>1 r : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>tab\<^sub>1 r : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 r\<guillemotright>"
proof -
show "\<guillemotleft>tab\<^sub>1 r : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 r\<guillemotright>"
using assms rep_props left_adjoint_is_ide by auto
have "trg (tab\<^sub>1 r) = trg r"
using assms rep_props
by (metis ideD(1) isomorphic_implies_hpar(1) isomorphic_implies_hpar(4) trg_hcomp)
moreover have "src (tab\<^sub>0 r) = src (tab\<^sub>1 r)"
using assms rep_props by fastforce
ultimately show "\<guillemotleft>tab\<^sub>1 r : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
using assms rep_props left_adjoint_is_ide
by (intro in_hhomI, auto)
qed
lemma tab\<^sub>1_simps [simp]:
assumes "ide r"
shows "ide (tab\<^sub>1 r)"
and "is_left_adjoint (tab\<^sub>1 r)"
and "src (tab\<^sub>1 r) = src (tab\<^sub>0 r)" and "trg (tab\<^sub>1 r) = trg r"
and "dom (tab\<^sub>1 r) = tab\<^sub>1 r" and "cod (tab\<^sub>1 r) = tab\<^sub>1 r"
using assms tab\<^sub>1_in_hom rep_props ide_dom left_adjoint_is_ide by auto
lemma rep_in_hom [intro]:
assumes "ide r"
shows "\<guillemotleft>rep r : src r \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
proof -
show "\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>"
using assms rep_props by auto
thus "\<guillemotleft>rep r : src r \<rightarrow> trg r\<guillemotright>"
using src_cod trg_cod by fastforce
qed
lemma rep_simps [simp]:
assumes "ide r"
shows "arr (rep r)"
and "src (rep r) = src r" and "trg (rep r) = trg r"
and "dom (rep r) = tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*" and "cod (rep r) = r"
using assms rep_in_hom by auto
lemma iso_rep:
assumes "ide r"
shows "iso (rep r)"
using assms rep_props by simp
end
text \<open>
Next, we assign a specific tabulation to each 1-cell r.
We can't just do this any old way if we ultimately expect to obtain a mapping that is
functorial with respect to vertical composition. What we have to do is to assign the
representative \<open>tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<close> its canonical tabulation, obtained as the adjoint
transpose of the identity, and then translate this to a tabulation of \<open>r\<close> via the chosen
isomorphism \<open>\<guillemotleft>rep r : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>.
\<close>
locale identity_in_bicategory_of_spans =
bicategory_of_spans +
fixes r :: 'a
assumes is_ide: "ide r"
begin
interpretation tab\<^sub>0: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>0 r\<close>
using is_ide rep_props by unfold_locales auto
interpretation tab\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>1 r\<close>
using is_ide rep_props by unfold_locales auto
text \<open>
A tabulation \<open>(tab\<^sub>0 r, tab, tab\<^sub>1 r)\<close> of \<open>r\<close> can be obtained as the adjoint transpose
of the isomorphism \<open>\<guillemotleft>rep r : (tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>. It is essential to define
it this way if we expect the mapping from 2-cells of the underlying bicategory
to arrows of spans to preserve vertical composition.
\<close>
definition tab
where "tab \<equiv> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) (rep r)"
text \<open>
In view of \<open>BS2'\<close>, the 1-cell \<open>(tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*\<close> has the canonical tabulation
obtained via adjoint transpose of an identity. In fact, this tabulation generates the
chosen tabulation of \<open>r\<close> in the same isomorphism class by translation along the
isomorphism \<open>\<guillemotleft>rep r : (tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> r\<guillemotright>\<close>. This fact is used to show that the
mapping from 2-cells to arrows of spans preserves identities.
\<close>
lemma canonical_tabulation:
shows "tabulation V H \<a> \<i> src trg
((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) (tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)) (tab\<^sub>0 r) (tab\<^sub>1 r)"
proof -
have "\<exists>\<rho>. tabulation V H \<a> \<i> src trg ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<rho> (tab\<^sub>0 r) (tab\<^sub>1 r)"
by (simp add: bicategory_of_spans.BS2' bicategory_of_spans_axioms is_ide
isomorphic_reflexive)
thus ?thesis
using is_ide tab\<^sub>0.canonical_tabulation by simp
qed
lemma tab_def_alt:
shows "tab = (rep r \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
and "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) (rep r \<cdot> ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*))"
using tab_def is_ide rep_in_hom [of r] comp_arr_dom by auto
also have "... = (rep r \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
using is_ide tab\<^sub>0.trnr\<^sub>\<eta>_comp by auto
finally show 1: "tab = (rep r \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)" by simp
have "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> tab =
((inv (rep r) \<star> tab\<^sub>0 r) \<cdot> (rep r \<star> tab\<^sub>0 r)) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
unfolding 1 using comp_assoc by presburger
also have "... = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have 1: "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> (rep r \<star> tab\<^sub>0 r) = ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<star> tab\<^sub>0 r"
using whisker_right [of "tab\<^sub>0 r" "inv (rep r)" "rep r"] iso_rep rep_in_hom
inv_is_inverse comp_inv_arr
by (simp add: comp_inv_arr' is_ide)
show ?thesis
proof -
have "\<guillemotleft>tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) :
tab\<^sub>1 r \<Rightarrow> (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) \<star> tab\<^sub>0 r\<guillemotright>"
by (meson canonical_tabulation tabulation_data.tab_in_hom(2) tabulation_def)
hence "((tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) \<star> tab\<^sub>0 r) \<cdot> tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*) =
tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
using 1 comp_cod_arr by blast
thus ?thesis
using 1 by simp
qed
qed
finally show "(inv (rep r) \<star> tab\<^sub>0 r) \<cdot> tab = tab\<^sub>0.trnr\<^sub>\<eta> (tab\<^sub>1 r) ((tab\<^sub>1 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
by blast
qed
lemma tab_is_tabulation:
shows "tabulation V H \<a> \<i> src trg r tab (tab\<^sub>0 r) (tab\<^sub>1 r)"
by (metis bicategory_of_spans.iso_rep bicategory_of_spans.rep_in_hom(2)
bicategory_of_spans_axioms is_ide canonical_tabulation tab_def_alt(1)
tabulation.is_preserved_by_base_iso)
(*
* TODO: If I pull the interpretation "tab" out of the following, Isabelle warns that
* the lemma is a redundant introduction rule and is being "ignored" for that purpose.
* However, the redundancy is only in the present context: if the enclosing locale is
* interpreted elsewhere, then the rule is not redundant. In order to make sure that
* the rule is not "ignored", I have put the interpretation "tab" into the proof to
* avoid the warning.
*)
lemma tab_in_hom [intro]:
shows "\<guillemotleft>tab : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
and "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
proof -
interpret tab: tabulation V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
using tab_is_tabulation by simp
show "\<guillemotleft>tab : src (tab\<^sub>0 r) \<rightarrow> trg r\<guillemotright>"
using tab.tab_in_hom by auto
show "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
using tab.tab_in_hom by auto
qed
lemma tab_simps [simp]:
shows "arr tab"
and "src tab = src (tab\<^sub>0 r)" and "trg tab = trg r"
and "dom tab = tab\<^sub>1 r" and "cod tab = r \<star> tab\<^sub>0 r"
using tab_in_hom by auto
end
text \<open>
The following makes the chosen tabulation conveniently available whenever we are
considering a particular 1-cell.
\<close>
sublocale identity_in_bicategory_of_spans \<subseteq> tabulation V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
using is_ide tab_is_tabulation by simp
context identity_in_bicategory_of_spans
begin
interpretation tab\<^sub>0: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>0 r\<close>
using is_ide rep_props by unfold_locales auto
interpretation tab\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>1 r\<close>
using is_ide rep_props by unfold_locales auto
text \<open>
The fact that adjoint transpose is a bijection allows us to invert the definition
of \<open>tab\<close> in terms of \<open>rep\<close> to express rep in terms of tab.
\<close>
lemma rep_in_terms_of_tab:
shows "rep r = T0.trnr\<^sub>\<epsilon> r tab"
using is_ide T0.adjoint_transpose_right(3) [of r "tab\<^sub>1 r" "rep r"] tab_def
by fastforce
lemma isomorphic_implies_same_tab:
assumes "isomorphic r r'"
shows "tab\<^sub>0 r = tab\<^sub>0 r'" and "tab\<^sub>1 r = tab\<^sub>1 r'"
using assms tab\<^sub>0_def tab\<^sub>1_def iso_class_eqI by auto
text \<open>
``Every 1-cell has a tabulation as a span of maps.''
Has a nice simple ring to it, but maybe not so useful for us, since we generally
really need to know that the tabulation has a specific form.
\<close>
lemma has_tabulation:
shows "\<exists>\<rho> f g. is_left_adjoint f \<and> is_left_adjoint g \<and> tabulation V H \<a> \<i> src trg r \<rho> f g"
using is_ide tab_is_tabulation rep_props by blast
end
subsection "Tabulations in a Bicategory of Spans"
context bicategory_of_spans
begin
abbreviation tab_of_ide
where "tab_of_ide r \<equiv> identity_in_bicategory_of_spans.tab V H \<a> \<i> src trg r"
abbreviation prj\<^sub>0
where "prj\<^sub>0 h k \<equiv> tab\<^sub>0 (k\<^sup>* \<star> h)"
abbreviation prj\<^sub>1
where "prj\<^sub>1 h k \<equiv> tab\<^sub>1 (k\<^sup>* \<star> h)"
lemma prj_in_hom [intro]:
assumes "ide h" and "is_left_adjoint k" and "trg h = trg k"
shows "\<guillemotleft>prj\<^sub>0 h k : src (prj\<^sub>0 h k) \<rightarrow> src h\<guillemotright>"
and "\<guillemotleft>prj\<^sub>1 h k : src (prj\<^sub>0 h k) \<rightarrow> src k\<guillemotright>"
and "\<guillemotleft>prj\<^sub>0 h k : prj\<^sub>0 h k \<Rightarrow> prj\<^sub>0 h k\<guillemotright>"
and "\<guillemotleft>prj\<^sub>1 h k : prj\<^sub>1 h k \<Rightarrow> prj\<^sub>1 h k\<guillemotright>"
by (intro in_hhomI, auto simp add: assms(1-3))
lemma prj_simps [simp]:
assumes "ide h" and "is_left_adjoint k" and "trg h = trg k"
shows "trg (prj\<^sub>0 h k) = src h"
and "src (prj\<^sub>1 h k) = src (prj\<^sub>0 h k)" and "trg (prj\<^sub>1 h k) = src k"
and "dom (prj\<^sub>0 h k) = prj\<^sub>0 h k" and "cod (prj\<^sub>0 h k) = prj\<^sub>0 h k"
and "dom (prj\<^sub>1 h k) = prj\<^sub>1 h k" and "cod (prj\<^sub>1 h k) = prj\<^sub>1 h k"
and "is_left_adjoint (prj\<^sub>0 h k)" and "is_left_adjoint (prj\<^sub>1 h k)"
using assms prj_in_hom by auto
end
text \<open>
Many of the commutativity conditions that we would otherwise have to worry about
when working with tabulations in a bicategory of spans reduce to trivialities.
The following locales try to exploit this to make our life more manageable.
\<close>
locale span_of_maps =
bicategory_of_spans +
fixes leg\<^sub>0 :: 'a
and leg\<^sub>1 :: 'a
assumes leg0_is_map: "is_left_adjoint leg\<^sub>0"
and leg1_is_map : "is_left_adjoint leg\<^sub>1"
text \<open>
The purpose of the somewhat strange-looking assumptions in this locale is
to cater to the form of data that we obtain from \<open>T1\<close>. Under the assumption
that we are in a bicategory of spans and that the legs of \<open>r\<close> and \<open>s\<close> are maps,
the hypothesized 2-cells will be uniquely determined isomorphisms, and an
arrow of spans \<open>w\<close> from \<open>r\<close> to \<open>s\<close> will be a map. We want to prove this once and
for all under the weakest assumptions we can manage.
\<close>
locale arrow_of_spans_of_maps =
bicategory_of_spans V H \<a> \<i> src trg +
r: span_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 +
s: span_of_maps V H \<a> \<i> src trg s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and w :: 'a +
assumes is_ide: "ide w"
and leg0_lax: "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
and leg1_iso: "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
begin
notation isomorphic (infix "\<cong>" 50)
lemma composite_leg1_is_map:
shows "is_left_adjoint (s\<^sub>1 \<star> w)"
using r.leg1_is_map leg1_iso left_adjoint_preserved_by_iso' isomorphic_def
isomorphic_symmetric
by auto
lemma is_map:
shows "is_left_adjoint w"
using is_ide composite_leg1_is_map s.leg1_is_map BS4 [of s\<^sub>1 w] by auto
lemma hseq_leg\<^sub>0:
shows "hseq s\<^sub>0 w"
by (metis ideD(1) ide_dom in_homE leg0_lax)
lemma composite_with_leg0_is_map:
shows "is_left_adjoint (s\<^sub>0 \<star> w)"
using left_adjoints_compose is_map s.leg0_is_map hseq_leg\<^sub>0 by blast
lemma leg0_uniquely_isomorphic:
shows "s\<^sub>0 \<star> w \<cong> r\<^sub>0"
and "\<exists>!\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
proof -
show 1: "s\<^sub>0 \<star> w \<cong> r\<^sub>0"
using leg0_lax composite_with_leg0_is_map r.leg0_is_map BS3 [of "s\<^sub>0 \<star> w" r\<^sub>0]
isomorphic_def
by auto
have "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> iso \<theta>"
using 1 isomorphic_def by simp
moreover have "\<And>\<theta> \<theta>'. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<Longrightarrow> \<guillemotleft>\<theta>' : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<Longrightarrow> \<theta> = \<theta>'"
using BS3 r.leg0_is_map composite_with_leg0_is_map by blast
ultimately show "\<exists>!\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" by blast
qed
lemma leg1_uniquely_isomorphic:
shows "r\<^sub>1 \<cong> s\<^sub>1 \<star> w"
and "\<exists>!\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
proof -
show 1: "r\<^sub>1 \<cong> s\<^sub>1 \<star> w"
using leg1_iso isomorphic_def by auto
have "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
using leg1_iso isomorphic_def isomorphic_symmetric by simp
moreover have "\<And>\<nu> \<nu>'. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<Longrightarrow> \<guillemotleft>\<nu>' : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<Longrightarrow> \<nu> = \<nu>'"
using BS3 r.leg1_is_map composite_leg1_is_map by blast
ultimately show "\<exists>!\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" by blast
qed
definition the_\<theta>
where "the_\<theta> \<equiv> THE \<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
definition the_\<nu>
where "the_\<nu> \<equiv> THE \<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
lemma the_\<theta>_props:
shows "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" and "iso the_\<theta>"
proof -
show "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
unfolding the_\<theta>_def
using the1I2 [of "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>" "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"]
leg0_uniquely_isomorphic
by simp
thus "iso the_\<theta>"
using BS3 r.leg0_is_map composite_with_leg0_is_map by simp
qed
lemma the_\<theta>_in_hom [intro]:
shows "\<guillemotleft>the_\<theta> : src r\<^sub>0 \<rightarrow> trg r\<^sub>0\<guillemotright>"
and "\<guillemotleft>the_\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright>"
using the_\<theta>_props apply auto
by (metis cod_trg in_hhom_def in_homE isomorphic_implies_hpar(3) leg0_uniquely_isomorphic(1)
src_dom trg.preserves_cod)
lemma the_\<theta>_simps [simp]:
shows "arr the_\<theta>"
and "src the_\<theta> = src r\<^sub>0" and "trg the_\<theta> = trg r\<^sub>0"
and "dom the_\<theta> = s\<^sub>0 \<star> w" and "cod the_\<theta> = r\<^sub>0"
using the_\<theta>_in_hom by auto
lemma the_\<nu>_props:
shows "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" and "iso the_\<nu>"
proof -
show "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
unfolding the_\<nu>_def
using the1I2 [of "\<lambda>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>" "\<lambda>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"]
leg1_uniquely_isomorphic
by simp
thus "iso the_\<nu>"
using BS3 r.leg1_is_map composite_leg1_is_map by simp
qed
lemma the_\<nu>_in_hom [intro]:
shows "\<guillemotleft>the_\<nu> : src r\<^sub>1 \<rightarrow> trg r\<^sub>1\<guillemotright>"
and "\<guillemotleft>the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright>"
using the_\<nu>_props apply auto
by (metis in_hhom_def in_homE isomorphic_implies_hpar(3) leg1_uniquely_isomorphic(1)
src_cod trg_dom)
lemma the_\<nu>_simps [simp]:
shows "arr the_\<nu>"
and "src the_\<nu> = src r\<^sub>1" and "trg the_\<nu> = trg r\<^sub>1"
and "dom the_\<nu> = r\<^sub>1" and "cod the_\<nu> = s\<^sub>1 \<star> w"
using the_\<nu>_in_hom by auto
end
(*
* TODO: I could probably avoid repeating the declarations of the locale parameters
* if I was willing to accept them being given in their order of appearance.
*)
locale arrow_of_spans_of_maps_to_tabulation_data =
bicategory_of_spans V H \<a> \<i> src trg +
arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w +
\<sigma>: tabulation_data V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and w :: 'a
text \<open>
The following declaration allows us to inherit the rules and other facts defined in
locale @{locale uw\<theta>}. It is tedious to prove very much without these in place.
\<close>
sublocale arrow_of_spans_of_maps_to_tabulation_data \<subseteq> uw\<theta> V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 r\<^sub>0 w the_\<theta>
using \<sigma>.tab_in_hom is_ide the_\<theta>_props by unfold_locales auto
locale arrow_of_spans_of_maps_to_tabulation =
arrow_of_spans_of_maps_to_tabulation_data +
tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
locale tabulation_in_maps =
span_of_maps V H \<a> \<i> src trg s\<^sub>0 s\<^sub>1 +
tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
sublocale tabulation_in_maps \<subseteq> tabulation V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 ..
sublocale identity_in_bicategory_of_spans \<subseteq>
tabulation_in_maps V H \<a> \<i> src trg r tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
using is_ide rep_props by unfold_locales auto
locale cospan_of_maps_in_bicategory_of_spans =
bicategory_of_spans +
fixes h :: 'a
and k :: 'a
assumes h_is_map: "is_left_adjoint h"
and k_is_map: "is_left_adjoint k"
and cospan: "trg h = trg k"
begin
text \<open>
The following sublocale declaration is perhaps pushing the limits of sensibility,
but the purpose is, given a cospan of maps \<open>(h, k)\<close>, to obtain ready access to the
composite \<open>k\<^sup>* \<star> h\<close> and its chosen tabulation.
\<close>
sublocale identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close>
using h_is_map k_is_map cospan left_adjoint_is_ide
by unfold_locales auto
notation isomorphic (infix "\<cong>" 50)
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
interpretation h: map_in_bicategory V H \<a> \<i> src trg h
using h_is_map by unfold_locales auto
interpretation k: map_in_bicategory V H \<a> \<i> src trg k
using k_is_map by unfold_locales auto
text \<open>
Our goal here is to reformulate the biuniversal properties of the chosen tabulation
of \<open>k\<^sup>* \<star> h\<close> in terms of its transpose, which yields a 2-cell from \<open>k \<star> tab\<^sub>1 (k\<^sup>* \<star> h)\<close>
to \<open>h \<star> tab\<^sub>0 (k\<^sup>* \<star> h)\<close>. These results do not depend on \<open>BS3\<close>.
\<close>
abbreviation p\<^sub>0
where "p\<^sub>0 \<equiv> prj\<^sub>0 h k"
abbreviation p\<^sub>1
where "p\<^sub>1 \<equiv> prj\<^sub>1 h k"
lemma p\<^sub>0_in_hom [intro]:
shows "\<guillemotleft>p\<^sub>0 : src p\<^sub>0 \<rightarrow> src h\<guillemotright>"
by auto
lemma p\<^sub>1_in_hom [intro]:
shows "\<guillemotleft>p\<^sub>1 : src p\<^sub>0 \<rightarrow> src k\<guillemotright>"
using prj_in_hom cospan h.ide_left k_is_map by blast
lemma p\<^sub>0_simps [simp]:
shows "trg p\<^sub>0 = src h"
by simp
lemma p\<^sub>1_simps [simp]:
shows "trg p\<^sub>1 = src k"
using k.antipar(1) by auto
definition \<phi>
where "\<phi> \<equiv> k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab)"
lemma \<phi>_in_hom [intro]:
shows "\<guillemotleft>\<phi> : src p\<^sub>0 \<rightarrow> trg h\<guillemotright>"
and "\<guillemotleft>\<phi> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<phi> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
unfolding \<phi>_def
using k.antipar cospan k.adjoint_transpose_left(2) [of "h \<star> p\<^sub>0" "p\<^sub>1"]
by fastforce
show "\<guillemotleft>\<phi> : src p\<^sub>0 \<rightarrow> trg h\<guillemotright>"
using 1 k.antipar src_dom trg_cod by fastforce
qed
lemma \<phi>_simps [simp]:
shows "arr \<phi>"
and "src \<phi> = src p\<^sub>0" and "trg \<phi> = trg h"
and "dom \<phi> = k \<star> p\<^sub>1" and "cod \<phi> = h \<star> p\<^sub>0"
using \<phi>_in_hom by auto
lemma transpose_\<phi>:
shows "tab = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 \<phi>"
proof -
have "\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 \<phi> = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab"
unfolding \<phi>_def
using k.antipar cospan
k.adjoint_transpose_left(4)
[of "h \<star> p\<^sub>0" "p\<^sub>1" "\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> tab"]
by fastforce
also have "... = (\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot> tab"
using comp_assoc by presburger
also have "... = tab"
using k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp
finally show ?thesis by simp
qed
lemma transpose_triangle:
assumes "ide w"
and "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>" and "\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright>"
shows "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>) =
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
proof -
have u: "ide u"
using assms(2) by auto
have v: "ide v"
using assms(3) by auto
have 0: "src p\<^sub>0 = trg w"
by (metis assms(2) hseqE ideD(1) src.preserves_reflects_arr u vconn_implies_hpar(3))
have 1: "src h = trg u"
using assms(1-2) 0 trg_dom trg_cod vconn_implies_hpar(4) by auto
have 2: "src k = trg v"
using assms(1,3) 0 trg_dom trg_cod hseqI'
by (metis ideD(1) leg1_simps(2) leg1_simps(3) p\<^sub>1_simps trg_hcomp vconn_implies_hpar(4))
have 3: "src u = src v \<and> src u = src w"
using assms 0 k.antipar src_dom src_cod hseqI'
by (metis ideD(1) leg0_simps(2) leg1_simps(2) leg1_simps(3) src_hcomp
vconn_implies_hpar(3))
have 4: "src h = trg \<theta>"
using assms 1 k.antipar by auto
define \<chi> where "\<chi> = \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)"
have \<chi>: "\<guillemotleft>\<chi> : p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> p\<^sub>0 \<star> w\<guillemotright>"
unfolding \<chi>_def
using assms 0 k.antipar cospan by (intro comp_in_homI, auto)
have "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>) =
k.trnl\<^sub>\<epsilon> (h \<star> u) ((k\<^sup>* \<star> h \<star> \<theta>) \<cdot> \<chi> \<cdot> \<nu>)"
unfolding \<chi>_def
using assms 1 k.antipar cospan assoc_naturality [of "k\<^sup>*" h \<theta>] comp_assoc
by (metis "4" h.ide_left ide_char in_homE k.ide_right)
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k\<^sup>* \<star> h \<star> \<theta>) \<cdot> (k \<star> \<chi> \<cdot> \<nu>)"
proof -
have "ide (h \<star> u)"
using "1" u assms h.ide_left by blast
moreover have "seq (k\<^sup>* \<star> h \<star> \<theta>) (\<chi> \<cdot> \<nu>)"
using assms 1 k.antipar cospan \<chi> seqI'
apply (intro seqI)
apply auto
apply blast
proof -
have "dom (k\<^sup>* \<star> h \<star> \<theta>) = k\<^sup>* \<star> h \<star> p\<^sub>0 \<star> w"
using assms
by (metis "4" cospan hcomp_simps(2-3) h.ide_left hseqI' ide_char in_homE
k.antipar(2) k.ide_right)
also have "... = cod \<chi>"
using \<chi> by auto
finally show "dom (k\<^sup>* \<star> h \<star> \<theta>) = cod \<chi>" by simp
qed
moreover have "src k = trg (k\<^sup>* \<star> h \<star> \<theta>)"
using assms k.antipar cospan calculation(2) by auto
ultimately show ?thesis
using k.trnl\<^sub>\<epsilon>_comp by simp
qed
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k\<^sup>* \<star> h \<star> \<theta>) \<cdot> (k \<star> \<chi>) \<cdot> (k \<star> \<nu>)"
using assms u \<chi> whisker_left
by (metis k.ide_left seqI')
also have
"... = (\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>)) \<cdot> (k \<star> \<chi>) \<cdot> (k \<star> \<nu>)"
unfolding k.trnl\<^sub>\<epsilon>_def by simp
also have "... = (h \<star> \<theta>) \<cdot>
(\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<chi>)) \<cdot>
(k \<star> \<nu>)"
proof -
have "\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>) =
\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
using assms 4 k.antipar cospan assoc'_naturality [of k "k\<^sup>*" "h \<star> \<theta>"]
by fastforce
also have "... = \<l>[h \<star> u] \<cdot> ((k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>)) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
using comp_assoc by presburger
also have "... = (\<l>[h \<star> u] \<cdot> (trg k \<star> h \<star> \<theta>)) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
proof -
have "(k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) = k.\<epsilon> \<cdot> (k \<star> k\<^sup>*) \<star> (h \<star> u) \<cdot> (h \<star> \<theta>)"
using assms 1 k.antipar cospan interchange comp_arr_dom comp_cod_arr
by fastforce
also have "... = k.\<epsilon> \<star> h \<star> \<theta>"
using assms k.antipar cospan comp_arr_dom comp_cod_arr k.counit_in_hom
whisker_left
by (metis h.ide_left in_homE)
also have "... = (trg k \<star> h \<star> \<theta>) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)"
using assms 4 k.antipar cospan whisker_left comp_arr_dom comp_cod_arr
interchange [of "trg k" k.\<epsilon> "h \<star> \<theta>" "h \<star> p\<^sub>0 \<star> w"]
by auto
finally have "(k.\<epsilon> \<star> h \<star> u) \<cdot> ((k \<star> k\<^sup>*) \<star> h \<star> \<theta>) = (trg k \<star> h \<star> \<theta>) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
proof -
have "\<l>[h \<star> u] \<cdot> (trg k \<star> h \<star> \<theta>) = (h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w]"
using assms 1 4 k.antipar cospan lunit_naturality [of "h \<star> \<theta>"]
by (metis hcomp_simps(3-4) h.ide_left hseqI' ide_char in_homE trg_hcomp)
thus ?thesis
using comp_assoc by presburger
qed
finally have "\<l>[h \<star> u] \<cdot> (k.\<epsilon> \<star> h \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> u] \<cdot> (k \<star> k\<^sup>* \<star> h \<star> \<theta>) =
(h \<star> \<theta>) \<cdot> \<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w]"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
proof -
have "\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)) =
\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot>
(trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
proof -
have "\<l>[h \<star> p\<^sub>0 \<star> w] =
\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
proof -
have "\<a>[h, p\<^sub>0, w] \<cdot> \<l>[(h \<star> p\<^sub>0) \<star> w] \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) =
\<a>[h, p\<^sub>0, w] \<cdot> \<ll> ((h \<star> p\<^sub>0) \<star> w) \<cdot> (trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
using assms 0 k.antipar cospan comp_cod_arr \<ll>_ide_simp by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> \<ll> (\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
using assms 0 k.antipar cospan \<ll>.is_natural_2 [of "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w] \<cdot> \<ll> (h \<star> p\<^sub>0 \<star> w)"
using assms 0 k.antipar cospan \<ll>.is_natural_1 [of "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"] comp_assoc
by simp
also have "... = (\<a>[h, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<ll> (h \<star> p\<^sub>0 \<star> w)"
using comp_assoc by presburger
also have "... = \<ll> (h \<star> p\<^sub>0 \<star> w)"
using assms 0 k.antipar cospan comp_cod_arr comp_assoc_assoc' by simp
also have "... = \<l>[h \<star> p\<^sub>0 \<star> w]"
using assms 0 k.antipar cospan \<ll>_ide_simp by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[trg h, h \<star> p\<^sub>0, w] \<cdot>
((trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w)) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using assms 0 k.antipar cospan lunit_hcomp comp_assoc by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
(\<a>\<^sup>-\<^sup>1[trg h, h \<star> p\<^sub>0, w] \<cdot> (k.\<epsilon> \<star> (h \<star> p\<^sub>0) \<star> w)) \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
proof -
have "(trg h \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) =
(k.\<epsilon> \<star> (h \<star> p\<^sub>0) \<star> w) \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w])"
using assms 0 k.antipar cospan comp_arr_dom comp_cod_arr
interchange [of "trg h" k.\<epsilon> "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]" "h \<star> p\<^sub>0 \<star> w"]
interchange [of k.\<epsilon> "k \<star> k\<^sup>*" "(h \<star> p\<^sub>0) \<star> w" "\<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using assms 0 k.antipar cospan assoc'_naturality [of k.\<epsilon> "h \<star> p\<^sub>0" w] comp_assoc
by simp
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
(k \<star> tab \<star> w)"
proof -
have "k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) =
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
(k \<star> tab \<star> w)"
proof -
have "seq \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] (\<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w))"
using \<chi>_def assms 0 k.antipar cospan \<chi> by blast
thus ?thesis
using assms 0 k.antipar cospan whisker_left by auto
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> (h \<star> p\<^sub>0)) \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot>
((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot> \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot> (k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot>
\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w]) \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
proof -
have "k \<star> tab \<star> w = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> (k \<star> tab \<star> w)"
proof -
have "\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> (k \<star> tab \<star> w) =
(\<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w]) \<cdot> (k \<star> tab \<star> w)"
using comp_assoc by presburger
also have "... = (k \<star> ((k\<^sup>* \<star> h) \<star> p\<^sub>0) \<star> w) \<cdot> (k \<star> tab \<star> w)"
using assms k.antipar 0 comp_assoc_assoc' by simp
also have "... = k \<star> tab \<star> w"
using assms k.antipar 0 comp_cod_arr by simp
finally show ?thesis by simp
qed
also have "... = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
using assms 0 k.antipar cospan assoc'_naturality [of k tab w] by simp
finally have "k \<star> tab \<star> w = \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] \<cdot> ((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<l>[h \<star> p\<^sub>0] \<star> w) \<cdot>
((k.\<epsilon> \<star> h \<star> p\<^sub>0) \<star> w) \<cdot>
(\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w) \<cdot>
((k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot>
(k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot> \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] =
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w"
proof -
have "\<a>\<^sup>-\<^sup>1[k \<star> k\<^sup>*, h \<star> p\<^sub>0, w] \<cdot> ((k \<star> k\<^sup>*) \<star> \<a>\<^sup>-\<^sup>1[h, p\<^sub>0, w]) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w]) \<cdot>
(k \<star> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w]) \<cdot> \<a>[k, (k\<^sup>* \<star> h) \<star> p\<^sub>0, w] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> ((\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, (\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using assms 0 k.antipar cospan \<alpha>_def \<a>'_def by simp
also have "... = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>k\<^bold>\<rangle>, \<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>k\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>k\<^sup>*\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<rbrace>"
using assms 0 k.antipar cospan
by (intro E.eval_eqI, simp_all)
also have "... = \<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<star> w"
using assms 0 k.antipar cospan \<alpha>_def \<a>'_def by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[h, p\<^sub>0, w] \<cdot>
(\<l>[h \<star> p\<^sub>0] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0] \<cdot> (k \<star> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot>
(k \<star> tab) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
using assms 0 k.antipar cospan comp_assoc whisker_right by auto
also have "... = \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
unfolding \<phi>_def k.trnl\<^sub>\<epsilon>_def
using assms 0 k.antipar cospan comp_assoc whisker_left by simp
finally have "\<l>[h \<star> p\<^sub>0 \<star> w] \<cdot> (k.\<epsilon> \<star> h \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, k\<^sup>*, h \<star> p\<^sub>0 \<star> w] \<cdot>
(k \<star> \<a>[k\<^sup>*, h, p\<^sub>0 \<star> w] \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)) =
\<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
by blast
thus ?thesis
using \<chi>_def comp_assoc by simp
qed
finally show ?thesis by simp
qed
text \<open>
\<open>BS3\<close> implies that \<open>\<phi>\<close> is the unique 2-cell from \<open>k \<star> p\<^sub>1\<close> to \<open>h \<star> p\<^sub>0\<close> and is an isomorphism.
\<close>
lemma \<phi>_uniqueness:
shows "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright> \<Longrightarrow> \<mu> = \<phi>" and "iso \<phi>"
proof -
have 2: "is_left_adjoint (k \<star> p\<^sub>1)"
using k.antipar cospan left_adjoints_compose by (simp add: k_is_map)
have 3: "is_left_adjoint (h \<star> p\<^sub>0)"
using k.antipar cospan left_adjoints_compose by (simp add: h_is_map)
show "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright> \<Longrightarrow> \<mu> = \<phi>"
using \<phi>_in_hom 2 3 BS3 by simp
show "iso \<phi>"
using \<phi>_in_hom 2 3 BS3 by simp
qed
text \<open>
As a consequence, the chosen tabulation of \<open>k\<^sup>* \<star> h\<close> is the unique 2-cell from
\<open>p\<^sub>1\<close> to \<open>(k\<^sup>* \<star> h) \<star> p\<^sub>0\<close>, and therefore if we are given any such 2-cell we may
conclude it yields a tabulation of \<open>k\<^sup>* \<star> h\<close>.
\<close>
lemma tab_uniqueness:
assumes "\<guillemotleft>\<tau> : p\<^sub>1 \<Rightarrow> (k\<^sup>* \<star> h) \<star> p\<^sub>0\<guillemotright>"
shows "\<tau> = tab"
proof -
have "\<guillemotleft>k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>) : k \<star> p\<^sub>1 \<Rightarrow> h \<star> p\<^sub>0\<guillemotright>"
using assms k.antipar cospan k.adjoint_transpose_left(2) [of "h \<star> p\<^sub>0" "p\<^sub>1"]
assoc_in_hom
by force
hence "tab = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> k.trnl\<^sub>\<eta> p\<^sub>1 (k.trnl\<^sub>\<epsilon> (h \<star> p\<^sub>0) (\<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>))"
using transpose_\<phi> \<phi>_uniqueness(1) by auto
also have "... = \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0] \<cdot> \<tau>"
using assms k.antipar cospan k.adjoint_transpose_left(4) assoc_in_hom by auto
also have "... = (\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, p\<^sub>0] \<cdot> \<a>[k\<^sup>*, h, p\<^sub>0]) \<cdot> \<tau>"
using comp_assoc by presburger
also have "... = \<tau>"
using assms k.antipar cospan comp_cod_arr comp_assoc_assoc' by auto
finally show ?thesis by simp
qed
text \<open>
The following lemma reformulates the biuniversal property of the canonical tabulation
of \<open>k\<^sup>* \<star> h\<close> as a biuniversal property of \<open>\<phi>\<close>, regarded as a square that commutes up to
isomorphism.
\<close>
lemma \<phi>_biuniversal_prop:
assumes "ide u" and "ide v"
shows "\<And>\<mu>. \<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
and "\<And>w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w';
\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>;
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>) \<rbrakk>
\<Longrightarrow> \<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<beta>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright>"
have 1: "src h = trg u"
using assms \<mu> ide_cod
by (metis ide_def in_homE seq_if_composable)
have 2: "src k = trg v"
using assms \<mu> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
let ?\<omega> = "\<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u] \<cdot> k.trnl\<^sub>\<eta> v \<mu>"
have \<omega>: "\<guillemotleft>?\<omega> : v \<Rightarrow> (k\<^sup>* \<star> h) \<star> u\<guillemotright>"
using assms \<mu> 1 2 k.antipar cospan k.adjoint_transpose_left(1) [of "h \<star> u" v]
assoc_in_hom
by auto
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = ?\<omega>"
using assms \<omega> T1 [of u ?\<omega>] comp_assoc by (metis in_homE)
have 0: "src p\<^sub>0 = trg w"
using w\<theta>\<nu> ide_dom
by (metis hseqE ideD(1) in_homE)
have "\<mu> = k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>)"
proof -
have "\<mu> = k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ?\<omega>)"
proof -
have "k.trnl\<^sub>\<epsilon> (h \<star> u) (\<a>[k\<^sup>*, h, u] \<cdot> ?\<omega>) =
k.trnl\<^sub>\<epsilon> (h \<star> u) ((\<a>[k\<^sup>*, h, u] \<cdot> \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \<cdot> k.trnl\<^sub>\<eta> v \<mu>)"
using comp_assoc by presburger
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) (k.trnl\<^sub>\<eta> v \<mu>)"
proof -
have "(\<a>[k\<^sup>*, h, u] \<cdot> \<a>\<^sup>-\<^sup>1[k\<^sup>*, h, u]) \<cdot> k.trnl\<^sub>\<eta> v \<mu> = (k\<^sup>* \<star> h \<star> u) \<cdot> k.trnl\<^sub>\<eta> v \<mu>"
using comp_assoc_assoc'
by (simp add: "1" assms(1) cospan k.antipar(2))
also have "... = k.trnl\<^sub>\<eta> v \<mu>"
using "1" \<omega> assms(1) comp_ide_arr cospan k.antipar(2) by fastforce
finally show ?thesis
by simp
qed
also have "... = \<mu>"
using assms \<mu> k.antipar cospan 1 2 k.adjoint_transpose_left(3) by simp
finally show ?thesis by simp
qed
thus ?thesis using w\<theta>\<nu> by simp
qed
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>)"
using assms k.antipar cospan w\<theta>\<nu> transpose_triangle [of w \<theta> u \<nu>] by auto
finally have "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
by simp
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
using w\<theta>\<nu> by blast
next
fix w w' \<theta> \<theta>' \<beta>
assume w: "ide w"
assume w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>"
assume eq: "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>)"
have 0: "src p\<^sub>0 = trg w"
using \<theta> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
interpret uw\<theta>w'\<theta>': uw\<theta>w'\<theta>' V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close>
u w \<theta> w' \<theta>'
using w \<theta> w' \<theta>' apply (unfold_locales) by auto
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<beta>"
proof -
let ?LHS = "\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w)"
let ?RHS = "\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>') \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w'] \<cdot> (tab \<star> w') \<cdot> \<beta>"
have eq': "?LHS = ?RHS"
proof -
have "k.trnl\<^sub>\<epsilon> (h \<star> u) ?LHS =
k.trnl\<^sub>\<epsilon> (h \<star> u)
(\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> (p\<^sub>1 \<star> w))"
using assms 0 w \<theta> \<beta> k.antipar cospan comp_arr_dom
by (metis tab_simps(1) tab_simps(4) whisker_right)
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> p\<^sub>1 \<star> w)"
using assms w \<theta> \<beta> transpose_triangle
by (metis arr_dom ide_hcomp ide_in_hom(2) in_homE ide_leg1 not_arr_null
seq_if_composable)
also have "... = (h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
using assms 0 w k.antipar cospan comp_arr_dom by simp
also have "... = (h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> \<beta>)"
using eq by blast
also have "... = k.trnl\<^sub>\<epsilon> (h \<star> u) ?RHS"
using assms w' \<theta>' \<beta> transpose_triangle by simp
finally have 4: "k.trnl\<^sub>\<epsilon> (h \<star> u) ?LHS = k.trnl\<^sub>\<epsilon> (h \<star> u) ?RHS"
by simp
have "src k = trg (p\<^sub>1 \<star> w)"
using assms 0 w k.antipar cospan by simp
moreover have "src k\<^sup>* = trg (h \<star> u)"
using assms 0 w k.antipar cospan by simp
moreover have "ide (h \<star> u)"
using assms 0 w k.antipar cospan by simp
moreover have "ide (p\<^sub>1 \<star> w)"
using assms 0 w k.antipar cospan by simp
ultimately have "inj_on (k.trnl\<^sub>\<epsilon> (h \<star> u)) (hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u))"
using assms 0 w w' k.antipar cospan k.adjoint_transpose_left(6) bij_betw_imp_inj_on
by blast
moreover have "?LHS \<in> hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"
proof -
have "\<guillemotleft>\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>) \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w] \<cdot> (tab \<star> w) :
p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> u\<guillemotright>"
using k.antipar cospan
apply (intro comp_in_homI) by auto
thus ?thesis by simp
qed
moreover have "?RHS \<in> hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"
proof -
have "\<guillemotleft>\<a>[k\<^sup>*, h, u] \<cdot> ((k\<^sup>* \<star> h) \<star> \<theta>') \<cdot> \<a>[k\<^sup>* \<star> h, p\<^sub>0, w'] \<cdot>
(tab \<star> w') \<cdot> \<beta> : p\<^sub>1 \<star> w \<Rightarrow> k\<^sup>* \<star> h \<star> u\<guillemotright>"
using \<beta> k.antipar cospan
apply (intro comp_in_homI) by auto
thus ?thesis by blast
qed
ultimately show "?LHS = ?RHS"
using assms 4 k.antipar cospan bij_betw_imp_inj_on
inj_on_def [of "k.trnl\<^sub>\<epsilon> (h \<star> u)" "hom (p\<^sub>1 \<star> w) (k\<^sup>* \<star> h \<star> u)"]
by simp
qed
moreover have "seq \<a>[k\<^sup>*, h, u] (composite_cell w \<theta>)"
using assms k.antipar cospan tab_in_hom by fastforce
moreover have "seq \<a>[k\<^sup>*, h, u] (composite_cell w' \<theta>' \<cdot> \<beta>)"
using assms \<beta> k.antipar cospan tab_in_hom by fastforce
ultimately have "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta>"
using assms 0 w w' \<beta> k.antipar cospan iso_assoc iso_is_section section_is_mono
monoE [of "\<a>[k\<^sup>*, h, u]" "composite_cell w \<theta>" "composite_cell w' \<theta>' \<cdot> \<beta>"]
comp_assoc
by simp
thus ?thesis
using w w' \<theta> \<theta>' \<beta> eq' T2 [of w w' \<theta> u \<theta>' \<beta>] by metis
qed
qed
text \<open>
Using the uniqueness properties established for \<open>\<phi>\<close>, we obtain yet another reformulation
of the biuniversal property associated with the chosen tabulation of \<open>k\<^sup>* \<star> h\<close>,
this time as a kind of pseudo-pullback. We will use this to show that the
category of isomorphism classes of maps has pullbacks.
\<close>
lemma has_pseudo_pullback:
assumes "is_left_adjoint u" and "is_left_adjoint v" and "isomorphic (k \<star> v) (h \<star> u)"
shows "\<exists>w. is_left_adjoint w \<and> isomorphic (p\<^sub>0 \<star> w) u \<and> isomorphic v (p\<^sub>1 \<star> w)"
and "\<And>w w'. \<lbrakk> is_left_adjoint w; is_left_adjoint w';
p\<^sub>0 \<star> w \<cong> u; v \<cong> p\<^sub>1 \<star> w; p\<^sub>0 \<star> w' \<cong> u; v \<cong> p\<^sub>1 \<star> w' \<rbrakk> \<Longrightarrow> w \<cong> w'"
proof -
interpret u: map_in_bicategory V H \<a> \<i> src trg u
using assms(1) by unfold_locales auto
interpret v: map_in_bicategory V H \<a> \<i> src trg v
using assms(2) by unfold_locales auto
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : k \<star> v \<Rightarrow> h \<star> u\<guillemotright> \<and> iso \<mu>"
using assms(3) by auto
obtain w \<theta> \<nu> where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] \<cdot> (k \<star> \<nu>) = \<mu>"
using assms \<mu> \<phi>_biuniversal_prop(1) [of u v \<mu>] by auto
have "is_left_adjoint w \<and> isomorphic (p\<^sub>0 \<star> w) u \<and> isomorphic v (p\<^sub>1 \<star> w)"
proof (intro conjI)
show 1: "is_left_adjoint w"
using assms(2) w\<theta>\<nu> left_adjoint_preserved_by_iso' isomorphic_def BS4 leg1_is_map
by blast
show "v \<cong> p\<^sub>1 \<star> w"
using w\<theta>\<nu> isomorphic_def by blast
show "p\<^sub>0 \<star> w \<cong> u"
proof -
have "src p\<^sub>0 = trg w"
using w\<theta>\<nu> ide_dom
by (metis ideD(1) in_homE not_arr_null seq_if_composable)
hence "is_left_adjoint (p\<^sub>0 \<star> w)"
using 1 left_adjoints_compose by simp
thus ?thesis
using assms w\<theta>\<nu> 1 BS3 isomorphic_def by metis
qed
qed
thus "\<exists>w. is_left_adjoint w \<and> p\<^sub>0 \<star> w \<cong> u \<and> v \<cong> p\<^sub>1 \<star> w"
by blast
show "\<And>w w'. \<lbrakk> is_left_adjoint w; is_left_adjoint w';
p\<^sub>0 \<star> w \<cong> u; v \<cong> p\<^sub>1 \<star> w; p\<^sub>0 \<star> w' \<cong> u; v \<cong> p\<^sub>1 \<star> w' \<rbrakk> \<Longrightarrow> w \<cong> w'"
proof -
fix w w'
assume w: "is_left_adjoint w" and w': "is_left_adjoint w'"
assume 1: "p\<^sub>0 \<star> w \<cong> u"
assume 2: "v \<cong> p\<^sub>1 \<star> w"
assume 3: "p\<^sub>0 \<star> w' \<cong> u"
assume 4: "v \<cong> p\<^sub>1 \<star> w'"
obtain \<theta> where \<theta>: "\<guillemotleft>\<theta> : p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
using 1 by auto
obtain \<theta>' where \<theta>': "\<guillemotleft>\<theta>' : p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
using 3 by auto
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu>: v \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>"
using 2 by blast
obtain \<nu>' where \<nu>': "\<guillemotleft>\<nu>': v \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>'"
using 4 by blast
let ?\<beta> = "\<nu>' \<cdot> inv \<nu>"
have \<beta>: "\<guillemotleft>?\<beta> : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>"
using \<nu> \<nu>' by (elim conjE in_homE, auto)
interpret uw\<theta>: uw\<theta> V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close> u w \<theta>
using w \<theta> left_adjoint_is_ide
by unfold_locales auto
interpret uw'\<theta>': uw\<theta> V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close>
u w' \<theta>'
using w' \<theta>' left_adjoint_is_ide
by unfold_locales auto
interpret uw\<theta>w'\<theta>': uw\<theta>w'\<theta>' V H \<a> \<i> src trg \<open>k\<^sup>* \<star> h\<close> tab \<open>p\<^sub>0\<close> \<open>p\<^sub>1\<close> u w \<theta> w' \<theta>'
using w w' \<theta> \<theta>' left_adjoint_is_ide by unfold_locales
have "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w] =
(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot>
(k \<star> ?\<beta>)"
proof -
let ?LHS = "(h \<star> \<theta>) \<cdot> \<a>[h, p\<^sub>0, w] \<cdot> (\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w]"
let ?RHS = "(h \<star> \<theta>') \<cdot> \<a>[h, p\<^sub>0, w'] \<cdot> (\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[k, p\<^sub>1, w'] \<cdot> (k \<star> ?\<beta>)"
have "\<guillemotleft>?LHS : k \<star> p\<^sub>1 \<star> w \<Rightarrow> h \<star> u\<guillemotright>"
using w k.antipar by fastforce
moreover have "\<guillemotleft>?RHS : k \<star> p\<^sub>1 \<star> w \<Rightarrow> h \<star> u\<guillemotright>"
using w k.antipar \<beta> by fastforce
moreover have "is_left_adjoint (k \<star> p\<^sub>1 \<star> w)"
using w k.is_map left_adjoints_compose by simp
moreover have "is_left_adjoint (h \<star> u)"
using assms h.is_map left_adjoints_compose by auto
ultimately show "?LHS = ?RHS"
using BS3 by blast
qed
hence "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<theta> = \<theta>' \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = ?\<beta>"
using assms left_adjoint_is_ide w w' \<theta> \<theta>' \<nu> \<nu>' \<beta>
\<phi>_biuniversal_prop(2) [of u v w w' \<theta> \<theta>' ?\<beta>]
by presburger
thus "w \<cong> w'"
using w w' BS3 isomorphic_def by metis
qed
qed
end
subsubsection "Tabulations in Maps"
text \<open>
Here we focus our attention on the properties of tabulations in a bicategory of spans,
in the special case in which both legs are maps.
\<close>
context tabulation_in_maps
begin
text \<open>
The following are the conditions under which \<open>w\<close> is a 1-cell induced via \<open>T1\<close> by
a 2-cell \<open>\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>\<close>: \<open>w\<close> is an arrow of spans and \<open>\<omega>\<close> is obtained by
composing the tabulation \<open>\<sigma>\<close> with \<open>w\<close> and the isomorphisms that witness \<open>w\<close> being
an arrow of spans.
\<close>
abbreviation is_induced_by_cell
where "is_induced_by_cell w r\<^sub>0 \<omega> \<equiv>
arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<omega>) s\<^sub>0 s\<^sub>1 w \<and>
composite_cell w (arrow_of_spans_of_maps.the_\<theta> V H r\<^sub>0 s\<^sub>0 w) \<cdot>
(arrow_of_spans_of_maps.the_\<nu> V H (dom \<omega>) s\<^sub>1 w) = \<omega>"
lemma induced_map_unique:
assumes "is_induced_by_cell w r\<^sub>0 \<omega>" and "is_induced_by_cell w' r\<^sub>0 \<omega>"
shows "isomorphic w w'"
proof -
interpret w: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s\<^sub>0 s\<^sub>1 w
using assms(1) by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w
..
interpret w': arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s\<^sub>0 s\<^sub>1 w'
using assms(2) by auto
interpret w': arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<omega>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w'
..
let ?\<beta> = "w'.the_\<nu> \<cdot> inv w.the_\<nu>"
have \<beta>: "\<guillemotleft>?\<beta> : s\<^sub>1 \<star> w \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using w.the_\<nu>_props w'.the_\<nu>_props arr_iff_in_hom by auto
have 1: "composite_cell w w.the_\<theta> = composite_cell w' w'.the_\<theta> \<cdot> (w'.the_\<nu> \<cdot> inv w.the_\<nu>)"
proof -
have "composite_cell w' w'.the_\<theta> \<cdot> (w'.the_\<nu> \<cdot> inv w.the_\<nu>) =
((composite_cell w' w'.the_\<theta>) \<cdot> w'.the_\<nu>) \<cdot> inv w.the_\<nu>"
using comp_assoc by presburger
also have "... = \<omega> \<cdot> inv w.the_\<nu>"
using assms(2) comp_assoc by simp
also have "... = (composite_cell w w.the_\<theta> \<cdot> w.the_\<nu>) \<cdot> inv w.the_\<nu>"
using assms(1) comp_assoc by simp
also have "... = composite_cell w w.the_\<theta> \<cdot> w.the_\<nu> \<cdot> inv w.the_\<nu>"
using comp_assoc by presburger
also have "... = composite_cell w w.the_\<theta>"
proof -
have "w.the_\<nu> \<cdot> inv w.the_\<nu> = s\<^sub>1 \<star> w"
using w.the_\<nu>_props comp_arr_inv inv_is_inverse by auto
thus ?thesis
using composite_cell_in_hom w.ide_w w.the_\<theta>_props comp_arr_dom
by (metis composite_cell_in_hom in_homE w.w_in_hom(1))
qed
finally show ?thesis by auto
qed
have "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>"
using 1 \<beta> w.is_ide w'.is_ide w.the_\<theta>_props w'.the_\<theta>_props
T2 [of w w' w.the_\<theta> r\<^sub>0 w'.the_\<theta> ?\<beta>]
by blast
thus ?thesis
using BS3 w.is_map w'.is_map by blast
qed
text \<open>
The object src \<open>s\<^sub>0\<close> forming the apex of the tabulation satisfies the conditions for
being a map induced via \<open>T1\<close> by the 2-cell \<open>\<sigma>\<close> itself. This is ultimately required
for the map from 2-cells to arrows of spans to be functorial with respect to vertical
composition.
\<close>
lemma apex_is_induced_by_cell:
shows "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \<sigma>"
proof -
have 1: "arrow_of_spans_of_maps V H \<a> \<i> src trg s\<^sub>0 s\<^sub>1 s\<^sub>0 s\<^sub>1 (src s\<^sub>0)"
using iso_runit [of s\<^sub>0] iso_runit [of s\<^sub>1] tab_in_hom
apply unfold_locales
apply simp
using ide_leg0 isomorphic_def
apply blast
using ide_leg1 isomorphic_def leg1_simps(3) runit'_in_vhom [of s\<^sub>1 "src s\<^sub>0"] iso_runit'
by blast
interpret w: arrow_of_spans_of_maps V H \<a> \<i> src trg s\<^sub>0 \<open>dom \<sigma>\<close> s\<^sub>0 s\<^sub>1 \<open>src s\<^sub>0\<close>
using 1 tab_in_hom by simp
interpret w: arrow_of_spans_of_maps_to_tabulation
V H \<a> \<i> src trg s\<^sub>0 \<open>dom \<sigma>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 \<open>src s\<^sub>0\<close>
..
show "is_induced_by_cell (src s\<^sub>0) s\<^sub>0 \<sigma>"
proof (intro conjI)
show "arrow_of_spans_of_maps V H \<a> \<i> src trg s\<^sub>0 (dom \<sigma>) s\<^sub>0 s\<^sub>1 (src s\<^sub>0)"
using w.arrow_of_spans_of_maps_axioms by simp
show "composite_cell (src s\<^sub>0) w.the_\<theta> \<cdot> w.the_\<nu> = \<sigma>"
proof -
have \<theta>: "w.the_\<theta> = \<r>[s\<^sub>0]"
using iso_runit [of s\<^sub>0] w.leg0_uniquely_isomorphic w.the_\<theta>_props
the1_equality [of "\<lambda>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> src s\<^sub>0 \<Rightarrow> s\<^sub>0\<guillemotright> \<and> iso \<theta>"]
by auto
have \<nu>: "w.the_\<nu> = \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using iso_runit' w.leg1_uniquely_isomorphic w.the_\<nu>_props leg1_simps(3)
the1_equality [of "\<lambda>\<nu>. \<guillemotleft>\<nu> : s\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> src s\<^sub>0\<guillemotright> \<and> iso \<nu>"] tab_in_vhom'
by auto
have "composite_cell (src s\<^sub>0) \<r>[s\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1] = \<sigma>"
proof -
have "composite_cell (src s\<^sub>0) \<r>[s\<^sub>0] \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1] =
((s \<star> \<r>[s\<^sub>0]) \<cdot> \<a>[s, s\<^sub>0, src s\<^sub>0]) \<cdot> (\<sigma> \<star> src s\<^sub>0) \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using comp_assoc by presburger
also have "... = (\<r>[s \<star> s\<^sub>0] \<cdot> (\<sigma> \<star> src s\<^sub>0)) \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using runit_hcomp comp_assoc by simp
also have "... = \<sigma> \<cdot> \<r>[s\<^sub>1] \<cdot> \<r>\<^sup>-\<^sup>1[s\<^sub>1]"
using runit_naturality tab_in_hom
by (metis tab_simps(1) tab_simps(2) tab_simps(4) tab_simps(5) comp_assoc)
also have "... = \<sigma>"
using iso_runit tab_in_hom comp_arr_dom comp_arr_inv inv_is_inverse by simp
finally show ?thesis by simp
qed
thus ?thesis
using \<theta> \<nu> comp_assoc by simp
qed
qed
qed
end
subsubsection "Composing Tabulations"
text \<open>
Given tabulations \<open>(r\<^sub>0, \<rho>, r\<^sub>1)\<close> of \<open>r\<close> and \<open>(s\<^sub>0, \<sigma>, s\<^sub>1)\<close> of \<open>s\<close> in a bicategory of spans,
where \<open>(r\<^sub>0, r\<^sub>1)\<close> and \<open>(s\<^sub>0, s\<^sub>1)\<close> are spans of maps and 1-cells \<open>r\<close> and \<open>s\<close> are composable,
we can construct a 2-cell that yields a tabulation of \<open>r \<star> s\<close>.
The proof uses the fact that the 2-cell \<open>\<phi>\<close> associated with the cospan \<open>(r\<^sub>0, s\<^sub>1)\<close>
is an isomorphism, which we have proved above
(\<open>cospan_of_maps_in_bicategory_of_spans.\<phi>_uniqueness\<close>) using \<open>BS3\<close>.
However, this is the only use of \<open>BS3\<close> in the proof, and it seems plausible that it would be
possible to establish that \<open>\<phi>\<close> is an isomorphism in more general situations in which the
subbicategory of maps is not locally essentially discrete. Alternatively, more general
situations could be treated by adding the assertion that \<open>\<phi>\<close> is an isomorphism as part of
a weakening of \<open>BS3\<close>.
\<close>
locale composite_tabulation_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a +
assumes composable: "src r = trg s"
begin
text \<open>
Interpret \<open>(r\<^sub>0, s\<^sub>1)\<close> as a @{locale cospan_of_maps_in_bicategory_of_spans},
to obtain the isomorphism \<open>\<phi>\<close> in the central diamond, along with the assertion
that it is unique.
\<close>
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map composable by unfold_locales auto
text \<open>
We need access to simps, etc. in the preceding interpretation, yet trying to declare
it as a sublocale introduces too many conflicts at the moment.
As it confusing elsewhere to figure out exactly how, in other contexts, to express
the particular interpretation that is needed, to make things easier we include the
following lemma. Then we can just recall the lemma to find out how to express
the interpretation required in a given context.
\<close>
lemma r\<^sub>0s\<^sub>1_is_cospan:
shows "cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0"
..
text \<open>
The following define the projections associated with the natural tabulation of \<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>.
\<close>
abbreviation p\<^sub>0
where "p\<^sub>0 \<equiv> r\<^sub>0s\<^sub>1.p\<^sub>0"
abbreviation p\<^sub>1
where "p\<^sub>1 \<equiv> r\<^sub>0s\<^sub>1.p\<^sub>1"
text \<open>
$$
\xymatrix{
&& {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
\<close>
text \<open>
Next, we define the 2-cell that is the composite of the tabulation \<open>\<sigma>\<close>, the tabulation \<open>\<rho>\<close>,
and the central diamond that commutes up to unique isomorphism \<open>\<phi>\<close>.
\<close>
definition tab
where "tab \<equiv> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1)"
lemma tab_in_hom [intro]:
shows "\<guillemotleft>tab : r\<^sub>1 \<star> p\<^sub>1 \<Rightarrow> (r \<star> s) \<star> s\<^sub>0 \<star> p\<^sub>0\<guillemotright>"
using \<rho>.T0.antipar(1) r\<^sub>0s\<^sub>1.\<phi>_in_hom composable \<rho>.leg0_in_hom(1) \<sigma>.leg1_in_hom(1)
composable tab_def
by auto
interpretation tabulation_data V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using composable tab_in_hom
by unfold_locales auto
text \<open>
In the subsequent proof we will use coherence to shortcut a few of the calculations.
\<close>
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
text \<open>
The following is applied twice in the proof of property \<open>T2\<close> for the composite
tabulation. It's too long to repeat.
\<close>
lemma technical:
assumes "ide w"
and "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
and "w\<^sub>r = p\<^sub>1 \<star> w"
and "\<theta>\<^sub>r = (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
shows "\<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r = \<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
text \<open>
$$
\xymatrix{
&& X \ar[d]^{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}} \\
&& {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
\<close>
proof -
interpret uw\<theta>: uw\<theta> V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close> u w \<theta>
using assms(1-2) composable
by unfold_locales auto
show ?thesis
proof -
have "\<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
(\<a>[r, s, u] \<cdot> ((r \<star> s) \<star> \<theta>)) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (r \<star> s \<star> \<theta>) \<cdot> \<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using assoc_naturality [of r s \<theta>] composable comp_assoc by simp
also have "... = (r \<star> s \<star> \<theta>) \<cdot> \<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
((\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
\<rho>.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
unfolding tab_def
using comp_assoc by presburger
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<cdot> \<rho>.composite_cell p\<^sub>1 r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using composable \<rho>.T0.antipar(1) comp_assoc whisker_right by auto
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
(\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using composable \<rho>.T0.antipar(1) whisker_right tab_def tab_in_hom(2)
composable comp_assoc
by force
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
((\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]) \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using assoc'_naturality [of \<rho> p\<^sub>1 w] \<rho>.T0.antipar(1) r\<^sub>0s\<^sub>1.base_simps(2) comp_assoc
by auto
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w]) \<cdot> \<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
proof -
have "(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) =
\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using pentagon' \<rho>.T0.antipar(1) comp_assoc by simp
moreover have 1: "seq (\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w)(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]))"
using \<rho>.T0.antipar(1)
by (intro seqI hseqI, auto)
ultimately
have "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
((\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using \<rho>.T0.antipar(1) iso_assoc
invert_side_of_triangle(2)
[of "(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])"
"\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]" "\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1 \<star> w]"]
by fastforce
hence "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] =
(\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"
using comp_assoc by presburger
moreover have "seq (inv (\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w)) \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) 1 by fastforce
ultimately show ?thesis
using \<rho>.T0.antipar(1) iso_assoc
invert_side_of_triangle(1)
[of "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]" "\<a>\<^sup>-\<^sup>1[r, r\<^sub>0, p\<^sub>1] \<star> w"
"\<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]"]
by fastforce
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot> ((\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w))) \<cdot>
(((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w]) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, r\<^sub>0 \<star> p\<^sub>1, w] = \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w] \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using assoc'_naturality [of r r\<^sub>0s\<^sub>1.\<phi> w] r\<^sub>0s\<^sub>1.cospan by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w]) \<cdot>
(r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<rho>.composite_cell (p\<^sub>1 \<star> w) \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, s\<^sub>1 \<star> p\<^sub>0, w] =
\<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w)"
using assoc'_naturality [of r "\<sigma> \<star> p\<^sub>0" w]
by (simp add: composable)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
((r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
proof -
have "\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] =
r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]"
proof -
have "\<a>[r, s, (s\<^sub>0 \<star> p\<^sub>0) \<star> w] \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, (\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<alpha>_def \<a>'_def composable by simp
also have "... = \<lbrace>\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using composable
by (intro E.eval_eqI, simp_all)
also have "... = r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]"
using \<alpha>_def \<a>'_def composable by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (r \<star> s \<star> \<theta>) \<cdot>
(r \<star> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<rho>.composite_cell (p\<^sub>1 \<star> w)
(((\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])"
using assms(3) arrI \<rho>.T0.antipar(1) whisker_left by auto
also have "... = (r \<star> (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot> ((\<sigma> \<star> p\<^sub>0) \<star> w)) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using \<rho>.T0.antipar(1) comp_assoc whisker_left by auto
also have "... = (r \<star> (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w)"
using assoc_naturality [of \<sigma> p\<^sub>0 w] comp_assoc by simp
finally show ?thesis
using assms(3-4) by simp
qed
qed
lemma composite_is_tabulation:
shows "tabulation V H \<a> \<i> src trg (r \<star> s) tab (s\<^sub>0 \<star> p\<^sub>0) (r\<^sub>1 \<star> p\<^sub>1)"
proof
show "\<And>u \<omega>. \<lbrakk> ide u; \<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> (r \<star> s) \<star> u\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
proof -
fix u \<omega>
assume u: "ide u"
assume \<omega>: "\<guillemotleft>\<omega> : dom \<omega> \<Rightarrow> (r \<star> s) \<star> u\<guillemotright>"
let ?v = "dom \<omega>"
have 1: "\<guillemotleft>\<a>[r, s, u] \<cdot> \<omega> : ?v \<Rightarrow> r \<star> s \<star> u\<guillemotright>"
proof -
(*
* TODO: I think this highlights the current issue with assoc_in_hom:
* it can't be applied automatically here because there isn't any way to
* obtain the equations src r = trg s and src s = trg u from assumption \<omega>.
* Maybe this can be improved with a little bit of thought, but not while
* a lot of other stuff is being changed, too.
*)
have "src r = trg s \<and> src s = trg u"
by (metis \<omega> arr_cod hseq_char in_homE hcomp_simps(1))
thus ?thesis
using u \<omega> by fastforce
qed
obtain w\<^sub>r \<theta>\<^sub>r \<nu>\<^sub>r
where w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r: "ide w\<^sub>r \<and> \<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright> \<and>
\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright> \<and> iso \<nu>\<^sub>r \<and>
\<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r \<cdot> \<nu>\<^sub>r = \<a>[r, s, u] \<cdot> \<omega>"
using u \<omega> \<rho>.T1 [of "s \<star> u" "\<a>[r, s, u] \<cdot> \<omega>"]
by (metis 1 \<rho>.ide_base \<sigma>.ide_base arr_cod composable ide_hcomp in_homE
match_1 not_arr_null seq_if_composable)
text \<open>
$$
\xymatrix{
&& X \ar@ {.>}[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
\ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}} \\
&& \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& \\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
\<close>
text \<open>We need some simps, etc., otherwise the subsequent diagram chase is very painful.\<close>
have w\<^sub>r: "ide w\<^sub>r"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have [simp]: "src w\<^sub>r = src u"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<omega> 1 comp_arr_dom in_homE seqE hcomp_simps(1) vseq_implies_hpar(1)
by (metis src_hcomp)
have [simp]: "trg w\<^sub>r = src \<rho>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r
by (metis 1 arrI not_arr_null seqE seq_if_composable)
have \<theta>\<^sub>r_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have \<theta>\<^sub>r_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>r : src u \<rightarrow> trg s\<guillemotright>"
using \<theta>\<^sub>r_in_hom src_cod [of \<theta>\<^sub>r] trg_cod [of \<theta>\<^sub>r]
by (metis \<open>src w\<^sub>r = src u\<close> \<sigma>.leg1_simps(4) arr_dom in_hhomI in_homE r\<^sub>0s\<^sub>1.cospan
src_hcomp trg_hcomp vconn_implies_hpar(1) vconn_implies_hpar(2))
have [simp]: "src \<theta>\<^sub>r = src u" using \<theta>\<^sub>r_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>r = trg s" using \<theta>\<^sub>r_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>r = r\<^sub>0 \<star> w\<^sub>r" using \<theta>\<^sub>r_in_hom by blast
have [simp]: "cod \<theta>\<^sub>r = s \<star> u" using \<theta>\<^sub>r_in_hom by blast
have \<nu>\<^sub>r_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright>" using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
have \<nu>\<^sub>r_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>r : src u \<rightarrow> trg r\<guillemotright>"
using \<nu>\<^sub>r_in_hom src_dom [of \<nu>\<^sub>r] trg_dom [of \<nu>\<^sub>r]
by (metis \<open>src w\<^sub>r = src u\<close> \<rho>.leg1_simps(4) arr_cod in_hhomI in_homE
src_hcomp trg_hcomp vconn_implies_hpar(3) vconn_implies_hpar(4))
have [simp]: "src \<nu>\<^sub>r = src u" using \<nu>\<^sub>r_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>r = trg r" using \<nu>\<^sub>r_in_hhom by auto
have [simp]: "dom \<nu>\<^sub>r = ?v" using \<nu>\<^sub>r_in_hom by auto
have [simp]: "cod \<nu>\<^sub>r = r\<^sub>1 \<star> w\<^sub>r" using \<nu>\<^sub>r_in_hom by auto
have iso_\<nu>\<^sub>r: "iso \<nu>\<^sub>r" using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by simp
obtain w\<^sub>s \<theta>\<^sub>s \<nu>\<^sub>s
where w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s: "ide w\<^sub>s \<and> \<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright> \<and> \<guillemotleft>\<nu>\<^sub>s : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s\<^sub>1 \<star> w\<^sub>s\<guillemotright> \<and> iso \<nu>\<^sub>s \<and>
\<sigma>.composite_cell w\<^sub>s \<theta>\<^sub>s \<cdot> \<nu>\<^sub>s = \<theta>\<^sub>r"
using u w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<sigma>.T1 [of u \<theta>\<^sub>r] by auto
text \<open>
$$
\xymatrix{
&& X \ar[ddl]^{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
\ar@/ur20pt/[dddrr]^{u} \ar@ {.>}[ddr]_{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}}
\xtwocell[ddd]{}\omit{^{<1>\nu_s}} \\
&& \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
\<close>
have w\<^sub>s: "ide w\<^sub>s"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have [simp]: "src w\<^sub>s = src u"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s src_cod
by (metis \<sigma>.leg0_simps(2) \<sigma>.tab_simps(2) \<theta>\<^sub>r_in_hom arrI hseqI' ideD(1) seqE
seq_if_composable src_hcomp vconn_implies_hpar(3))
have [simp]: "trg w\<^sub>s = src \<sigma>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s
by (metis \<sigma>.tab_simps(2) arr_dom in_homE not_arr_null seq_if_composable)
have \<theta>\<^sub>s_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have \<theta>\<^sub>s_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>s : src u \<rightarrow> src s\<guillemotright>"
using \<theta>\<^sub>s_in_hom src_cod trg_cod
by (metis \<theta>\<^sub>r_in_hom arrI hseqE in_hhom_def seqE vconn_implies_hpar(1)
vconn_implies_hpar(3) w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s)
have [simp]: "src \<theta>\<^sub>s = src u" using \<theta>\<^sub>s_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>s = src s" using \<theta>\<^sub>s_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>s = s\<^sub>0 \<star> w\<^sub>s" using \<theta>\<^sub>s_in_hom by blast
have [simp]: "cod \<theta>\<^sub>s = u" using \<theta>\<^sub>s_in_hom by blast
have \<nu>\<^sub>s_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>s : r\<^sub>0 \<star> w\<^sub>r \<Rightarrow> s\<^sub>1 \<star> w\<^sub>s\<guillemotright>" using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
have \<nu>\<^sub>s_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>s : src u \<rightarrow> trg s\<guillemotright>"
using \<nu>\<^sub>s_in_hom src_dom trg_cod
by (metis \<open>src \<theta>\<^sub>r = src u\<close> \<open>trg \<theta>\<^sub>r = trg s\<close> \<theta>\<^sub>r_in_hom in_hhomI in_homE src_dom trg_dom)
have [simp]: "src \<nu>\<^sub>s = src u" using \<nu>\<^sub>s_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>s = trg s" using \<nu>\<^sub>s_in_hhom by auto
have [simp]: "dom \<nu>\<^sub>s = r\<^sub>0 \<star> w\<^sub>r" using \<nu>\<^sub>s_in_hom by auto
have [simp]: "cod \<nu>\<^sub>s = s\<^sub>1 \<star> w\<^sub>s" using \<nu>\<^sub>s_in_hom by auto
have iso_\<nu>\<^sub>s: "iso \<nu>\<^sub>s" using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
obtain w \<theta>\<^sub>t \<nu>\<^sub>t
where w\<theta>\<^sub>t\<nu>\<^sub>t: "ide w \<and> \<guillemotleft>\<theta>\<^sub>t : p\<^sub>0 \<star> w \<Rightarrow> w\<^sub>s\<guillemotright> \<and> \<guillemotleft>\<nu>\<^sub>t : w\<^sub>r \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu>\<^sub>t \<and>
(s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) = \<nu>\<^sub>s"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s iso_\<nu>\<^sub>s r\<^sub>0s\<^sub>1.\<phi>_biuniversal_prop(1) [of w\<^sub>s w\<^sub>r \<nu>\<^sub>s] by blast
text \<open>
$$
\xymatrix{
&& X \ar[ddl]_{w_r} \ar@/ul20pt/[dddll]_{v} \xtwocell[dddll]{}\omit{^{<1.5>\nu_r}}
\ar@/ur20pt/[dddrr]^{u} \ar[ddr]^{w_s} \xtwocell[dddrr]{}\omit{^{<-1.5>\theta_s}}
\ar@ {.>}[d]^{w} \xtwocell[ddl]{}\omit{^<-2>{\nu_t}} \xtwocell[ddr]{}\omit{^<2>{\theta_t}} \\
&& {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
$$
\<close>
text \<open>{\bf Note:} \<open>w\<close> is not necessarily a map.\<close>
have w: "ide w"
using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have [simp]: "src w = src u"
using w\<theta>\<^sub>t\<nu>\<^sub>t src_cod
by (metis \<nu>\<^sub>s_in_hom \<open>src \<nu>\<^sub>s = src u\<close> arrI seqE src_hcomp src_vcomp vseq_implies_hpar(1))
have [simp]: "trg w = src p\<^sub>0"
using w\<theta>\<^sub>t\<nu>\<^sub>t
by (metis \<nu>\<^sub>s_in_hom arrI not_arr_null r\<^sub>0s\<^sub>1.\<phi>_simps(2) seqE seq_if_composable)
have \<theta>\<^sub>t_in_hom [intro]: "\<guillemotleft>\<theta>\<^sub>t : p\<^sub>0 \<star> w \<Rightarrow> w\<^sub>s\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have \<theta>\<^sub>t_in_hhom [intro]: "\<guillemotleft>\<theta>\<^sub>t : src u \<rightarrow> src \<sigma>\<guillemotright>"
using \<theta>\<^sub>t_in_hom src_cod trg_cod \<open>src w\<^sub>s = src u\<close> \<open>trg w\<^sub>s = src \<sigma>\<close> by fastforce
have [simp]: "src \<theta>\<^sub>t = src u" using \<theta>\<^sub>t_in_hhom by auto
have [simp]: "trg \<theta>\<^sub>t = src \<sigma>" using \<theta>\<^sub>t_in_hhom by auto
have [simp]: "dom \<theta>\<^sub>t = p\<^sub>0 \<star> w" using \<theta>\<^sub>t_in_hom by blast
have (* [simp]: *) "cod \<theta>\<^sub>t = w\<^sub>s" using \<theta>\<^sub>t_in_hom by blast
have \<nu>\<^sub>t_in_hom [intro]: "\<guillemotleft>\<nu>\<^sub>t : w\<^sub>r \<Rightarrow> p\<^sub>1 \<star> w\<guillemotright>" using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
have \<nu>\<^sub>t_in_hhom [intro]: "\<guillemotleft>\<nu>\<^sub>t : src u \<rightarrow> src \<rho>\<guillemotright>"
using \<nu>\<^sub>t_in_hom src_dom trg_dom \<open>src w\<^sub>r = src u\<close> \<open>trg w\<^sub>r = src \<rho>\<close> by fastforce
have [simp]: "src \<nu>\<^sub>t = src u" using \<nu>\<^sub>t_in_hhom by auto
have [simp]: "trg \<nu>\<^sub>t = src \<rho>" using \<nu>\<^sub>t_in_hhom by auto
have (* [simp]: *) "dom \<nu>\<^sub>t = w\<^sub>r" using \<nu>\<^sub>t_in_hom by auto
have [simp]: "cod \<nu>\<^sub>t = p\<^sub>1 \<star> w" using \<nu>\<^sub>t_in_hom by auto
have iso_\<nu>\<^sub>t: "iso \<nu>\<^sub>t" using w\<theta>\<^sub>t\<nu>\<^sub>t by simp
define \<theta> where "\<theta> = \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
have \<theta>: "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
proof (unfold \<theta>_def, intro comp_in_homI)
show "\<guillemotleft>\<a>[s\<^sub>0, p\<^sub>0, w] : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> s\<^sub>0 \<star> p\<^sub>0 \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by auto
show "\<guillemotleft>s\<^sub>0 \<star> \<theta>\<^sub>t : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0 \<star> w\<^sub>s\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by auto
show "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> w\<^sub>s \<Rightarrow> u\<guillemotright>"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
qed
define \<nu> where "\<nu> = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
have \<nu>: "\<guillemotleft>\<nu> : ?v \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright>"
proof (unfold \<nu>_def, intro comp_in_homI)
show "\<guillemotleft>\<nu>\<^sub>r : ?v \<Rightarrow> r\<^sub>1 \<star> w\<^sub>r\<guillemotright>"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r by blast
show "\<guillemotleft>r\<^sub>1 \<star> \<nu>\<^sub>t : r\<^sub>1 \<star> w\<^sub>r \<Rightarrow> r\<^sub>1 \<star> p\<^sub>1 \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] : r\<^sub>1 \<star> p\<^sub>1 \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright>"
using w\<theta>\<^sub>t\<nu>\<^sub>t assoc_in_hom by (simp add: \<rho>.T0.antipar(1))
qed
have iso_\<nu>: "iso \<nu>"
using \<nu> w\<theta>\<^sub>t\<nu>\<^sub>t w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<rho>.T0.antipar(1)
by (unfold \<nu>_def, intro isos_compose) auto
have *: "arr ((s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot> (\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t \<theta>\<^sub>r_in_hom comp_assoc by auto
have "((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = \<omega>"
proof -
have "seq (r \<star> \<theta>\<^sub>r) (\<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r)"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r \<rho>.base_simps(2) composable by fastforce
hence "\<omega> = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> \<rho>.composite_cell w\<^sub>r \<theta>\<^sub>r \<cdot> \<nu>\<^sub>r"
using w\<^sub>r\<theta>\<^sub>r\<nu>\<^sub>r invert_side_of_triangle(1) iso_assoc
by (metis 1 \<rho>.ide_base \<sigma>.ide_base arrI assoc'_eq_inv_assoc composable hseq_char'
seqE seq_if_composable u vconn_implies_hpar(2) vconn_implies_hpar(4) w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s)
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> \<rho>.composite_cell w\<^sub>r (\<sigma>.composite_cell w\<^sub>s \<theta>\<^sub>s \<cdot> \<nu>\<^sub>s) \<cdot> \<nu>\<^sub>r"
using w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s by simp
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc by simp
text \<open>Rearrange to create \<open>\<theta>\<close> and \<open>\<nu>\<close>, leaving \<open>tab\<close> in the middle.\<close>
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> \<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
((\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t)) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (\<a>[s, s\<^sub>0, w\<^sub>s] \<cdot>
((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t)) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) = \<sigma> \<star> \<theta>\<^sub>t"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> \<sigma>.tab_simps(1) \<sigma>.tab_simps(4) arrI w\<theta>\<^sub>t\<nu>\<^sub>t)
also have "... = ((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
using comp_arr_dom comp_cod_arr interchange w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t \<sigma>.tab_in_hom
by (metis \<open>dom \<theta>\<^sub>t = p\<^sub>0 \<star> w\<close> \<sigma>.tab_simps(5) arrI)
finally have "(\<sigma> \<star> w\<^sub>s) \<cdot> (s\<^sub>1 \<star> \<theta>\<^sub>t) = ((s \<star> s\<^sub>0) \<star> \<theta>\<^sub>t) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot> (\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
using assoc_naturality [of s s\<^sub>0 \<theta>\<^sub>t] w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> arrI by force
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] \<cdot>
(\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) =
(r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
proof -
have "seq ((s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t))
(\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
proof -
have "seq (s \<star> \<theta>\<^sub>s)
((s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t))"
using \<open>\<guillemotleft>\<a>[r, s, u] \<cdot> \<omega> : dom \<omega> \<Rightarrow> r \<star> s \<star> u\<guillemotright>\<close> calculation by blast
thus ?thesis
using comp_assoc by presburger
qed
thus ?thesis
using whisker_left [of r "(s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
"\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t)"]
w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s w\<theta>\<^sub>t\<nu>\<^sub>t comp_assoc
by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> ((r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r]) \<cdot>
(\<rho> \<star> w\<^sub>r) \<cdot> \<nu>\<^sub>r"
proof -
have "seq (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) (r\<^sub>0 \<star> \<nu>\<^sub>t)"
using 1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\<theta>\<^sub>t\<nu>\<^sub>t
apply (intro seqI' comp_in_homI) by auto
hence "r \<star> (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> (r\<^sub>0 \<star> \<nu>\<^sub>t) =
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> (r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t)"
using whisker_left by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r)) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> r\<^sub>0 \<star> \<nu>\<^sub>t) \<cdot> \<a>[r, r\<^sub>0, w\<^sub>r] = \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> ((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t)"
using assoc_naturality [of r r\<^sub>0 \<nu>\<^sub>t] \<nu>\<^sub>t_in_hom by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t))) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r) = \<rho> \<star> \<nu>\<^sub>t"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>dom \<nu>\<^sub>t = w\<^sub>r\<close> \<rho>.tab_simps(1) \<rho>.tab_simps(5) arrI w\<theta>\<^sub>t\<nu>\<^sub>t)
also have "... = (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t)"
using comp_arr_dom comp_cod_arr interchange
by (metis \<open>cod \<nu>\<^sub>t = p\<^sub>1 \<star> w\<close> \<open>trg \<nu>\<^sub>t = src \<rho>\<close> \<rho>.T0.antipar(1) \<rho>.tab_simps(1)
\<rho>.tab_simps(2) \<rho>.tab_simps(4) r\<^sub>0s\<^sub>1.base_simps(2) trg.preserves_reflects_arr
trg_hcomp)
finally have "((r \<star> r\<^sub>0) \<star> \<nu>\<^sub>t) \<cdot> (\<rho> \<star> w\<^sub>r) = (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t)" by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> ((\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) =
((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
proof -
have "seq (s \<star> \<theta>\<^sub>s) (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using \<theta>\<^sub>s_in_hom \<theta>\<^sub>s_in_hhom \<theta>\<^sub>t_in_hom \<theta>\<^sub>t_in_hhom 1 calculation by blast
moreover have "src r = trg (s \<star> \<theta>\<^sub>s)"
using composable hseqI by force
ultimately
have "\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> (s \<star> \<theta>\<^sub>s) \<cdot> (s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) =
(\<a>\<^sup>-\<^sup>1[r, s, u] \<cdot> (r \<star> s \<star> \<theta>\<^sub>s)) \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using whisker_left comp_assoc by simp
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> w\<^sub>s] \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)"
using assoc_naturality comp_assoc
by (metis \<open>cod \<theta>\<^sub>s = u\<close> \<open>dom \<theta>\<^sub>s = s\<^sub>0 \<star> w\<^sub>s\<close> \<open>trg \<theta>\<^sub>s = src s\<close>
\<rho>.base_simps(2-4) \<sigma>.base_simps(2-4) arrI assoc'_naturality composable w\<^sub>s\<theta>\<^sub>s\<nu>\<^sub>s)
also have "... = (((r \<star> s) \<star> \<theta>\<^sub>s) \<cdot> ((r \<star> s) \<star> s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> w\<^sub>s] \<cdot> (r \<star> s \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) =
((r \<star> s) \<star> s\<^sub>0 \<star> \<theta>\<^sub>t) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
using arrI hseq_char assoc'_naturality [of r s "s\<^sub>0 \<star> \<theta>\<^sub>t"] \<open>cod \<theta>\<^sub>t = w\<^sub>s\<close> composable
by auto
thus ?thesis
using comp_assoc by auto
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w]"
using \<theta>_def \<theta> whisker_left
by (metis (full_types) arrI cod_comp ide_base seqE seqI)
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot>
((r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot>
((\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w])) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] = \<a>[s \<star> s\<^sub>0, p\<^sub>0, w] \<cdot> ((\<sigma> \<star> p\<^sub>0) \<star> w)"
using assoc_naturality [of \<sigma> p\<^sub>0 w] by (simp add: w\<theta>\<^sub>t\<nu>\<^sub>t)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]) \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
using r\<^sub>0s\<^sub>1.p\<^sub>1_simps w\<theta>\<^sub>t\<nu>\<^sub>t whisker_left comp_assoc by force
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
(\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]) \<cdot>
(\<rho> \<star> p\<^sub>1 \<star> w)) \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] =
\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
proof -
have 1: "(r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) =
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using pentagon
by (simp add: \<rho>.T0.antipar(1) w)
moreover have 2: "seq \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w by simp
moreover have "inv (r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]) = r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w by simp
ultimately have "\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) =
((r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]) \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w comp_assoc
invert_side_of_triangle(1)
[of "\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] \<cdot> \<a>[r \<star> r\<^sub>0, p\<^sub>1, w]" "r \<star> \<a>[r\<^sub>0, p\<^sub>1, w]"
"\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)"]
by simp
hence "(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w] =
(\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w]"
using \<rho>.T0.antipar(1) w
invert_side_of_triangle(2)
[of "\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w)"
"(r \<star> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w]" "\<a>[r \<star> r\<^sub>0, p\<^sub>1, w]"]
using \<open>trg w = src p\<^sub>0\<close> by simp
thus ?thesis
using comp_assoc by presburger
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> (r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>[r, r\<^sub>0 \<star> p\<^sub>1, w]) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "\<a>\<^sup>-\<^sup>1[r \<star> r\<^sub>0, p\<^sub>1, w] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w) = ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using assoc'_naturality [of \<rho> p\<^sub>1 w] by (simp add: \<rho>.T0.antipar(1) w\<theta>\<^sub>t\<nu>\<^sub>t)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> ((r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w]) \<cdot>
((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>[r, r\<^sub>0 \<star> p\<^sub>1, w] = \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w] \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w)"
using assoc_naturality [of r r\<^sub>0s\<^sub>1.\<phi> w] r\<^sub>0s\<^sub>1.cospan w\<theta>\<^sub>t\<nu>\<^sub>t by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot>
(((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot>
((\<rho> \<star> p\<^sub>1) \<star> w)) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "(r \<star> (\<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> \<a>[r, s\<^sub>1 \<star> p\<^sub>0, w] =
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w)"
proof -
have "arr w \<and> dom w = w \<and> cod w = w"
using ide_char w by blast
then show ?thesis
using assoc_naturality [of r "\<sigma> \<star> p\<^sub>0" w] composable by auto
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot>
(r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot>
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w)) \<cdot>
(tab \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> (r\<^sub>1 \<star> \<nu>\<^sub>t) \<cdot> \<nu>\<^sub>r"
proof -
have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot> ((\<rho> \<star> p\<^sub>1) \<star> w) =
(r \<star> \<sigma> \<star> p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
using w \<rho>.T0.antipar(1) composable whisker_right by auto
also have "... = (((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
proof -
have "((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) =
r \<star> \<sigma> \<star> p\<^sub>0"
proof -
have "((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot>
(r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) =
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> ((r \<star> s \<star> s\<^sub>0 \<star> p\<^sub>0) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]))) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using comp_assoc_assoc' by (simp add: composable)
also have "... = ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0])) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using comp_cod_arr by (simp add: composable)
also have "... = ((r \<star> (s \<star> s\<^sub>0) \<star> p\<^sub>0)) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0)"
using whisker_left comp_assoc_assoc' assoc_in_hom hseqI'
by (metis \<rho>.ide_base \<sigma>.base_simps(2) \<sigma>.ide_base \<sigma>.ide_leg0
\<sigma>.leg0_simps(2-3) \<sigma>.leg1_simps(3) r\<^sub>0s\<^sub>1.ide_leg0
r\<^sub>0s\<^sub>1.leg0_simps(2) r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1))
also have "... = r \<star> \<sigma> \<star> p\<^sub>0"
using comp_cod_arr
by (simp add: composable)
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = (r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> \<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0]) \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0]) \<cdot> (r \<star> \<sigma> \<star> p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1] \<cdot> (\<rho> \<star> p\<^sub>1) \<star> w"
using comp_assoc by presburger
also have "... = (r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<cdot> \<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<cdot> tab \<star> w"
using tab_def by simp
also have "... = ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) \<cdot> (tab \<star> w)"
using w \<rho>.T0.antipar(1) composable comp_assoc whisker_right by auto
finally have "((r \<star> \<sigma> \<star> p\<^sub>0) \<star> w) \<cdot> ((r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<star> w) \<cdot> (\<a>[r, r\<^sub>0, p\<^sub>1] \<star> w) \<cdot>
((\<rho> \<star> p\<^sub>1) \<star> w) =
((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) \<cdot> (tab \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w])) \<cdot>
\<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot> (\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) =
((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> p\<^sub>0 \<star> w] \<cdot> (r \<star> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (r \<star> \<a>[s \<star> s\<^sub>0, p\<^sub>0, w]) \<cdot>
\<a>[r, (s \<star> s\<^sub>0) \<star> p\<^sub>0, w] \<cdot> ((r \<star> \<a>\<^sup>-\<^sup>1[s, s\<^sub>0, p\<^sub>0]) \<star> w) \<cdot>
(\<a>[r, s, s\<^sub>0 \<star> p\<^sub>0] \<star> w) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, (\<^bold>\<langle>s\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>) \<^bold>\<cdot> (\<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle>, \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>w\<^bold>\<rangle>)\<rbrace>"
using w comp_assoc \<a>'_def \<alpha>_def composable by simp
also have "... = \<lbrace>((\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>s\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>r\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>s\<^bold>\<rangle>, \<^bold>\<langle>s\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>w\<^bold>\<rangle>\<^bold>]\<rbrace>"
using w composable
apply (intro E.eval_eqI) by simp_all
also have "... = ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w]"
using w comp_assoc \<a>'_def \<alpha>_def composable by simp
finally show ?thesis by simp
qed
thus ?thesis
using \<nu>_def comp_assoc by simp
qed
also have "... = ((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu>"
proof -
have "((r \<star> s) \<star> \<theta>\<^sub>s \<cdot> (s\<^sub>0 \<star> \<theta>\<^sub>t)) \<cdot> ((r \<star> s) \<star> \<a>[s\<^sub>0, p\<^sub>0, w]) = (r \<star> s) \<star> \<theta>"
using \<theta>_def w whisker_left composable
by (metis \<theta> arrI ide_base comp_assoc)
thus ?thesis
using comp_assoc by presburger
qed
finally show "((r \<star> s) \<star> \<theta>) \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w] \<cdot> (tab \<star> w) \<cdot> \<nu> = \<omega>"
by simp
qed
thus "\<exists>w \<theta> \<nu>. ide w \<and> \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright> \<and>
\<guillemotleft>\<nu> : dom \<omega> \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w\<guillemotright> \<and> iso \<nu> \<and>
composite_cell w \<theta> \<cdot> \<nu> = \<omega>"
using w\<theta>\<^sub>t\<nu>\<^sub>t \<theta> \<nu> iso_\<nu> comp_assoc by metis
qed
show "\<And>u w w' \<theta> \<theta>' \<beta>.
\<lbrakk> ide w; ide w'; \<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>; \<guillemotleft>\<theta>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w' \<Rightarrow> u\<guillemotright>;
\<guillemotleft>\<beta> : (r\<^sub>1 \<star> p\<^sub>1) \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w'\<guillemotright>;
composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta> \<rbrakk> \<Longrightarrow>
\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
fix u w w' \<theta> \<theta>' \<beta>
assume w: "ide w"
assume w': "ide w'"
assume \<theta>: "\<guillemotleft>\<theta> : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> u\<guillemotright>"
assume \<theta>': "\<guillemotleft>\<theta>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w' \<Rightarrow> u\<guillemotright>"
assume \<beta>: "\<guillemotleft>\<beta> : (r\<^sub>1 \<star> p\<^sub>1) \<star> w \<Rightarrow> (r\<^sub>1 \<star> p\<^sub>1) \<star> w'\<guillemotright>"
assume eq: "composite_cell w \<theta> = composite_cell w' \<theta>' \<cdot> \<beta>"
interpret uw\<theta>w'\<theta>'\<beta>: uw\<theta>w'\<theta>'\<beta> V H \<a> \<i> src trg
\<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close> u w \<theta> w' \<theta>' \<beta>
using w w' \<theta> \<theta>' \<beta> eq composable tab_in_hom comp_assoc
by unfold_locales auto
text \<open>
$$
\begin{array}{ll}
\xymatrix{
&& X \ar[d]_{w'} \xtwocell[ddl]{}\omit{^{\beta}}
\ar@/ul20pt/[dddll]_<>(0.25){w}|<>(0.33)@ {>}_<>(0.5){p_1}|<>(0.67)@ {>}_<>(0.75){r_1}
\ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta'}}\\
&& {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\\
\hspace{5cm}
=
\qquad
\xy/50pt/
\xymatrix{
&& X \ar[d]_{w} \ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta}}\\
&& {\rm src}~\phi \ar[dl]^{p_1} \ar[dr]_{p_0} \ddtwocell\omit{^\phi} \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
\<close>
text \<open>
First apply property \<open>\<rho>.T2\<close> using \<open>\<guillemotleft>\<beta>\<^sub>r : r\<^sub>1 \<star> p\<^sub>1 \<star> w \<Rightarrow> r\<^sub>1 \<star> p\<^sub>1 \<star> w'\<guillemotright>\<close>
(obtained by composing \<open>\<beta>\<close> with associativities) and ``everything to the right''
as \<open>\<theta>\<^sub>r\<close> and \<open>\<theta>\<^sub>r'\<close>. This yields \<open>\<guillemotleft>\<gamma>\<^sub>r : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>\<close>.
Next, apply property \<open>\<rho>.T2\<close> to obtain \<open>\<guillemotleft>\<gamma>\<^sub>s : p\<^sub>0 \<star> w \<Rightarrow> p\<^sub>0 \<star> w'\<guillemotright>\<close>.
For this use \<open>\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>\<close> and \<open>\<guillemotleft>\<theta>\<^sub>s' : s\<^sub>0 \<star> p\<^sub>0 \<star> w'\<guillemotright>\<close>
obtained by composing \<open>\<theta>\<close> and \<open>\<theta>'\<close> with associativities.
We also need \<open>\<guillemotleft>\<beta>\<^sub>s : s\<^sub>1 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>1 \<star> p\<^sub>0 \<star> w'\<guillemotright>\<close>.
To get this, transport \<open>r\<^sub>0 \<star> \<gamma>\<^sub>r\<close> across \<open>\<phi>\<close>; we need \<open>\<phi>\<close> to be an isomorphism
in order to do this.
Finally, apply the biuniversal property of \<open>\<phi>\<close> to get \<open>\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright>\<close>
and verify the required equation.
\<close>
let ?w\<^sub>r = "p\<^sub>1 \<star> w"
have w\<^sub>r: "ide ?w\<^sub>r" by simp
let ?w\<^sub>r' = "p\<^sub>1 \<star> w'"
have w\<^sub>r': "ide ?w\<^sub>r'" by simp
define \<theta>\<^sub>r where "\<theta>\<^sub>r = (s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
have \<theta>\<^sub>r: "\<guillemotleft>\<theta>\<^sub>r : r\<^sub>0 \<star> ?w\<^sub>r \<Rightarrow> s \<star> u\<guillemotright>"
unfolding \<theta>\<^sub>r_def
using \<rho>.T0.antipar(1) by fastforce
define \<theta>\<^sub>r' where "\<theta>\<^sub>r' = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']"
have \<theta>\<^sub>r': "\<guillemotleft>\<theta>\<^sub>r' : r\<^sub>0 \<star> ?w\<^sub>r' \<Rightarrow> s \<star> u\<guillemotright>"
unfolding \<theta>\<^sub>r'_def
using \<rho>.T0.antipar(1) by fastforce
define \<beta>\<^sub>r where "\<beta>\<^sub>r = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
have \<beta>\<^sub>r: "\<guillemotleft>\<beta>\<^sub>r : r\<^sub>1 \<star> ?w\<^sub>r \<Rightarrow> r\<^sub>1 \<star> ?w\<^sub>r'\<guillemotright>"
unfolding \<beta>\<^sub>r_def
using \<rho>.T0.antipar(1) by force
have eq\<^sub>r: "\<rho>.composite_cell ?w\<^sub>r \<theta>\<^sub>r = \<rho>.composite_cell ?w\<^sub>r' \<theta>\<^sub>r' \<cdot> \<beta>\<^sub>r"
text \<open>
$$
\begin{array}{ll}
\xymatrix{
&& X \ar[ddl]^{w_r'} \xtwocell[dddll]{}\omit{^<2>{\beta_r}}
\ar@/ul20pt/[dddll]_<>(0.33){w_r}|<>(0.67)@ {>}_<>(0.75){r_1}
\ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r'}}\\
&& \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& \\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\\
\hspace{5cm}
=\qquad
\xy/50pt/
\xymatrix{
&& X \ar[ddl]^{w_r} \ar@/ur20pt/[dddrr]^{u} \xtwocell[dddr]{}\omit{^{\theta_r}}\\
&& \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \dtwocell\omit{^\rho}
&& \\
{\rm trg}~r && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
\<close>
proof -
have "\<rho>.composite_cell ?w\<^sub>r \<theta>\<^sub>r = \<a>[r, s, u] \<cdot> composite_cell w \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<theta>\<^sub>r_def technical uw\<theta>w'\<theta>'\<beta>.uw\<theta>.uw\<theta> by blast
also have "... = \<a>[r, s, u] \<cdot> (((r \<star> s) \<star> \<theta>') \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w'] \<cdot>
(tab \<star> w') \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using eq comp_assoc by simp
also have "... = (r \<star> \<theta>\<^sub>r') \<cdot> \<a>[r, r\<^sub>0, ?w\<^sub>r'] \<cdot> (\<rho> \<star> ?w\<^sub>r') \<cdot> \<beta>\<^sub>r"
proof -
have "\<a>[r, s, u] \<cdot> (composite_cell w' \<theta>' \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (r \<star> (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot>
\<a>[r, r\<^sub>0, p\<^sub>1 \<star> w'] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w') \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
proof -
have "\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] =
\<a>[r, s, u] \<cdot> composite_cell w' \<theta>' \<cdot>
((\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta>) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
proof -
have "(\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta> = \<beta>"
using comp_cod_arr \<rho>.T0.antipar(1) \<beta> comp_assoc_assoc' by simp
thus ?thesis by argo
qed
also have "... = (\<a>[r, s, u] \<cdot> ((r \<star> s) \<star> \<theta>') \<cdot> \<a>[r \<star> s, s\<^sub>0 \<star> p\<^sub>0, w'] \<cdot> (tab \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = ((r \<star> (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> \<a>[r, r\<^sub>0, p\<^sub>1 \<star> w'] \<cdot> (\<rho> \<star> p\<^sub>1 \<star> w')) \<cdot>
\<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<theta>\<^sub>r'_def technical [of w' \<theta>' u ?w\<^sub>r' \<theta>\<^sub>r'] comp_assoc by fastforce
finally show ?thesis
using comp_assoc by simp
qed
finally show ?thesis
using \<theta>\<^sub>r'_def \<beta>\<^sub>r_def comp_assoc by auto
qed
finally show ?thesis
using comp_assoc by presburger
qed
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>"
using eq\<^sub>r \<rho>.T2 [of ?w\<^sub>r ?w\<^sub>r' \<theta>\<^sub>r "s \<star> u" \<theta>\<^sub>r' \<beta>\<^sub>r] w\<^sub>r w\<^sub>r' \<theta>\<^sub>r \<theta>\<^sub>r' \<beta>\<^sub>r by blast
obtain \<gamma>\<^sub>r where \<gamma>\<^sub>r: "\<guillemotleft>\<gamma>\<^sub>r : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>\<^sub>r"
using 1 by blast
let ?w\<^sub>s = "p\<^sub>0 \<star> w"
have w\<^sub>s: "ide ?w\<^sub>s" by simp
let ?w\<^sub>s' = "p\<^sub>0 \<star> w'"
have w\<^sub>s': "ide ?w\<^sub>s'" by simp
define \<theta>\<^sub>s where "\<theta>\<^sub>s = \<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
have \<theta>\<^sub>s: "\<guillemotleft>\<theta>\<^sub>s : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> u\<guillemotright>"
using \<theta>\<^sub>s_def by auto
define \<theta>\<^sub>s' where "\<theta>\<^sub>s' = \<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']"
have \<theta>\<^sub>s': "\<guillemotleft>\<theta>\<^sub>s' : s\<^sub>0 \<star> p\<^sub>0 \<star> w' \<Rightarrow> u\<guillemotright>"
using \<theta>\<^sub>s'_def by auto
define \<beta>\<^sub>s where "\<beta>\<^sub>s = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
have \<beta>\<^sub>s: "\<guillemotleft>\<beta>\<^sub>s : s\<^sub>1 \<star> ?w\<^sub>s \<Rightarrow> s\<^sub>1 \<star> ?w\<^sub>s'\<guillemotright>"
unfolding \<beta>\<^sub>s_def
using \<gamma>\<^sub>r r\<^sub>0s\<^sub>1.\<phi>_in_hom(2) r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1)
apply (intro comp_in_homI)
apply auto
by auto
have eq\<^sub>s: "\<sigma>.composite_cell (p\<^sub>0 \<star> w) (\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) =
\<sigma>.composite_cell (p\<^sub>0 \<star> w') (\<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<beta>\<^sub>s"
text \<open>
$$
\begin{array}{ll}
\xy/67pt/
\xymatrix{
&& X \ar[d]^{w'} \ar@/l10pt/[dl]_{w} \ddltwocell\omit{^{\gamma_r}}
\ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s'}}\\
& {\rm src}~\phi \ar[dr]_{p_1} \ar[d]_{p_0}
& {\rm src}~\phi \ar[d]^{p_1} \ar[dr]_{p_0} \ddrtwocell\omit{^\phi} \xtwocell[ddl]{}\omit{^\;\;\;\;\phi^{-1}} \\
& {\rm src}~\sigma \ar[dr]_{s_1} & {\rm src}~\rho \ar[d]^{r_0}
& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
&& {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s}
}
\endxy
\\
\hspace{5cm}=
\xy/50pt/
\xymatrix{
& X \ar@/dl15pt/[ddr]_<>(0.5){w_s}
\ar@/ur20pt/[dddrr]^{u} \xtwocell[ddr]{}\omit{^{\theta_s}}\\
& \\
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^\sigma}\\
& {\rm src}~r = {\rm trg}~s && {\rm src}~s \ar[ll]^{s}
}
\endxy
\end{array}
$$
\<close>
proof -
have "\<sigma>.composite_cell (p\<^sub>0 \<star> w') (\<theta>' \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<beta>\<^sub>s =
(\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<beta>\<^sub>s_def \<theta>\<^sub>r'_def whisker_left comp_assoc by simp
also have "... = \<theta>\<^sub>r \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<gamma>\<^sub>r by simp
also have "... = ((s \<star> \<theta>) \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \<cdot> \<a>[s, s\<^sub>0, ?w\<^sub>s] \<cdot> (\<sigma> \<star> ?w\<^sub>s) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using \<theta>\<^sub>r_def comp_assoc by simp
also have "... = (s \<star> \<theta>) \<cdot> \<sigma>.composite_cell (p\<^sub>0 \<star> w) \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] =
\<sigma> \<star> p\<^sub>0 \<star> w"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1) comp_assoc_assoc' by simp
text \<open>Here the fact that \<open>\<phi>\<close> is a retraction is used.\<close>
moreover have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = cod \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_arr_inv' whisker_right [of w r\<^sub>0s\<^sub>1.\<phi> "inv r\<^sub>0s\<^sub>1.\<phi>"]
by simp
moreover have "\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] = dom (\<sigma> \<star> p\<^sub>0 \<star> w)"
using r\<^sub>0s\<^sub>1.base_simps(2) hseq_char comp_assoc_assoc' by auto
moreover have "hseq (inv r\<^sub>0s\<^sub>1.\<phi>) w"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2)
by (intro hseqI, auto)
moreover have "hseq \<sigma> (p\<^sub>0 \<star> w)"
by (intro hseqI, auto)
ultimately show ?thesis
using comp_arr_dom comp_cod_arr by simp
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<sigma>.composite_cell (p\<^sub>0 \<star> w) (\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
using \<theta>\<^sub>s_def whisker_left
by (metis \<sigma>.ide_base \<theta>\<^sub>s arrI comp_assoc)
finally show ?thesis by simp
qed
hence 2: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>"
using \<sigma>.T2 [of ?w\<^sub>s ?w\<^sub>s' \<theta>\<^sub>s u \<theta>\<^sub>s' \<beta>\<^sub>s] w\<^sub>s w\<^sub>s' \<theta>\<^sub>s \<theta>\<^sub>s' \<beta>\<^sub>s
by (metis \<theta>\<^sub>s'_def \<theta>\<^sub>s_def)
obtain \<gamma>\<^sub>s where \<gamma>\<^sub>s: "\<guillemotleft>\<gamma>\<^sub>s : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>\<^sub>s) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>\<^sub>s"
using 2 by blast
have eq\<^sub>t: "(s\<^sub>1 \<star> \<gamma>\<^sub>s) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
text \<open>
$$
\xy/78pt/
\xymatrix{
& X \ar[d]^{w'} \ar@/ul15pt/[ddl]_{w_r} \xtwocell[ddl]{}\omit{^{\gamma_r}} \\
& {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\
{\rm src}~\rho \ar[dr]^{r_0}
&& {\rm src}~\sigma \ar[dl]_{s_1} \\
& {\rm src}~r = {\rm trg}~s &
}
\endxy
\qquad = \qquad
\xy/78pt/
\xymatrix{
& X \ar[d]^{w} \ar@/ur15pt/[ddr]^{w_s'} \xtwocell[ddr]{}\omit{^{\gamma_s}} \\
& {\rm src}~\phi \ar[dl]_{p_1} \ar[dr]^{p_0} \ddtwocell\omit{^\phi} \\
{\rm src}~\rho \ar[dr]^{r_0}
&& {\rm src}~\sigma \ar[dl]_{s_1} \\
& {\rm src}~r = {\rm trg}~s &
}
\endxy
$$
\<close>
proof -
have "(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) =
\<beta>\<^sub>s \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "\<beta>\<^sub>s \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<beta>\<^sub>s_def comp_assoc by metis
also have "... = (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
proof -
have "(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
r\<^sub>0 \<star> \<gamma>\<^sub>r"
proof -
have "(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w]) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] =
(r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> ((inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_assoc_assoc' comp_cod_arr by simp
(* Used here that \<phi> is a section. *)
also have "... = (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_inv_arr' \<rho>.T0.antipar(1)
whisker_right [of w "inv r\<^sub>0s\<^sub>1.\<phi>" r\<^sub>0s\<^sub>1.\<phi>] comp_cod_arr
by simp
also have "... = r\<^sub>0 \<star> \<gamma>\<^sub>r"
proof -
have "hseq r\<^sub>0 \<gamma>\<^sub>r"
using \<beta>\<^sub>s \<beta>\<^sub>s_def by blast
thus ?thesis
using comp_assoc_assoc' comp_arr_dom
by (metis (no_types) \<gamma>\<^sub>r \<rho>.ide_leg0 comp_assoc_assoc'(1) hcomp_simps(3)
hseq_char ide_char in_homE r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps w w\<^sub>r)
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
using comp_assoc by presburger
also have "... = (s\<^sub>1 \<star> ?w\<^sub>s') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r)"
proof -
have "(s\<^sub>1 \<star> ?w\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] = \<a>[s\<^sub>1, p\<^sub>0, w']"
using comp_cod_arr by simp
thus ?thesis
using comp_assoc by metis
qed
finally show ?thesis by simp
qed
also have "... = (s\<^sub>1 \<star> \<gamma>\<^sub>s) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>\<^sub>s by simp
finally show ?thesis by simp
qed
have 3: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
using w w' w\<^sub>s' w\<^sub>r \<gamma>\<^sub>r \<gamma>\<^sub>s eq\<^sub>t
r\<^sub>0s\<^sub>1.\<phi>_biuniversal_prop(2) [of ?w\<^sub>s' ?w\<^sub>r w w' \<gamma>\<^sub>s "p\<^sub>0 \<star> w'" \<gamma>\<^sub>r]
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
using 3 by blast
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
proof -
have "\<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>) = (\<theta>\<^sub>s' \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w']) \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
using \<theta>\<^sub>s'_def comp_arr_dom comp_assoc comp_assoc_assoc'(2) by auto
also have "... = (\<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>)) \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
using assoc_naturality [of s\<^sub>0 p\<^sub>0 \<gamma>] comp_assoc
by (metis \<gamma> \<gamma>\<^sub>r \<sigma>.leg0_simps(4-5) r\<^sub>0s\<^sub>1.leg0_simps(4-5)
r\<^sub>0s\<^sub>1.leg1_simps(3) hseqE in_homE leg0_simps(2))
also have "... = \<theta>\<^sub>s \<cdot> \<a>[s\<^sub>0, p\<^sub>0, w]"
by (metis \<gamma> \<gamma>\<^sub>s arrI comp_ide_arr w\<^sub>s')
also have "... = \<theta>"
using \<theta>\<^sub>s_def comp_assoc comp_arr_dom comp_assoc_assoc' by simp
finally show ?thesis by simp
qed
moreover have "\<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>"
proof -
have "\<beta> = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta>\<^sub>r \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<beta>\<^sub>r \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w] =
(\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> \<beta> \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
using \<beta>\<^sub>r_def comp_assoc by simp
also have "... = \<beta>"
using comp_arr_dom comp_cod_arr
by (metis \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 comp_assoc_assoc'(2) hseqE ideD(1)
uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(4-5) leg1_simps(2) w w' w\<^sub>r w\<^sub>r')
finally show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> (r\<^sub>1 \<star> \<gamma>\<^sub>r) \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w]"
using \<gamma>\<^sub>r by simp
also have "... = \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>)"
using assoc_naturality [of r\<^sub>1 p\<^sub>1 \<gamma>]
by (metis \<gamma> \<gamma>\<^sub>r \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE
ide_char in_homE leg1_simps(2))
also have "... = (\<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>[r\<^sub>1, p\<^sub>1, w']) \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>)"
using comp_assoc by presburger
also have "... = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>"
using comp_cod_arr
by (metis \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 calculation comp_assoc_assoc'(2) comp_ide_arr
hseqE ideD(1) ide_cod local.uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) local.uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(5)
w' w\<^sub>r')
finally show ?thesis by simp
qed
ultimately show "\<exists>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma> \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>)"
using \<gamma> by blast
qed
moreover have "\<And>\<gamma>'. \<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>')
\<Longrightarrow> \<gamma>' = \<gamma>"
proof -
fix \<gamma>'
assume \<gamma>': "\<guillemotleft>\<gamma>' : w \<Rightarrow> w'\<guillemotright> \<and> \<beta> = (r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>' \<and> \<theta> = \<theta>' \<cdot> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>')"
show "\<gamma>' = \<gamma>"
proof -
let ?P\<^sub>r = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright> \<and> \<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>r = r\<^sub>1 \<star> \<gamma>"
let ?P\<^sub>s = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright> \<and> \<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> \<gamma>) \<and> \<beta>\<^sub>s = s\<^sub>1 \<star> \<gamma>"
let ?\<gamma>\<^sub>r' = "p\<^sub>1 \<star> \<gamma>'"
let ?\<gamma>\<^sub>s' = "p\<^sub>0 \<star> \<gamma>'"
let ?P\<^sub>t = "\<lambda>\<gamma>. \<guillemotleft>\<gamma> : w \<Rightarrow> w'\<guillemotright> \<and> \<gamma>\<^sub>s = (p\<^sub>0 \<star> w') \<cdot> (p\<^sub>0 \<star> \<gamma>) \<and> p\<^sub>1 \<star> \<gamma> = \<gamma>\<^sub>r"
have "hseq p\<^sub>0 \<gamma>'"
proof (intro hseqI)
show "\<guillemotleft>\<gamma>' : src \<theta> \<rightarrow> src p\<^sub>0\<guillemotright>"
using \<gamma>'
by (metis hseqE hseqI' in_hhom_def uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) src_hcomp
src_vcomp leg0_simps(2) leg1_simps(3)
uw\<theta>w'\<theta>'\<beta>.uw\<theta>.\<theta>_simps(1) vseq_implies_hpar(1))
show "\<guillemotleft>p\<^sub>0 : src p\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>" by simp
qed
hence "hseq p\<^sub>1 \<gamma>'"
using hseq_char by simp
have "\<guillemotleft>?\<gamma>\<^sub>r' : ?w\<^sub>r \<Rightarrow> ?w\<^sub>r'\<guillemotright>"
using \<gamma>' by auto
moreover have "\<theta>\<^sub>r = \<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"
proof -
text \<open>
Note that @{term \<theta>\<^sub>r} is the composite of ``everything to the right''
of @{term "\<rho> \<star> ?w\<^sub>r"}, and similarly for @{term \<theta>\<^sub>r'}.
We can factor @{term \<theta>\<^sub>r} as @{term "(s \<star> \<theta>) \<cdot> X w"}, where @{term "X w"}
is a composite of @{term \<sigma>} and @{term \<phi>}. We can similarly factor @{term \<theta>\<^sub>r'}
as @{term "(s \<star> \<theta>') \<cdot> X w'"}.
Then @{term "\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r') = (s \<star> \<theta>') \<cdot> X w' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"},
which equals @{term "(s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> ?\<gamma>\<^sub>r') \<cdot> X w = \<theta>\<^sub>r"}.
\<close>
let ?X = "\<lambda>w. (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
have "\<theta>\<^sub>r' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r') = (s \<star> \<theta>') \<cdot> ?X w' \<cdot> (r\<^sub>0 \<star> ?\<gamma>\<^sub>r')"
using \<theta>\<^sub>r'_def comp_assoc by simp
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> ?X w"
proof -
have "(s \<star> \<theta>') \<cdot> ((s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w']) \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>') =
(s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot>
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>')"
using comp_assoc by presburger
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot>
((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>')) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using assoc'_naturality [of r\<^sub>0 p\<^sub>1 \<gamma>'] comp_assoc
by (metis \<gamma>' \<open>\<guillemotleft>p\<^sub>1 \<star> \<gamma>' : p\<^sub>1 \<star> w \<Rightarrow> p\<^sub>1 \<star> w'\<guillemotright>\<close> \<rho>.T0.antipar(1)
\<rho>.leg0_in_hom(2) r\<^sub>0s\<^sub>1.leg1_simps(4-6)
r\<^sub>0s\<^sub>1.base_simps(2) hcomp_in_vhomE in_homE trg_hcomp)
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>')) \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') = r\<^sub>0s\<^sub>1.\<phi> \<star> \<gamma>'"
using \<gamma>' interchange [of r\<^sub>0s\<^sub>1.\<phi> "r\<^sub>0 \<star> p\<^sub>1" w' \<gamma>'] comp_arr_dom comp_cod_arr
by auto
also have "... = ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using \<gamma>' interchange \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr
by (metis comp_arr_ide r\<^sub>0s\<^sub>1.\<phi>_simps(1,5) seqI'
uw\<theta>w'\<theta>'\<beta>.uw\<theta>.w_in_hom(2) w)
finally have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') =
((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> \<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
((\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc_naturality [of s\<^sub>1 p\<^sub>0 \<gamma>'] comp_assoc
by (metis \<sigma>.leg1_simps(2) \<sigma>.leg1_simps(3,5-6) r\<^sub>0s\<^sub>1.leg0_simps(4-5)
hcomp_in_vhomE hseqE in_homE uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1)
leg0_in_hom(2) leg1_simps(3))
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w'] \<cdot>
((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>') = \<sigma> \<star> p\<^sub>0 \<star> \<gamma>'"
using \<gamma>' interchange [of \<sigma> s\<^sub>1 "p\<^sub>0 \<star> w'" "p\<^sub>0 \<star> \<gamma>'"]
whisker_left \<open>hseq p\<^sub>0 \<gamma>'\<close>comp_arr_dom comp_cod_arr
by (metis \<sigma>.tab_simps(1) \<sigma>.tab_simps(4) hcomp_simps(4) in_homE
r\<^sub>0s\<^sub>1.leg0_simps(5))
also have "... = ((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>') \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
using \<gamma>' interchange [of "s \<star> s\<^sub>0" \<sigma> "p\<^sub>0 \<star> \<gamma>'" "p\<^sub>0 \<star> w"]
whisker_left comp_arr_dom comp_cod_arr \<open>hseq p\<^sub>0 \<gamma>'\<close>
by (metis \<sigma>.tab_simps(1) \<sigma>.tab_simps(5) hcomp_simps(3) in_homE
r\<^sub>0s\<^sub>1.leg0_simps(4))
finally have "(\<sigma> \<star> p\<^sub>0 \<star> w') \<cdot> (s\<^sub>1 \<star> p\<^sub>0 \<star> \<gamma>') =
((s \<star> s\<^sub>0) \<star> p\<^sub>0 \<star> \<gamma>') \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> ((s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w]) \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc_naturality [of s s\<^sub>0 "p\<^sub>0 \<star> \<gamma>'"] \<open>hseq p\<^sub>0 \<gamma>'\<close> by force
also have "... = (s \<star> \<theta>') \<cdot> ((s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')) \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using comp_assoc by presburger
also have "... = (s \<star> \<theta>') \<cdot> ((s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])) \<cdot>
\<a>[s, s\<^sub>0, p\<^sub>0 \<star> w] \<cdot> (\<sigma> \<star> p\<^sub>0 \<star> w) \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot>
(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
proof -
have "(s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') =
(s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
proof -
have "(s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w']) \<cdot> (s \<star> s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>') =
s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] \<cdot> (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')"
proof -
have "seq \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] (s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>')"
proof
(* It seems to be too time-consuming for auto to solve these. *)
show "\<guillemotleft>s\<^sub>0 \<star> p\<^sub>0 \<star> \<gamma>' : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0 \<star> p\<^sub>0 \<star> w'\<guillemotright>"
using \<gamma>'
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w'] : s\<^sub>0 \<star> p\<^sub>0 \<star> w' \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w'\<guillemotright>"
by auto
qed
thus ?thesis
using w w' \<gamma>' whisker_left by simp
qed
also have "... = s \<star> ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> assoc'_naturality [of s\<^sub>0 p\<^sub>0 \<gamma>'] by fastforce
also have "... = (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (s \<star> \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w])"
proof -
have "seq ((s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w]"
proof
(* Same here. *)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[s\<^sub>0, p\<^sub>0, w] : s\<^sub>0 \<star> p\<^sub>0 \<star> w \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w\<guillemotright>"
by auto
show "\<guillemotleft>(s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>' : (s\<^sub>0 \<star> p\<^sub>0) \<star> w \<Rightarrow> (s\<^sub>0 \<star> p\<^sub>0) \<star> w'\<guillemotright>"
using \<gamma>' by (intro hcomp_in_vhom, auto)
qed
thus ?thesis
using w w' \<gamma>' whisker_left by simp
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = (s \<star> \<theta>') \<cdot> (s \<star> (s\<^sub>0 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> ?X w"
using comp_assoc by presburger
finally show ?thesis by simp
qed
also have "... = \<theta>\<^sub>r"
using \<theta>\<^sub>r_def \<gamma>' uw\<theta>w'\<theta>'\<beta>.uw\<theta>.\<theta>_simps(1) whisker_left \<sigma>.ide_base comp_assoc
by simp
finally show ?thesis by simp
qed
moreover have "\<beta>\<^sub>r = r\<^sub>1 \<star> ?\<gamma>\<^sub>r'"
proof -
have "\<beta>\<^sub>r = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> ((r\<^sub>1 \<star> p\<^sub>1) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w]"
using \<beta>\<^sub>r_def \<gamma>' by simp
also have "... = \<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w'] \<cdot> (r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>')"
using \<gamma>' assoc'_naturality
by (metis \<rho>.leg1_simps(5-6) r\<^sub>0s\<^sub>1.leg1_simps(5-6)
hcomp_in_vhomE hseqE in_homE uw\<theta>w'\<theta>'\<beta>.\<beta>_simps(1) leg1_in_hom(2))
also have "... = (\<a>[r\<^sub>1, p\<^sub>1, w'] \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>1, p\<^sub>1, w']) \<cdot> (r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>')"
using comp_assoc by presburger
also have "... = r\<^sub>1 \<star> p\<^sub>1 \<star> \<gamma>'"
using comp_cod_arr
by (metis (no_types, lifting) \<beta>\<^sub>r \<rho>.ide_leg1 r\<^sub>0s\<^sub>1.ide_leg1 arrI calculation
comp_assoc_assoc'(1) comp_ide_arr ide_hcomp hseq_char'
ideD(1) seq_if_composable hcomp_simps(2) leg1_simps(2) w' w\<^sub>r')
finally show ?thesis by simp
qed
ultimately have P\<^sub>r': "?P\<^sub>r ?\<gamma>\<^sub>r'"
by simp
have eq\<^sub>r: "\<gamma>\<^sub>r = ?\<gamma>\<^sub>r'"
using 1 \<gamma>\<^sub>r P\<^sub>r' by blast
have "\<guillemotleft>?\<gamma>\<^sub>s' : ?w\<^sub>s \<Rightarrow> ?w\<^sub>s'\<guillemotright>"
using \<gamma>' by auto
moreover have "\<theta>\<^sub>s = \<theta>\<^sub>s' \<cdot> (s\<^sub>0 \<star> ?\<gamma>\<^sub>s')"
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> \<sigma>.leg0_simps(2,4-5) \<sigma>.leg1_simps(3) \<theta>\<^sub>s'_def \<theta>\<^sub>s_def
assoc'_naturality hseqE in_homE comp_assoc r\<^sub>0s\<^sub>1.leg0_simps(4-5)
r\<^sub>0s\<^sub>1.p\<^sub>0_simps
by metis
moreover have "\<beta>\<^sub>s = s\<^sub>1 \<star> ?\<gamma>\<^sub>s'"
proof -
have "\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> \<gamma>\<^sub>r) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w] =
\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>')) \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using eq\<^sub>r comp_assoc by simp
also have "... = \<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>')) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w'] \<cdot> (r\<^sub>0 \<star> p\<^sub>1 \<star> \<gamma>') = ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') \<cdot> \<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w]"
using \<gamma>' assoc'_naturality \<open>hseq p\<^sub>1 \<gamma>'\<close>
by (metis \<rho>.leg0_simps(2,4-5) \<rho>.leg1_simps(3)
r\<^sub>0s\<^sub>1.leg1_simps(5-6) hseqE in_homE leg1_simps(2))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>')) \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot>
\<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') = r\<^sub>0s\<^sub>1.\<phi> \<star> \<gamma>'"
using \<gamma>' interchange [of r\<^sub>0s\<^sub>1.\<phi> "r\<^sub>0 \<star> p\<^sub>1" w' \<gamma>']
comp_arr_dom comp_cod_arr
by auto
also have "... = ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using \<gamma>' interchange \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr
by (metis in_homE r\<^sub>0s\<^sub>1.\<phi>_simps(1,5))
finally have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w') \<cdot> ((r\<^sub>0 \<star> p\<^sub>1) \<star> \<gamma>') =
((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') \<cdot> (r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (s\<^sub>1 \<star> ?\<gamma>\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w] \<cdot> ((r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot>
\<a>[r\<^sub>0, p\<^sub>1, w]) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)) \<cdot> \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
proof -
have "\<a>[s\<^sub>1, p\<^sub>0, w'] \<cdot> ((s\<^sub>1 \<star> p\<^sub>0) \<star> \<gamma>') = (s\<^sub>1 \<star> ?\<gamma>\<^sub>s') \<cdot> \<a>[s\<^sub>1, p\<^sub>0, w]"
using \<gamma>' assoc_naturality [of s\<^sub>1 p\<^sub>0 \<gamma>'] \<open>hseq p\<^sub>0 \<gamma>'\<close> by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = s\<^sub>1 \<star> ?\<gamma>\<^sub>s'"
proof -
have "\<a>\<^sup>-\<^sup>1[r\<^sub>0, p\<^sub>1, w] \<cdot> \<a>[r\<^sub>0, p\<^sub>1, w] = cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w)"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) \<rho>.T0.antipar(1) comp_assoc_assoc'
by simp
text \<open>Here the fact that \<open>\<phi>\<close> is a retraction is used.\<close>
moreover have "(r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = cod \<a>\<^sup>-\<^sup>1[s\<^sub>1, p\<^sub>0, w]"
using r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) comp_arr_inv'
whisker_right [of w r\<^sub>0s\<^sub>1.\<phi> "inv r\<^sub>0s\<^sub>1.\<phi>"]
by simp
moreover have "cod (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) \<cdot> (inv r\<^sub>0s\<^sub>1.\<phi> \<star> w) = inv r\<^sub>0s\<^sub>1.\<phi> \<star> w"
using \<beta>\<^sub>s_def \<beta>\<^sub>s
by (meson arrI comp_cod_arr seqE)
ultimately show ?thesis
using \<gamma>' \<open>hseq p\<^sub>0 \<gamma>'\<close> comp_arr_dom comp_cod_arr comp_assoc_assoc'
whisker_left [of s\<^sub>1 "p\<^sub>0 \<star> \<gamma>'" "p\<^sub>0 \<star> w"] whisker_left [of p\<^sub>0]
by (metis \<sigma>.ide_leg1 assoc'_simps(1) hseqE ideD(1) in_homE r\<^sub>0s\<^sub>1.ide_leg0
r\<^sub>0s\<^sub>1.p\<^sub>0_simps w w\<^sub>s)
qed
finally show ?thesis
using \<beta>\<^sub>s_def by simp
qed
ultimately have P\<^sub>s': "?P\<^sub>s ?\<gamma>\<^sub>s'"
by simp
have eq\<^sub>s: "\<gamma>\<^sub>s = ?\<gamma>\<^sub>s'"
using 2 \<gamma>\<^sub>s P\<^sub>s' by blast
have "?P\<^sub>t \<gamma>'"
using \<gamma>' comp_cod_arr \<open>\<guillemotleft>p\<^sub>0 \<star> \<gamma>' : p\<^sub>0 \<star> w \<Rightarrow> p\<^sub>0 \<star> w'\<guillemotright>\<close> eq\<^sub>r eq\<^sub>s by auto
thus "\<gamma>' = \<gamma>"
using 3 \<gamma> by blast
qed
qed
ultimately show ?thesis by blast
qed
qed
qed
end
sublocale composite_tabulation_in_maps \<subseteq>
tabulation V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using composite_is_tabulation by simp
sublocale composite_tabulation_in_maps \<subseteq>
tabulation_in_maps V H \<a> \<i> src trg \<open>r \<star> s\<close> tab \<open>s\<^sub>0 \<star> p\<^sub>0\<close> \<open>r\<^sub>1 \<star> p\<^sub>1\<close>
using T0.is_map \<rho>.leg1_is_map \<rho>.T0.antipar(2) composable \<rho>.leg1_is_map \<rho>.T0.antipar
apply unfold_locales
apply simp
apply (intro left_adjoints_compose)
by auto
subsection "The Classifying Category of Maps"
text \<open>
\sloppypar
We intend to show that if \<open>B\<close> is a bicategory of spans, then \<open>B\<close> is biequivalent to
\<open>Span(Maps(B))\<close>, for a specific category \<open>Maps(B)\<close> derived from \<open>B\<close>.
The category \<open>Maps(B)\<close> is constructed in this section as the ``classifying category'' of
maps of \<open>B\<close>, which has the same objects as \<open>B\<close> and which has as 1-cells the isomorphism classes
of maps of \<open>B\<close>. We show that, if \<open>B\<close> is a bicategory of spans, then \<open>Maps(B)\<close> has pullbacks.
\<close>
locale maps_category =
B: bicategory_of_spans
begin
no_notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>B _\<guillemotright>")
notation B.isomorphic (infix "\<cong>\<^sub>B" 50)
notation B.iso_class ("\<lbrakk>_\<rbrakk>\<^sub>B")
text \<open>
I attempted to modularize the construction here, by refactoring ``classifying category''
out as a separate locale, but it ended up causing extra work because to apply it we
first need to obtain the full sub-bicategory of 2-cells between maps, then construct its
classifying category, and then we have to re-prove everything about it, to get rid of
any mention of the sub-bicategory construction. So the construction is being done
here as a ``one-off'' special case construction, with the necessary properties proved
just once.
\<close>
text \<open>
The ``hom-categories'' of \<open>Maps(C)\<close> have as arrows the isomorphism classes of maps of \<open>B\<close>.
\<close>
abbreviation Hom
where "Hom a b \<equiv> {F. \<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B}"
lemma in_HomD:
assumes "F \<in> Hom a b"
shows "F \<noteq> {}"
and "B.is_iso_class F"
and "f \<in> F \<Longrightarrow> B.ide f"
and "f \<in> F \<Longrightarrow> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
and "f \<in> F \<Longrightarrow> B.is_left_adjoint f"
and "f \<in> F \<Longrightarrow> F = \<lbrakk>f\<rbrakk>\<^sub>B"
proof -
show "F \<noteq> {}"
using assms B.ide_in_iso_class B.left_adjoint_is_ide B.iso_class_is_nonempty by auto
show 1: "B.is_iso_class F"
using assms B.is_iso_classI B.left_adjoint_is_ide by fastforce
show "f \<in> F \<Longrightarrow> B.ide f"
using assms 1 B.iso_class_memb_is_ide by blast
obtain f' where f': "\<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f' \<and> F = \<lbrakk>f'\<rbrakk>\<^sub>B"
using assms by auto
show "f \<in> F \<Longrightarrow> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
using assms f' B.iso_class_def B.isomorphic_implies_hpar by auto
show "f \<in> F \<Longrightarrow> B.is_left_adjoint f"
using assms f' B.iso_class_def B.left_adjoint_preserved_by_iso [of f'] by auto
show "f \<in> F \<Longrightarrow> F = \<lbrakk>f\<rbrakk>\<^sub>B"
by (metis B.adjoint_pair_antipar(1) f' B.ide_in_iso_class B.is_iso_classI
B.iso_class_elems_isomorphic B.iso_class_eqI)
qed
definition Comp
where "Comp G F \<equiv> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h)}"
lemma in_CompI [intro]:
assumes "B.is_iso_class F" and "B.is_iso_class G"
and "f \<in> F" and "g \<in> G" and "g \<star> f \<cong>\<^sub>B h"
shows "h \<in> Comp G F"
unfolding Comp_def
using assms by auto
lemma in_CompE [elim]:
assumes "h \<in> Comp G F"
and "\<And>f g. \<lbrakk> B.is_iso_class F; B.is_iso_class G; f \<in> F; g \<in> G; g \<star> f \<cong>\<^sub>B h \<rbrakk> \<Longrightarrow> T"
shows T
using assms Comp_def by auto
lemma is_iso_class_Comp:
assumes "Comp G F \<noteq> {}"
shows "B.is_iso_class (Comp G F)"
proof -
obtain h where h: "h \<in> Comp G F"
using assms by auto
have ide_h: "B.ide h"
using h Comp_def B.isomorphic_implies_hpar(2) by auto
obtain f g where fg: "B.is_iso_class F \<and> B.is_iso_class G \<and> f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using h Comp_def by auto
have "Comp G F = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<and> B.ide (g \<star> f)"
proof (intro conjI)
show "B.ide (g \<star> f)"
using fg B.iso_class_memb_is_ide B.isomorphic_implies_ide(1) by auto
show "Comp G F = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
show "\<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<subseteq> Comp G F"
unfolding Comp_def B.iso_class_def
using fg by auto
show "Comp G F \<subseteq> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
fix h'
assume h': "h' \<in> Comp G F"
obtain f' g' where f'g': "f' \<in> F \<and> g' \<in> G \<and> g' \<star> f' \<cong>\<^sub>B h'"
using h' Comp_def by auto
have 1: "f' \<cong>\<^sub>B f \<and> g' \<cong>\<^sub>B g"
using f'g' fg B.iso_class_elems_isomorphic by auto
moreover have "B.ide f \<and> B.ide f' \<and> B.ide g \<and> B.ide g'"
using 1 B.isomorphic_implies_hpar by auto
ultimately have "g' \<star> f' \<cong>\<^sub>B g \<star> f"
using f'g' fg B.hcomp_isomorphic_ide B.hcomp_ide_isomorphic
B.isomorphic_transitive B.isomorphic_implies_hpar
by (meson B.hseqE B.ideD(1))
hence "h' \<cong>\<^sub>B g \<star> f"
using f'g' B.isomorphic_symmetric B.isomorphic_transitive by blast
thus "h' \<in> B.iso_class (g \<star> f)"
using B.iso_class_def B.isomorphic_symmetric by simp
qed
qed
qed
thus ?thesis
using assms B.is_iso_class_def B.ide_in_iso_class by auto
qed
lemma Comp_is_extensional:
assumes "Comp G F \<noteq> {}"
shows "B.is_iso_class F" and "B.is_iso_class G"
and "F \<noteq> {}" and "G \<noteq> {}"
using assms Comp_def by auto
lemma Comp_eqI [intro]:
assumes "h \<in> Comp G F" and "h' \<in> Comp G' F'" and "h \<cong>\<^sub>B h'"
shows "Comp G F = Comp G' F'"
proof -
obtain f g where fg: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms comp_def by auto
obtain f' g' where f'g': "f' \<in> F' \<and> g' \<in> G' \<and> g' \<star> f' \<cong>\<^sub>B h'"
using assms by auto
have "h \<in> Comp G F \<inter> Comp G' F'"
by (meson IntI assms in_CompE in_CompI B.isomorphic_symmetric B.isomorphic_transitive)
hence "Comp G F \<inter> Comp G' F' \<noteq> {}"
by auto
thus ?thesis
using assms is_iso_class_Comp
by (metis empty_iff B.iso_class_eq)
qed
lemma Comp_eq_iso_class_memb:
assumes "h \<in> Comp G F"
shows "Comp G F = \<lbrakk>h\<rbrakk>\<^sub>B"
proof
show "Comp G F \<subseteq> \<lbrakk>h\<rbrakk>\<^sub>B"
proof
fix h'
assume h': "h' \<in> Comp G F"
obtain f g where fg: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms by auto
obtain f' g' where f'g': "f' \<in> F \<and> g' \<in> G \<and> g' \<star> f' \<cong>\<^sub>B h'"
using h' by auto
have "f \<cong>\<^sub>B f' \<and> g \<cong>\<^sub>B g'"
using assms fg f'g' in_HomD(6) B.iso_class_elems_isomorphic by auto
moreover have "B.ide f \<and> B.ide f' \<and> B.ide g \<and> B.ide g'"
using assms fg f'g' in_HomD [of F] in_HomD [of G]
by (meson calculation B.isomorphic_implies_ide(1) B.isomorphic_implies_ide(2))
moreover have "src g = trg f \<and> src g = trg f' \<and> src g' = trg f \<and> src g' = trg f'"
using fg f'g'
by (metis B.seq_if_composable calculation(1) B.ideD(1)
B.isomorphic_implies_hpar(1) B.isomorphic_implies_hpar(3) B.not_arr_null)
ultimately have "g \<star> f \<cong>\<^sub>B g' \<star> f'"
using fg f'g' B.hcomp_ide_isomorphic B.hcomp_isomorphic_ide B.isomorphic_transitive
by metis
thus "h' \<in> \<lbrakk>h\<rbrakk>\<^sub>B"
using fg f'g' B.isomorphic_symmetric B.isomorphic_transitive B.iso_class_def [of h]
by blast
qed
show "\<lbrakk>h\<rbrakk>\<^sub>B \<subseteq> Comp G F"
proof (unfold B.iso_class_def Comp_def)
obtain f g where 1: "f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h"
using assms in_HomD Comp_def
by (meson in_CompE)
show "{h'. B.isomorphic h h'} \<subseteq>
{h. B.is_iso_class F \<and> B.is_iso_class G \<and> (\<exists>f g. f \<in> F \<and> g \<in> G \<and> g \<star> f \<cong>\<^sub>B h)}"
using assms 1 B.isomorphic_transitive by blast
qed
qed
interpretation concrete_category \<open>Collect B.obj\<close> Hom B.iso_class \<open>\<lambda>_ _ _. Comp\<close>
proof
show "\<And>a. a \<in> Collect B.obj \<Longrightarrow> \<lbrakk>a\<rbrakk>\<^sub>B \<in> Hom a a"
by (metis (mono_tags, lifting) B.ide_in_hom(1) mem_Collect_eq B.objE
B.obj_is_self_adjoint(1))
show "\<And>a b c F G. \<lbrakk> a \<in> Collect B.obj; b \<in> Collect B.obj; c \<in> Collect B.obj;
F \<in> Hom a b; G \<in> Hom b c \<rbrakk> \<Longrightarrow> Comp G F \<in> Hom a c"
proof -
fix a b c F G
assume a: "a \<in> Collect B.obj" and b: "b \<in> Collect B.obj" and c: "c \<in> Collect B.obj"
and F: "F \<in> Hom a b" and G: "G \<in> Hom b c"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by blast
obtain g where g: "\<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> B.is_left_adjoint g \<and> G = \<lbrakk>g\<rbrakk>\<^sub>B"
using G by blast
have "{h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)} =
\<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
show "{h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)}
\<subseteq> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof
fix h
assume "h \<in> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)}"
hence h: "B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)"
by simp
show "h \<in> \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof -
obtain f' g' where f'g': "g' \<in> G \<and> f' \<in> F \<and> g' \<star> f' \<cong>\<^sub>B h"
using h by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : f \<Rightarrow>\<^sub>B f'\<guillemotright> \<and> B.iso \<phi>"
using f f'g' F B.iso_class_def by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : g \<Rightarrow>\<^sub>B g'\<guillemotright> \<and> B.iso \<psi>"
using g f'g' G B.iso_class_def by auto
have 1: "\<guillemotleft>\<psi> \<star> \<phi> : g \<star> f \<Rightarrow>\<^sub>B g' \<star> f'\<guillemotright>"
using f g \<phi> \<psi> by auto
moreover have "B.iso (\<psi> \<star> \<phi>)"
using f g \<phi> \<psi> 1 B.iso_hcomp by auto
ultimately show ?thesis
using f'g' B.iso_class_def B.isomorphic_def by auto
qed
qed
show "\<lbrakk>g \<star> f\<rbrakk>\<^sub>B \<subseteq> {h. B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B h)}"
using f g B.iso_class_def B.isomorphic_reflexive B.left_adjoint_is_ide B.is_iso_classI
by blast
qed
hence 1: "\<And>gf. gf \<in> B.iso_class (g \<star> f) \<Longrightarrow>
B.is_iso_class F \<and> B.is_iso_class G \<and>
(\<exists>f g. f \<in> F \<and> \<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> g \<in> G \<and> \<guillemotleft>g : b \<rightarrow>\<^sub>B c\<guillemotright> \<and> g \<star> f \<cong>\<^sub>B gf)"
by blast
show "Comp G F \<in> Hom a c"
proof -
have gf: "B.is_left_adjoint (g \<star> f)"
by (meson f g B.hseqE B.hseqI B.left_adjoints_compose)
obtain gf' where gf': "B.adjoint_pair (g \<star> f) gf'"
using gf by blast
hence "\<lbrakk>g \<star> f\<rbrakk>\<^sub>B = Comp G F"
using 1 Comp_eq_iso_class_memb B.ide_in_iso_class B.left_adjoint_is_ide by blast
thus ?thesis
using f g gf' by blast
qed
qed
show "\<And>a b F. \<lbrakk> a \<in> Collect B.obj; F \<in> Hom a b \<rbrakk> \<Longrightarrow> Comp F \<lbrakk>a\<rbrakk>\<^sub>B = F"
proof -
fix a b F
assume a: "a \<in> Collect B.obj"
assume F: "F \<in> Hom a b"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by auto
have *: "\<And>f'. f' \<in> F \<Longrightarrow> \<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f' \<and> f \<cong>\<^sub>B f'"
using f B.iso_class_def by force
show "Comp F \<lbrakk>a\<rbrakk>\<^sub>B = F"
proof
show "Comp F \<lbrakk>a\<rbrakk>\<^sub>B \<subseteq> F"
proof
fix h
assume "h \<in> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
hence h: "\<exists>f' a'. f' \<in> F \<and> a' \<in> \<lbrakk>a\<rbrakk>\<^sub>B \<and> f' \<star> a' \<cong>\<^sub>B h"
unfolding Comp_def by auto
obtain f' a' where f'a': "f' \<in> F \<and> a' \<in> \<lbrakk>a\<rbrakk>\<^sub>B \<and> f' \<star> a' \<cong>\<^sub>B h"
using h by auto
have "B.isomorphic f h"
proof -
have "B.isomorphic f (f \<star> a)"
using f B.iso_runit' [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "f \<star> a \<cong>\<^sub>B f' \<star> a"
using f f'a' B.iso_class_def B.hcomp_isomorphic_ide
apply (elim conjE B.in_hhomE) by auto
also have "f' \<star> a \<cong>\<^sub>B f' \<star> a'"
using f'a' * [of f'] B.iso_class_def B.hcomp_ide_isomorphic by auto
also have "f' \<star> a' \<cong>\<^sub>B h"
using f'a' by simp
finally show ?thesis by blast
qed
thus "h \<in> F"
using f B.iso_class_def by simp
qed
show "F \<subseteq> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
proof
fix h
assume h: "h \<in> F"
have "f \<in> F"
using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
moreover have "a \<in> B.iso_class a"
using a B.iso_class_def B.isomorphic_reflexive by auto
moreover have "f \<star> a \<cong>\<^sub>B h"
proof -
have "f \<star> a \<cong>\<^sub>B f"
using f B.iso_runit [of f] B.isomorphic_def B.left_adjoint_is_ide by blast
also have "f \<cong>\<^sub>B h"
using h * by simp
finally show ?thesis by blast
qed
ultimately show "h \<in> Comp F \<lbrakk>a\<rbrakk>\<^sub>B"
unfolding Comp_def
using a f F B.is_iso_classI B.left_adjoint_is_ide by auto
qed
qed
qed
show "\<And>a b F. \<lbrakk> b \<in> Collect B.obj; F \<in> Hom a b \<rbrakk> \<Longrightarrow> Comp \<lbrakk>b\<rbrakk>\<^sub>B F = F"
proof -
fix a b F
assume b: "b \<in> Collect B.obj"
assume F: "F \<in> Hom a b"
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.is_left_adjoint f \<and> F = \<lbrakk>f\<rbrakk>\<^sub>B"
using F by auto
have *: "\<And>f'. f' \<in> F \<Longrightarrow> \<guillemotleft>f' : a \<rightarrow>\<^sub>B b\<guillemotright> \<and> B.ide f' \<and> f \<cong>\<^sub>B f'"
using f B.iso_class_def by force
show "Comp \<lbrakk>b\<rbrakk>\<^sub>B F = F"
proof
show "Comp \<lbrakk>b\<rbrakk>\<^sub>B F \<subseteq> F"
proof
fix h
assume "h \<in> Comp \<lbrakk>b\<rbrakk>\<^sub>B F"
hence h: "\<exists>b' f'. b' \<in> \<lbrakk>b\<rbrakk>\<^sub>B \<and> f' \<in> F \<and> b' \<star> f' \<cong>\<^sub>B h"
unfolding Comp_def by auto
obtain b' f' where b'f': "b' \<in> \<lbrakk>b\<rbrakk>\<^sub>B \<and> f' \<in> F \<and> b' \<star> f' \<cong>\<^sub>B h"
using h by auto
have "f \<cong>\<^sub>B h"
proof -
have "f \<cong>\<^sub>B b \<star> f"
using f B.iso_lunit' [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "... \<cong>\<^sub>B b \<star> f'"
using f b'f' B.iso_class_def B.hcomp_ide_isomorphic
by (elim conjE B.in_hhomE, auto)
also have "... \<cong>\<^sub>B b' \<star> f'"
using b'f' * [of f'] B.iso_class_def B.hcomp_isomorphic_ide by auto
also have "... \<cong>\<^sub>B h"
using b'f' by simp
finally show ?thesis by blast
qed
thus "h \<in> F"
using f B.iso_class_def by simp
qed
show "F \<subseteq> Comp \<lbrakk>b\<rbrakk>\<^sub>B F"
proof
fix h
assume h: "h \<in> F"
have "f \<in> F"
using f B.iso_class_def B.right_adjoint_determines_left_up_to_iso by auto
moreover have "b \<in> B.iso_class b"
using b B.iso_class_def B.isomorphic_reflexive by auto
moreover have "b \<star> f \<cong>\<^sub>B h"
proof -
have "b \<star> f \<cong>\<^sub>B f"
using f B.iso_lunit [of f] B.isomorphic_def B.left_adjoint_is_ide
by blast
also have "f \<cong>\<^sub>B h"
using h * by simp
finally show ?thesis by blast
qed
ultimately show "h \<in> Comp \<lbrakk>b\<rbrakk>\<^sub>B F"
unfolding Comp_def
using b f F B.is_iso_classI B.left_adjoint_is_ide by auto
qed
qed
qed
show "\<And>a b c d F G H.
\<lbrakk> a \<in> Collect B.obj; b \<in> Collect B.obj; c \<in> Collect B.obj; d \<in> Collect B.obj;
F \<in> Hom a b; G \<in> Hom b c; H \<in> Hom c d \<rbrakk> \<Longrightarrow>
Comp H (Comp G F) = Comp (Comp H G) F"
proof -
fix a b c d F G H
assume F: "F \<in> Hom a b" and G: "G \<in> Hom b c" and H: "H \<in> Hom c d"
show "Comp H (Comp G F) = Comp (Comp H G) F"
proof
show "Comp H (Comp G F) \<subseteq> Comp (Comp H G) F"
proof
fix x
assume x: "x \<in> Comp H (Comp G F)"
obtain f g h gf
where 1: "B.is_iso_class F \<and> B.is_iso_class G \<and> B.is_iso_class H \<and>
h \<in> H \<and> g \<in> G \<and> f \<in> F \<and> gf \<in> Comp G F \<and> g \<star> f \<cong>\<^sub>B gf \<and> h \<star> gf \<cong>\<^sub>B x"
using x unfolding Comp_def by blast
have hgf: "B.ide f \<and> B.ide g \<and> B.ide h"
using 1 F G H B.isomorphic_implies_hpar in_HomD B.left_adjoint_is_ide
by (metis (mono_tags, lifting))
have "h \<star> g \<star> f \<cong>\<^sub>B x"
proof -
have "h \<star> g \<star> f \<cong>\<^sub>B h \<star> gf"
using 1 hgf B.hcomp_ide_isomorphic
by (metis (full_types) B.isomorphic_implies_hpar(1) B.isomorphic_reflexive
B.isomorphic_symmetric B.seq_if_composable)
also have "h \<star> gf \<cong>\<^sub>B x"
using 1 by simp
finally show ?thesis by blast
qed
moreover have "(h \<star> g) \<star> f \<cong>\<^sub>B h \<star> g \<star> f"
using 1 hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def
by (metis B.hseq_char B.ideD(1-3) B.isomorphic_implies_hpar(1)
B.trg_hcomp calculation)
moreover have hg: "\<guillemotleft>h \<star> g : b \<rightarrow>\<^sub>B d\<guillemotright>"
using G H 1 in_HomD(4) by blast
moreover have "h \<star> g \<in> Comp H G"
unfolding Comp_def
using 1 hgf F G H in_HomD [of F a b] in_HomD [of G b c] in_HomD [of H c d]
B.isomorphic_reflexive B.hcomp_ide_isomorphic B.hseqI'
by (metis (no_types, lifting) B.hseqE B.hseqI mem_Collect_eq)
ultimately show "x \<in> Comp (Comp H G) F"
using 1 F G H hgf B.is_iso_class_def is_iso_class_Comp [of H G]
B.isomorphic_reflexive [of "h \<star> g"]
apply (intro in_CompI)
apply auto[7]
apply blast
apply simp
by (meson B.isomorphic_symmetric B.isomorphic_transitive)
qed
show "Comp (Comp H G) F \<subseteq> Comp H (Comp G F)"
proof
fix x
assume x: "x \<in> Comp (Comp H G) F"
obtain f g h hg
where 1: "B.is_iso_class F \<and> B.is_iso_class G \<and> B.is_iso_class H \<and>
h \<in> H \<and> g \<in> G \<and> f \<in> F \<and> hg \<in> Comp H G \<and> h \<star> g \<cong>\<^sub>B hg \<and> hg \<star> f \<cong>\<^sub>B x"
using x unfolding Comp_def by blast
have hgf: "B.ide f \<and> B.ide g \<and> B.ide h \<and> src h = trg g \<and> src g = trg f"
using 1 F G H in_HomD B.left_adjoint_is_ide
by (metis (no_types, lifting) B.hseq_char' B.ideD(1) B.ide_dom
B.in_homE B.isomorphic_def B.isomorphic_symmetric B.seqI'
B.seq_if_composable B.src_dom B.src_hcomp B.vseq_implies_hpar(1))
have 2: "(h \<star> g) \<star> f \<cong>\<^sub>B x"
proof -
have "(h \<star> g) \<star> f \<cong>\<^sub>B hg \<star> f"
using 1 F G H hgf
by (simp add: B.hcomp_isomorphic_ide)
also have "hg \<star> f \<cong>\<^sub>B x"
using 1 by simp
finally show ?thesis by blast
qed
moreover have "(h \<star> g) \<star> f \<cong>\<^sub>B h \<star> g \<star> f"
using hgf B.iso_assoc B.assoc_in_hom B.isomorphic_def by auto
moreover have "g \<star> f \<in> Comp G F"
using 1 F G hgf B.isomorphic_reflexive by blast
ultimately show "x \<in> Comp H (Comp G F)"
using 1 hgf is_iso_class_Comp [of G F] B.isomorphic_reflexive [of "g \<star> f"]
apply (intro in_CompI)
apply auto[6]
apply simp
apply auto[1]
by (meson B.isomorphic_symmetric B.isomorphic_transitive)
qed
qed
qed
qed
lemma is_concrete_category:
shows "concrete_category (Collect B.obj) Hom B.iso_class (\<lambda>_ _ _. Comp)"
..
sublocale concrete_category \<open>Collect B.obj\<close> Hom B.iso_class \<open>\<lambda>_ _ _. Comp\<close>
using is_concrete_category by simp
abbreviation comp (infixr "\<odot>" 55)
where "comp \<equiv> COMP"
notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
lemma Map_memb_in_hhom:
assumes "\<guillemotleft>F : A \<rightarrow> B\<guillemotright>" and "f \<in> Map F"
shows "\<guillemotleft>f : Dom A \<rightarrow>\<^sub>B Dom B\<guillemotright>"
proof -
have "\<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright>"
using assms arr_char [of F] in_HomD [of "Map F" "Dom F" "Cod F"] by blast
moreover have "Dom F = Dom A"
using assms by auto
moreover have "Cod F = Dom B"
using assms by auto
ultimately show ?thesis by simp
qed
lemma MkArr_in_hom':
assumes "B.is_left_adjoint f" and "\<guillemotleft>f : a \<rightarrow>\<^sub>B b\<guillemotright>"
shows "\<guillemotleft>MkArr a b \<lbrakk>f\<rbrakk>\<^sub>B : MkIde a \<rightarrow> MkIde b\<guillemotright>"
using assms MkArr_in_hom by blast
text \<open>
The isomorphisms in \<open>Maps(B)\<close> are the isomorphism classes of equivalence maps in \<open>B\<close>.
\<close>
lemma iso_char:
shows "iso F \<longleftrightarrow> arr F \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
proof
assume F: "iso F"
have "\<And>f. f \<in> Map F \<Longrightarrow> B.equivalence_map f"
proof -
fix f
assume f: "f \<in> Map F"
obtain G where G: "inverse_arrows F G"
using F by auto
obtain g where g: "g \<in> Map G"
using G arr_char [of G] in_HomD(1) by blast
have f: "f \<in> Map F \<and> \<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright> \<and> B.ide f \<and> B.is_left_adjoint f"
by (metis (mono_tags, lifting) F iso_is_arr is_concrete_category
concrete_category.arr_char f in_HomD(2,4-5) B.iso_class_memb_is_ide)
have g: "g \<in> Map G \<and> \<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright> \<and> B.ide g \<and> B.is_left_adjoint g"
by (metis (no_types, lifting) F G Cod_cod Cod_dom arr_inv cod_inv dom_inv
inverse_unique iso_is_arr is_concrete_category concrete_category.Map_in_Hom
g in_HomD(2,4-5) B.iso_class_memb_is_ide)
have "src (g \<star> f) \<cong>\<^sub>B g \<star> f"
proof -
have "g \<star> f \<in> Map (G \<odot> F)"
proof -
have 1: "f \<in> Map F \<and> g \<in> Map G \<and> g \<star> f \<cong>\<^sub>B g \<star> f"
using F G f g B.isomorphic_reflexive by force
have 2: "Dom G = Cod F \<and> Cod G = Dom F"
using F G arr_char
by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv
cod_char comp_inv_arr dom_inv inverse_unique
iso_is_arr is_concrete_category concrete_category.Cod_comp)
have "g \<star> f \<in> Comp (Map G) (Map F)"
using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2)
apply (intro in_CompI, auto)
proof -
show "B.is_iso_class (Map G)"
using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast
qed
thus ?thesis
using F G f g comp_char [of G F] by auto
qed
moreover have "Dom F \<in> Map (G \<odot> F)"
by (metis (no_types, lifting) F G Map_dom comp_inv_arr iso_is_arr
g B.ide_in_iso_class B.in_hhomE B.objE)
moreover have "Map (G \<odot> F) = \<lbrakk>Dom F\<rbrakk>\<^sub>B"
by (simp add: F G comp_inv_arr iso_is_arr)
moreover have "Dom F = src (g \<star> f)"
using f g by fastforce
ultimately show ?thesis
using f g B.iso_class_elems_isomorphic B.is_iso_classI
by (metis B.hseqI B.ide_src)
qed
moreover have "f \<star> g \<cong>\<^sub>B trg f"
proof -
have "f \<star> g \<in> Map (F \<odot> G)"
proof -
have 1: "f \<in> Map F \<and> g \<in> Map G \<and> f \<star> g \<cong>\<^sub>B f \<star> g"
using F G f g B.isomorphic_reflexive by force
have 2: "Dom G = Cod F \<and> Cod G = Dom F"
using F G arr_char
by (metis (no_types, lifting) Cod.simps(1) Cod_dom arr_inv
cod_char comp_inv_arr dom_inv inverse_unique
iso_is_arr is_concrete_category concrete_category.Cod_comp)
hence "f \<star> g \<in> Comp (Map F) (Map G)"
using 1 F iso_is_arr Map_in_Hom [of F] in_HomD(2)
apply (intro in_CompI, auto)
proof -
show "B.is_iso_class (Map G)"
using G iso_is_arr Map_in_Hom [of G] in_HomD(2) [of "Map G"] by blast
qed
thus ?thesis
using F G f g comp_char [of F G] by auto
qed
moreover have "Cod F \<in> Map (F \<odot> G)"
by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv
inverse_unique iso_is_arr g B.ide_in_iso_class B.in_hhomE B.src.preserves_ide)
moreover have "Map (F \<odot> G) = \<lbrakk>Cod F\<rbrakk>\<^sub>B"
by (metis (no_types, lifting) F G Map_cod comp_arr_inv dom_inv
inverse_unique iso_is_arr)
moreover have "Cod F = trg (f \<star> g)"
using f g by fastforce
ultimately show ?thesis
using B.iso_class_elems_isomorphic
by (metis f g B.is_iso_classI B.in_hhomE B.src.preserves_ide)
qed
ultimately show "B.equivalence_map f"
using f g B.equivalence_mapI B.quasi_inversesI [of f g] by fastforce
qed
thus "arr F \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
using F by blast
next
assume F: "arr F \<and> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
show "iso F"
proof -
obtain f where f: "f \<in> Map F"
using F arr_char in_HomD(1) by blast
have f_in_hhom: "\<guillemotleft>f : Dom F \<rightarrow>\<^sub>B Cod F\<guillemotright>"
using f F arr_char in_HomD(4) [of "Map F" "Dom F" "Cod F" f] by simp
have "Map F = B.iso_class f"
using f F arr_char in_HomD(6) [of "Map F" "Dom F" "Cod F" f] by simp
obtain g \<eta> \<epsilon>' where \<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using f F B.equivalence_map_def by auto
interpret \<epsilon>': equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using \<epsilon>' by auto
obtain \<epsilon> where \<epsilon>: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using f F \<epsilon>'.ide_right \<epsilon>'.antipar \<epsilon>'.unit_in_hom \<epsilon>'.unit_is_iso B.equivalence_map_def
B.equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
by auto
interpret \<epsilon>: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using \<epsilon> by simp
have g_in_hhom: "\<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright>"
using \<epsilon>.ide_right \<epsilon>.antipar f_in_hhom by auto
have fg: "B.quasi_inverses f g"
using B.quasi_inverses_def \<epsilon>.equivalence_in_bicategory_axioms by auto
have g: "\<guillemotleft>g : Cod F \<rightarrow>\<^sub>B Dom F\<guillemotright> \<and> B.is_left_adjoint g \<and> \<lbrakk>g\<rbrakk>\<^sub>B = \<lbrakk>g\<rbrakk>\<^sub>B"
using \<epsilon>'.dual_equivalence B.equivalence_is_left_adjoint B.equivalence_map_def
g_in_hhom
by blast
have F_as_MkArr: "F = MkArr (Dom F) (Cod F) \<lbrakk>f\<rbrakk>\<^sub>B"
using F MkArr_Map \<open>Map F = B.iso_class f\<close> by fastforce
have F_in_hom: "in_hom F (MkIde (Dom F)) (MkIde (Cod F))"
using F arr_char dom_char cod_char
by (intro in_homI, auto)
let ?G = "MkArr (Cod F) (Dom F) \<lbrakk>g\<rbrakk>\<^sub>B"
have "arr ?G"
using MkArr_in_hom' g by blast
hence G_in_hom: "\<guillemotleft>?G : MkIde (Cod F) \<rightarrow> MkIde (Dom F)\<guillemotright>"
using arr_char MkArr_in_hom by simp
have "inverse_arrows F ?G"
proof
show "ide (?G \<odot> F)"
proof -
have "?G \<odot> F = dom F"
proof (intro arr_eqI)
show 1: "seq ?G F"
using F_in_hom G_in_hom by blast
show "arr (dom F)"
using F_in_hom ide_dom by fastforce
show "Dom (?G \<odot> F) = Dom (dom F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Cod (?G \<odot> F) = Cod (dom F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Map (?G \<odot> F) = Map (dom F)"
proof -
have "Map (?G \<odot> F) = Comp \<lbrakk>g\<rbrakk>\<^sub>B \<lbrakk>f\<rbrakk>\<^sub>B"
- using 1 comp_char [of ?G F] `Map F = B.iso_class f` by simp
+ using 1 comp_char [of ?G F] \<open>Map F = B.iso_class f\<close> by simp
also have "... = \<lbrakk>g \<star> f\<rbrakk>\<^sub>B"
proof -
have "g \<star> f \<in> Comp \<lbrakk>g\<rbrakk>\<^sub>B \<lbrakk>f\<rbrakk>\<^sub>B"
by (metis \<epsilon>.ide_left \<epsilon>.ide_right \<epsilon>.unit_in_vhom \<epsilon>.unit_simps(3) B.arrI
B.ide_cod B.ide_in_iso_class in_CompI B.is_iso_classI
B.isomorphic_reflexive)
thus ?thesis
using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char
- `Map F = B.iso_class f`
+ \<open>Map F = B.iso_class f\<close>
by auto
qed
also have "... = \<lbrakk>src f\<rbrakk>\<^sub>B"
using \<epsilon>.unit_in_hom \<epsilon>.unit_is_iso B.isomorphic_def B.iso_class_def
B.isomorphic_symmetric
by (intro B.iso_class_eqI, blast)
also have "... = \<lbrakk>Dom F\<rbrakk>\<^sub>B"
using \<epsilon>.ide_left f_in_hhom B.ide_in_iso_class B.in_hhomE B.src.preserves_ide
B.isomorphic_reflexive
by (intro B.iso_class_eqI, fastforce)
also have "... = Map (dom F)"
using F_in_hom dom_char by auto
finally show ?thesis by blast
qed
qed
thus ?thesis
using F by simp
qed
show "ide (F \<odot> ?G)"
proof -
have "F \<odot> ?G = cod F"
proof (intro arr_eqI)
show 1: "seq F ?G"
using F_in_hom G_in_hom by blast
show "arr (cod F)"
using F_in_hom ide_cod by fastforce
show "Dom (F \<odot> ?G) = Dom (cod F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Cod (F \<odot> ?G) = Cod (cod F)"
using F_in_hom G_in_hom 1 comp_char by auto
show "Map (F \<odot> ?G) = Map (cod F)"
proof -
have "Map (F \<odot> ?G) = Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
- using 1 comp_char [of F ?G] `Map F = \<lbrakk>f\<rbrakk>\<^sub>B` by simp
+ using 1 comp_char [of F ?G] \<open>Map F = \<lbrakk>f\<rbrakk>\<^sub>B\<close> by simp
also have "... = \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
proof -
have "f \<star> g \<in> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
by (metis \<epsilon>.antipar(1) \<epsilon>.ide_left \<epsilon>.ide_right B.ide_hcomp
B.ide_in_iso_class in_CompI B.is_iso_classI B.isomorphic_reflexive)
thus ?thesis
using Comp_eq_iso_class_memb F_in_hom G_in_hom arr_char arr_char
- `Map F = \<lbrakk>f\<rbrakk>\<^sub>B`
+ \<open>Map F = \<lbrakk>f\<rbrakk>\<^sub>B\<close>
by auto
qed
also have "... = \<lbrakk>trg f\<rbrakk>\<^sub>B"
proof -
have "trg f \<in> \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
using \<epsilon>.counit_in_hom \<epsilon>.counit_is_iso B.isomorphic_def B.iso_class_def
B.isomorphic_symmetric
by blast
thus ?thesis
using B.iso_class_eqI
by (metis \<epsilon>.antipar(1) \<epsilon>.ide_left \<epsilon>.ide_right B.ide_hcomp
B.ide_in_iso_class B.is_iso_classI B.iso_class_elems_isomorphic)
qed
also have "... = \<lbrakk>Cod F\<rbrakk>\<^sub>B"
using f_in_hhom by auto
also have "... = Map (cod F)"
using F_in_hom dom_char by auto
finally show ?thesis by blast
qed
qed
thus ?thesis
using F by simp
qed
qed
thus ?thesis by auto
qed
qed
lemma iso_char':
shows "iso F \<longleftrightarrow> arr F \<and> (\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
proof -
have "arr F \<Longrightarrow> (\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f) \<longleftrightarrow>
(\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
proof
assume F: "arr F"
show "(\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f) \<Longrightarrow>
(\<exists>f. f \<in> Map F \<and> B.equivalence_map f)"
using F arr_char in_HomD(1) by blast
show "(\<exists>f. f \<in> Map F \<and> B.equivalence_map f) \<Longrightarrow>
(\<forall>f. f \<in> Map F \<longrightarrow> B.equivalence_map f)"
by (metis (no_types, lifting) F is_concrete_category concrete_category.arr_char
B.equivalence_map_preserved_by_iso in_HomD(2) B.iso_class_elems_isomorphic)
qed
thus ?thesis
using iso_char by blast
qed
text \<open>
The following mapping takes a map in \<open>B\<close> to the corresponding arrow of \<open>Maps(B)\<close>.
\<close>
abbreviation CLS ("\<lbrakk>\<lbrakk>_\<rbrakk>\<rbrakk>")
where "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<equiv> MkArr (src f) (trg f) \<lbrakk>f\<rbrakk>\<^sub>B"
lemma ide_CLS_obj:
assumes "B.obj a"
shows "ide \<lbrakk>\<lbrakk>a\<rbrakk>\<rbrakk>"
by (simp add: assms B.obj_simps)
lemma CLS_in_hom:
assumes "B.is_left_adjoint f"
shows "\<guillemotleft>\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> : \<lbrakk>\<lbrakk>src f\<rbrakk>\<rbrakk> \<rightarrow> \<lbrakk>\<lbrakk>trg f\<rbrakk>\<rbrakk>\<guillemotright>"
using assms B.left_adjoint_is_ide MkArr_in_hom MkArr_in_hom' by simp
lemma CLS_eqI:
assumes "B.ide f"
shows "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> \<longleftrightarrow> f \<cong>\<^sub>B g"
by (metis arr.inject assms B.ide_in_iso_class B.iso_class_def B.iso_class_eqI
B.isomorphic_implies_hpar(3) B.isomorphic_implies_hpar(4) B.isomorphic_symmetric
mem_Collect_eq)
lemma CLS_hcomp:
assumes "B.ide f" and "B.ide g" and "src f = trg g"
shows "\<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) (Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B)"
proof -
have "\<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
using assms B.left_adjoint_is_ide by simp
moreover have "\<lbrakk>f \<star> g\<rbrakk>\<^sub>B = Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
proof
show "\<lbrakk>f \<star> g\<rbrakk>\<^sub>B \<subseteq> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B"
by (metis assms(1-2) B.ide_in_iso_class in_CompI B.is_iso_classI
B.iso_class_def mem_Collect_eq subsetI)
show "Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B \<subseteq> \<lbrakk>f \<star> g\<rbrakk>\<^sub>B"
by (metis Comp_eq_iso_class_memb \<open>\<lbrakk>f \<star> g\<rbrakk>\<^sub>B \<subseteq> Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B\<close>
assms(1-3) B.ide_hcomp B.ide_in_iso_class subset_iff)
qed
ultimately show ?thesis by simp
qed
lemma comp_CLS:
assumes "B.is_left_adjoint f" and "B.is_left_adjoint g" and "f \<star> g \<cong>\<^sub>B h"
shows "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>h\<rbrakk>\<rbrakk>"
proof -
have hseq_fg: "B.hseq f g"
using assms B.isomorphic_implies_hpar(1) by simp
have "seq \<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk>"
using assms hseq_fg CLS_in_hom [of f] CLS_in_hom [of g]
apply (elim B.hseqE) by auto
hence "\<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>g\<rbrakk>\<rbrakk> = MkArr (src g) (trg f) (Comp \<lbrakk>f\<rbrakk>\<^sub>B \<lbrakk>g\<rbrakk>\<^sub>B)"
using comp_char [of "CLS f" "CLS g"] by simp
also have "... = \<lbrakk>\<lbrakk>f \<star> g\<rbrakk>\<rbrakk>"
using assms hseq_fg CLS_hcomp
apply (elim B.hseqE)
using B.adjoint_pair_antipar(1) by auto
also have "... = \<lbrakk>\<lbrakk>h\<rbrakk>\<rbrakk>"
using assms B.isomorphic_symmetric
by (simp add: assms(3) B.iso_class_eqI B.isomorphic_implies_hpar(3)
B.isomorphic_implies_hpar(4))
finally show ?thesis by simp
qed
text \<open>
The following mapping that takes an arrow of \<open>Maps(B)\<close> and chooses some
representative map in \<open>B\<close>.
\<close>
definition REP
where "REP F \<equiv> if arr F then SOME f. f \<in> Map F else B.null"
lemma REP_in_Map:
assumes "arr A"
shows "REP A \<in> Map A"
proof -
have "Map A \<noteq> {}"
using assms arr_char in_HomD(1) by blast
thus ?thesis
using assms REP_def someI_ex [of "\<lambda>f. f \<in> Map A"] by auto
qed
lemma REP_in_hhom [intro]:
assumes "in_hom F A B"
shows "\<guillemotleft>REP F : src (REP A) \<rightarrow>\<^sub>B src (REP B)\<guillemotright>"
and "B.is_left_adjoint (REP F)"
proof -
have "Map F \<noteq> {}"
using assms in_hom_char arr_char null_char in_HomD(1) [of "Map F" "Dom F" "Cod F"]
by blast
hence 1: "REP F \<in> Map F"
using assms REP_def someI_ex [of "\<lambda>f. f \<in> Map F"] by auto
hence 2: "B.arr (REP F)"
using assms 1 in_hom_char [of F A B] B.iso_class_def Map_memb_in_hhom B.in_hhom_def
by blast
hence "\<guillemotleft>REP F : Dom A \<rightarrow>\<^sub>B Dom B\<guillemotright>"
using assms 1 in_hom_char [of F A B] Map_memb_in_hhom by auto
thus "\<guillemotleft>REP F : src (REP A) \<rightarrow>\<^sub>B src (REP B)\<guillemotright>"
using assms
by (metis (no_types, lifting) Map_memb_in_hhom ideD(1)
in_homI in_hom_char REP_in_Map B.in_hhom_def)
have "REP F \<in> \<lbrakk>REP F\<rbrakk>\<^sub>B"
using assms 1 2 arr_char [of F] in_HomD(6) B.ide_in_iso_class B.iso_class_memb_is_ide
in_hom_char
by blast
thus "B.is_left_adjoint (REP F)"
using assms 1 2 arr_char in_HomD(5) [of "Map F" "Dom F" "Cod F" "REP F"]
by auto
qed
lemma ide_REP:
assumes "arr A"
shows "B.ide (REP A)"
using assms REP_in_hhom(2) B.adjoint_pair_antipar(1) by blast
lemma REP_simps [simp]:
assumes "arr A"
shows "B.ide (REP A)"
and "src (REP A) = Dom A" and "trg (REP A) = Cod A"
and "B.dom (REP A) = REP A" and "B.cod (REP A) = REP A"
proof -
show "B.ide (REP A)"
using assms ide_REP by blast
show "src (REP A) = Dom A"
using assms REP_in_hhom
by (metis (no_types, lifting) Map_memb_in_hhom Dom_dom in_homI
REP_in_Map B.in_hhom_def)
show "trg (REP A) = Cod A"
using assms REP_in_hhom
by (metis (no_types, lifting) Map_memb_in_hhom Dom_cod in_homI
REP_in_Map B.in_hhom_def)
show "B.dom (REP A) = REP A"
using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(2)
by blast
show "B.cod (REP A) = REP A"
using assms in_homI REP_in_hhom(2) B.adjoint_pair_antipar(1) B.ideD(3)
by blast
qed
lemma isomorphic_REP_src:
assumes "ide A"
shows "REP A \<cong>\<^sub>B src (REP A)"
using assms
by (metis (no_types, lifting) ideD(1) ide_char REP_in_Map ide_REP
REP_simps(2) B.is_iso_classI B.ide_in_iso_class B.iso_class_elems_isomorphic
B.src.preserves_ide)
lemma isomorphic_REP_trg:
assumes "ide A"
shows "REP A \<cong>\<^sub>B trg (REP A)"
using assms ide_char isomorphic_REP_src by auto
lemma CLS_REP:
assumes "arr F"
shows "\<lbrakk>\<lbrakk>REP F\<rbrakk>\<rbrakk> = F"
by (metis (mono_tags, lifting) MkArr_Map
is_concrete_category REP_in_Map REP_simps(2) REP_simps(3) assms
concrete_category.Map_in_Hom in_HomD(6))
lemma REP_CLS:
assumes "B.is_left_adjoint f"
shows "REP \<lbrakk>\<lbrakk>f\<rbrakk>\<rbrakk> \<cong>\<^sub>B f"
by (metis (mono_tags, lifting) CLS_in_hom Map.simps(1) in_homE REP_in_Map
assms B.iso_class_def B.isomorphic_symmetric mem_Collect_eq)
lemma isomorphic_hcomp_REP:
assumes "seq F G"
shows "REP F \<star> REP G \<cong>\<^sub>B REP (F \<odot> G)"
proof -
have 1: "Dom F = Cod G"
using assms seq_char by simp
have 2: "Map F = B.iso_class (REP F)"
using assms seq_char arr_char REP_in_Map in_HomD(6) by meson
have 3: "Map G = B.iso_class (REP G)"
using assms seq_char arr_char REP_in_Map
in_HomD(6) [of "Map G" "Dom G" "Cod G" "REP G"]
by auto
have "Map (F \<odot> G) = Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
using assms seq_char null_char
by (metis (no_types, lifting) CLS_REP Map.simps(1) Map_comp)
have "Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B = \<lbrakk>REP F \<star> REP G\<rbrakk>\<^sub>B"
proof -
have "REP F \<star> REP G \<in> Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
proof -
have "REP F \<in> Map F \<and> REP G \<in> Map G"
using assms seq_char REP_in_Map by auto
moreover have "REP F \<star> REP G \<cong>\<^sub>B REP F \<star> REP G"
using assms seq_char 2 B.isomorphic_reflexive by auto
ultimately show ?thesis
using assms 1 2 3 B.iso_class_def B.is_iso_class_def
by (intro in_CompI, auto)
qed
moreover have "\<lbrakk>REP F\<rbrakk>\<^sub>B \<in> Hom (Cod G) (Cod F)"
using assms 1 2 arr_char [of F] by auto
moreover have "\<lbrakk>REP G\<rbrakk>\<^sub>B \<in> Hom (Dom G) (Cod G)"
using assms 1 3 arr_char [of G] by auto
ultimately show ?thesis
using Comp_eq_iso_class_memb assms arr_char arr_char REP_in_Map
by (simp add: Comp_def)
qed
moreover have "REP (F \<odot> G) \<in> Comp \<lbrakk>REP F\<rbrakk>\<^sub>B \<lbrakk>REP G\<rbrakk>\<^sub>B"
proof -
have "Map (F \<odot> G) = Comp (Map F) (Map G)"
using assms 1 comp_char [of F G] by simp
thus ?thesis
using assms 1 2 3 REP_in_Map [of "F \<odot> G"] by simp
qed
ultimately show ?thesis
using B.iso_class_def by simp
qed
text \<open>
We now show that \<open>Maps(B)\<close> has pullbacks. For this we need to exhibit
functions \<open>PRJ\<^sub>0\<close> and \<open>PRJ\<^sub>1\<close> that produce the legs of the pullback of a cospan \<open>(H, K)\<close>
and verify that they have the required universal property. These are obtained by
choosing representatives \<open>h\<close> and \<open>k\<close> of \<open>H\<close> and \<open>K\<close>, respectively, and then taking
\<open>PRJ\<^sub>0 = CLS (tab\<^sub>0 (k\<^sup>* \<star> h))\<close> and \<open>PRJ\<^sub>1 = CLS (tab\<^sub>1 (k\<^sup>* \<star> h))\<close>. That these constitute a
pullback in \<open>Maps(B)\<close> follows from the fact that \<open>tab\<^sub>0 (k\<^sup>* \<star> h)\<close> and \<open>tab\<^sub>1 (k\<^sup>* \<star> h)\<close>
form a pseudo-pullback of \<open>(h, k)\<close> in the underlying bicategory.
For our purposes here, it is not sufficient simply to show that \<open>Maps(B)\<close> has pullbacks
and then to treat it as an abstract ``category with pullbacks'' where the pullbacks
are chosen arbitrarily. Instead, we have to retain the connection between a pullback
in Maps and the tabulation of \<open>k\<^sup>* \<star> h\<close> that is ultimately used to obtain it. If we don't
do this, then it becomes problematic to define the compositor of a pseudofunctor from
the underlying bicategory \<open>B\<close> to \<open>Span(Maps(B))\<close>, because the components will go from the
apex of a composite span (obtained by pullback) to the apex of a tabulation (chosen
arbitrarily using \<open>BS2\<close>) and these need not be in agreement with each other.
\<close>
definition PRJ\<^sub>0
where "PRJ\<^sub>0 \<equiv> \<lambda>K H. if cospan K H then \<lbrakk>\<lbrakk>B.tab\<^sub>0 ((REP K)\<^sup>* \<star> (REP H))\<rbrakk>\<rbrakk> else null"
definition PRJ\<^sub>1
where "PRJ\<^sub>1 \<equiv> \<lambda>K H. if cospan K H then \<lbrakk>\<lbrakk>B.tab\<^sub>1 ((REP K)\<^sup>* \<star> (REP H))\<rbrakk>\<rbrakk> else null"
interpretation elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1
proof
show "\<And>H K. \<not> cospan K H \<Longrightarrow> PRJ\<^sub>0 K H = null"
unfolding PRJ\<^sub>0_def by auto
show "\<And>H K. \<not> cospan K H \<Longrightarrow> PRJ\<^sub>1 K H = null"
unfolding PRJ\<^sub>1_def by auto
show "\<And>H K. cospan K H \<Longrightarrow> commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)"
proof -
fix H K
assume HK: "cospan K H"
define h where "h = REP H"
define k where "k = REP K"
have h: "h \<in> Map H"
using h_def HK REP_in_Map by blast
have k: "k \<in> Map K"
using k_def HK REP_in_Map by blast
have 1: "B.is_left_adjoint h \<and> B.is_left_adjoint k \<and> B.ide h \<and> B.ide k \<and> trg h = trg k"
using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide
in_HomD(5) [of "Map H" "Dom H" "Cod H" h]
in_HomD(5) [of "Map K" "Dom K" "Cod K" k]
apply auto
by (metis (no_types, lifting) HK Dom_cod)
interpret h: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h
using 1 by unfold_locales auto
interpret k: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg k
using 1 by unfold_locales auto
interpret hk: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h k
using 1 by unfold_locales auto
let ?f = "B.tab\<^sub>0 (k\<^sup>* \<star> h)"
let ?g = "B.tab\<^sub>1 (k\<^sup>* \<star> h)"
have span: "span \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
using dom_char CLS_in_hom [of ?f] CLS_in_hom [of ?g] by auto
have seq: "seq H \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
using HK seq_char hk.leg0_is_map CLS_in_hom h_def hk.p\<^sub>0_simps hk.satisfies_T0
by fastforce
have seq': "seq K \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
using HK k arr_char dom_char cod_char in_HomD(5) hk.leg1_is_map CLS_in_hom
by (metis (no_types, lifting) Cod.simps(1) seq_char REP_simps(2)
hk.p\<^sub>1_simps k_def span)
show "commutative_square K H (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)"
proof
show "cospan K H"
using HK by simp
show "dom K = cod (PRJ\<^sub>1 K H)"
using seq' PRJ\<^sub>1_def HK h_def k_def by auto
show "span (PRJ\<^sub>1 K H) (PRJ\<^sub>0 K H)"
unfolding PRJ\<^sub>0_def PRJ\<^sub>1_def using HK span h_def k_def by simp
show "K \<odot> PRJ\<^sub>1 K H = H \<odot> PRJ\<^sub>0 K H"
proof -
have iso: "h \<star> ?f \<cong>\<^sub>B k \<star> ?g"
using hk.\<phi>_uniqueness B.isomorphic_symmetric B.isomorphic_def by blast
have *: "Comp (Map H) \<lbrakk>?f\<rbrakk>\<^sub>B = Comp (Map K) \<lbrakk>?g\<rbrakk>\<^sub>B"
proof (intro Comp_eqI)
show "h \<star> ?f \<in> Comp (Map H) \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
proof (unfold Comp_def)
have "B.is_iso_class \<lbrakk>?f\<rbrakk>\<^sub>B"
by (simp add: B.is_iso_classI)
moreover have "B.is_iso_class (Map H)"
using CLS_REP HK Map.simps(1) B.is_iso_classI h.ide_left h_def
by (metis (no_types, lifting))
moreover have "?f \<in> \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
by (simp add: B.ide_in_iso_class(1))
moreover have "\<guillemotleft>?f : src (B.tab\<^sub>0 (k\<^sup>* \<star> h)) \<rightarrow>\<^sub>B Dom H\<guillemotright>"
using seq seq_char by simp
moreover have "h \<in> Map H"
by fact
moreover have "\<guillemotleft>h : Dom H \<rightarrow>\<^sub>B Cod H\<guillemotright>"
by (simp add: HK h_def)
moreover have "h \<star> ?f \<cong>\<^sub>B h \<star> ?f"
using B.isomorphic_reflexive by auto
ultimately show "h \<star> B.tab\<^sub>0 (k\<^sup>* \<star> h)
\<in> {h'. B.is_iso_class \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
B.is_iso_class (Map H) \<and>
(\<exists>f g. f \<in> \<lbrakk>B.tab\<^sub>0 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
g \<in> Map H \<and> g \<star> f \<cong>\<^sub>B h')}"
by auto
qed
show "k \<star> ?g \<in> Comp (Map K) \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
proof (unfold Comp_def)
have "B.is_iso_class \<lbrakk>?g\<rbrakk>\<^sub>B"
by (simp add: B.is_iso_classI)
moreover have "B.is_iso_class (Map K)"
by (metis (no_types, lifting) CLS_REP HK Map.simps(1)
B.is_iso_classI k.ide_left k_def)
moreover have "?g \<in> \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B"
by (simp add: B.ide_in_iso_class(1))
moreover have "\<guillemotleft>?g : src (B.tab\<^sub>1 (k\<^sup>* \<star> h)) \<rightarrow>\<^sub>B Dom K\<guillemotright>"
using seq seq_char B.in_hhom_def seq' by auto
moreover have "k \<in> Map K"
by fact
moreover have "\<guillemotleft>k : Dom K \<rightarrow>\<^sub>B Cod K\<guillemotright>"
by (simp add: HK k_def)
moreover have "k \<star> ?g \<cong>\<^sub>B k \<star> ?g"
using B.isomorphic_reflexive iso B.isomorphic_implies_hpar(2) by auto
ultimately show "k \<star> B.tab\<^sub>1 (k\<^sup>* \<star> h)
\<in> {h'. B.is_iso_class \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
B.is_iso_class (Map K) \<and>
(\<exists>f g. f \<in> \<lbrakk>B.tab\<^sub>1 (k\<^sup>* \<star> h)\<rbrakk>\<^sub>B \<and>
g \<in> Map K \<and> g \<star> f \<cong>\<^sub>B h')}"
by auto
qed
show "h \<star> ?f \<cong>\<^sub>B k \<star> ?g"
using iso by simp
qed
have "K \<odot> PRJ\<^sub>1 K H = K \<odot> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
unfolding PRJ\<^sub>1_def using HK h_def k_def by simp
also have "... = MkArr (src ?g) (Cod K) (Comp (Map K) \<lbrakk>?g\<rbrakk>\<^sub>B)"
using seq' comp_char [of K "\<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"] by simp
also have "... = MkArr (src ?f) (Cod H) (Comp (Map H) \<lbrakk>?f\<rbrakk>\<^sub>B)"
using * HK cod_char by auto
also have "... = comp H \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
using seq comp_char [of H "\<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"] by simp
also have "... = comp H (PRJ\<^sub>0 K H)"
unfolding PRJ\<^sub>0_def using HK h_def k_def by simp
finally show ?thesis by simp
qed
qed
qed
show "\<And>H K U V. commutative_square K H V U \<Longrightarrow>
\<exists>!E. comp (PRJ\<^sub>1 K H) E = V \<and> comp (PRJ\<^sub>0 K H) E = U"
proof -
fix H K U V
assume cs: "commutative_square K H V U"
have HK: "cospan K H"
using cs by auto
(* TODO: Is there any way to avoid this repetition? *)
define h where "h = REP H"
define k where "k = REP K"
have h: "h \<in> Map H"
using h_def HK REP_in_Map by blast
have k: "k \<in> Map K"
using k_def HK REP_in_Map by blast
have 1: "B.is_left_adjoint h \<and> B.is_left_adjoint k \<and> B.ide h \<and> B.ide k \<and> trg h = trg k"
using h k h_def k_def HK arr_char cod_char B.in_hhom_def B.left_adjoint_is_ide
in_HomD(5) [of "Map H" "Dom H" "Cod H" h]
in_HomD(5) [of "Map K" "Dom K" "Cod K" k]
apply auto
by (metis (no_types, lifting) HK Dom_cod)
interpret h: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h
using 1 by unfold_locales auto
interpret k: map_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg k
using 1 by unfold_locales auto
interpret hk: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg h k
using 1 by unfold_locales auto
let ?f = "B.tab\<^sub>0 (k\<^sup>* \<star> h)"
let ?g = "B.tab\<^sub>1 (k\<^sup>* \<star> h)"
have seq_HU: "seq H U"
using cs by auto
have seq_KV: "seq K V"
using cs by auto
let ?u = "REP U"
let ?v = "REP V"
have u: "B.ide ?u"
using ide_REP seq_HU by auto
have v: "B.ide ?v"
using ide_REP seq_KV by auto
have u_is_map: "B.is_left_adjoint ?u"
using u seq_HU REP_in_Map arr_char [of U]
in_HomD(5) [of "Map U" "Dom U" "Cod U" ?u]
by auto
have v_is_map: "B.is_left_adjoint ?v"
using v seq_KV REP_in_Map arr_char [of V]
in_HomD(5) [of "Map V" "Dom V" "Cod V" ?v]
by auto
have *: "h \<star> ?u \<cong>\<^sub>B k \<star> ?v"
proof -
have "h \<star> ?u \<cong>\<^sub>B REP (H \<odot> U)"
proof -
have "h \<star> ?u \<cong>\<^sub>B REP H \<star> ?u"
proof -
have "h \<cong>\<^sub>B REP H"
using h h_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic
in_HomD(5) [of "Map H" "Dom H" "Cod H" h] B.isomorphic_reflexive
by auto
thus ?thesis
using h_def seq_HU B.isomorphic_implies_hpar(1) B.isomorphic_reflexive
by (simp add: seq_char)
qed
also have "... \<cong>\<^sub>B REP (H \<odot> U)"
using seq_HU isomorphic_hcomp_REP isomorphic_def by blast
finally show ?thesis by blast
qed
also have "... \<cong>\<^sub>B REP (K \<odot> V)"
using seq_HU cs B.isomorphic_reflexive by auto
also have "... \<cong>\<^sub>B (k \<star> ?v)"
proof -
have "... \<cong>\<^sub>B REP K \<star> ?v"
using seq_KV isomorphic_hcomp_REP B.isomorphic_def B.isomorphic_symmetric
by blast
also have "... \<cong>\<^sub>B k \<star> ?v"
proof -
have "k \<cong>\<^sub>B REP K"
using k k_def HK arr_char REP_in_Map B.iso_class_elems_isomorphic
in_HomD(5) [of "Map K" "Dom K" "Cod K" k] B.isomorphic_reflexive
by auto
thus ?thesis
using k_def seq_KV B.isomorphic_implies_hpar(1) B.isomorphic_reflexive
by (simp add: seq_char)
qed
finally show ?thesis by blast
qed
finally show ?thesis by blast
qed
have hseq_hu: "src h = trg ?u"
using * B.isomorphic_implies_hpar
by (meson B.hseqE B.ideD(1))
have hseq_kv: "src k = trg ?v"
using * B.isomorphic_implies_hpar
by (meson B.hseqE B.ideD(1))
obtain w where w: "B.is_left_adjoint w \<and> ?f \<star> w \<cong>\<^sub>B ?u \<and> ?v \<cong>\<^sub>B (?g \<star> w)"
using * u_is_map v_is_map hk.has_pseudo_pullback [of ?u ?v] B.isomorphic_symmetric
by blast
have w_in_hom: "\<guillemotleft>w : src ?u \<rightarrow>\<^sub>B src ?f\<guillemotright> \<and> B.ide w"
using w B.left_adjoint_is_ide B.src_cod B.trg_cod B.isomorphic_def
by (metis B.hseqE B.ideD(1) B.in_hhom_def B.isomorphic_implies_hpar(3)
B.isomorphic_implies_ide(1) B.src_hcomp)
let ?W = "CLS w"
have W: "\<guillemotleft>?W : dom U \<rightarrow> dom \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>\<guillemotright>"
proof
show "arr ?W"
using w CLS_in_hom by blast
thus "dom ?W = dom U"
using w_in_hom dom_char REP_in_hhom(1) CLS_in_hom
by (metis (no_types, lifting) Dom.simps(1) commutative_squareE
dom_char REP_simps(2) cs B.in_hhomE)
show "cod ?W = dom \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>"
proof -
have "src ?f = trg w"
by (metis (lifting) B.in_hhomE w_in_hom)
thus ?thesis
using CLS_in_hom [of ?f] CLS_in_hom [of w] hk.satisfies_T0 w by fastforce
qed
qed
show "\<exists>!E. PRJ\<^sub>1 K H \<odot> E = V \<and> PRJ\<^sub>0 K H \<odot> E = U"
proof -
have "PRJ\<^sub>1 K H \<odot> ?W = V \<and> PRJ\<^sub>0 K H \<odot> ?W = U"
proof -
have "\<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<odot> ?W = U"
using w w_in_hom u CLS_in_hom comp_CLS
B.isomorphic_symmetric CLS_REP hk.leg0_is_map
by (metis (mono_tags, lifting) commutative_square_def cs)
moreover have "\<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<odot> ?W = V"
using w w_in_hom v CLS_in_hom comp_CLS
B.isomorphic_symmetric CLS_REP hk.leg1_is_map
by (metis (mono_tags, lifting) commutative_square_def cs)
ultimately show ?thesis
using HK h_def k_def PRJ\<^sub>0_def PRJ\<^sub>1_def by auto
qed
moreover have
"\<And>W'. PRJ\<^sub>1 K H \<odot> W' = V \<and> PRJ\<^sub>0 K H \<odot> W' = U \<Longrightarrow> W' = ?W"
proof -
fix W'
assume "PRJ\<^sub>1 K H \<odot> W' = V \<and> PRJ\<^sub>0 K H \<odot> W' = U"
hence W': "\<guillemotleft>W' : dom U \<rightarrow> dom \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk>\<guillemotright> \<and> \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<odot> W' = U \<and> \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<odot> W' = V"
using PRJ\<^sub>0_def PRJ\<^sub>1_def HK h_def k_def apply simp
using cs arr_iff_in_hom by blast
let ?w' = "REP W'"
have w': "B.ide ?w'"
using W' ide_REP by auto
have fw'_iso_u: "?f \<star> ?w' \<cong>\<^sub>B ?u"
proof -
have "?f \<star> ?w' \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<star> ?w'"
by (metis (no_types, lifting) Cod.simps(1) in_hom_char
REP_CLS REP_simps(3) W W' B.hcomp_isomorphic_ide hk.satisfies_T0
B.in_hhomE B.isomorphic_symmetric w' w_in_hom)
also have "REP \<lbrakk>\<lbrakk>?f\<rbrakk>\<rbrakk> \<star> ?w' \<cong>\<^sub>B ?u"
using W' isomorphic_hcomp_REP cs by blast
finally show ?thesis by blast
qed
have gw'_iso_v: "?g \<star> ?w' \<cong>\<^sub>B ?v"
proof -
have "?g \<star> ?w' \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk> \<star> ?w'"
proof -
have "?g \<cong>\<^sub>B REP \<lbrakk>\<lbrakk>?g\<rbrakk>\<rbrakk>"
using REP_CLS B.isomorphic_symmetric hk.leg1_is_map by blast
moreover have "B.ide (REP W')"
using W' by auto
moreover have "src ?f = trg ?w'"
using w_in_hom W W' in_hom_char arr_char B.in_hhom_def
by (meson fw'_iso_u B.hseqE B.ideD(1) B.isomorphic_implies_ide(1))
ultimately show ?thesis
using B.hcomp_isomorphic_ide by simp
qed
also have "... \<cong>\<^sub>B ?v"
using W' isomorphic_hcomp_REP cs by blast
finally show ?thesis by blast
qed
show "W' = ?W"
proof -
have "W' = \<lbrakk>\<lbrakk>?w'\<rbrakk>\<rbrakk>"
using w' W' CLS_REP by auto
also have "... = ?W"
proof -
have "?w' \<cong>\<^sub>B w"
using * w W' hk.has_pseudo_pullback(2) u_is_map v_is_map
B.isomorphic_symmetric fw'_iso_u gw'_iso_v
by blast
thus ?thesis
using CLS_eqI B.iso_class_eqI w' by blast
qed
finally show ?thesis by blast
qed
qed
ultimately show ?thesis by auto
qed
qed
qed
lemma is_elementary_category_with_pullbacks:
shows "elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1"
..
lemma is_category_with_pullbacks:
shows "category_with_pullbacks comp"
..
sublocale elementary_category_with_pullbacks comp PRJ\<^sub>0 PRJ\<^sub>1
using is_elementary_category_with_pullbacks by simp
end
text \<open>
Here we relate the projections of the chosen pullbacks in \<open>Maps(B)\<close> to the
projections associated with the chosen tabulations in \<open>B\<close>.
\<close>
context composite_tabulation_in_maps
begin
interpretation Maps: maps_category V H \<a> \<i> src trg
..
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map composable by unfold_locales auto
lemma prj_char:
shows "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
and "Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "Maps.arr (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
using \<sigma>.leg1_in_hom Maps.CLS_in_hom \<sigma>.leg1_is_map Maps.arr_char by auto
moreover have "Maps.arr (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>)"
using Maps.CLS_in_hom composable r\<^sub>0s\<^sub>1.k_is_map by fastforce
moreover have "Maps.cod (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>) =
Maps.cod (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
unfolding Maps.arr_char
using \<sigma>.leg1_in_hom \<rho>.leg0_in_hom
by (simp add: Maps.cod_char calculation(1) calculation(2))
ultimately have "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>))\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1 (Maps.CLS r\<^sub>0) (Maps.CLS s\<^sub>1) =
\<lbrakk>\<lbrakk>tab\<^sub>1 ((Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>0_def Maps.PRJ\<^sub>1_def
using Maps.CLS_in_hom \<sigma>.leg1_is_map \<rho>.leg0_is_map composable by simp
moreover have "r\<^sub>0\<^sup>* \<star> s\<^sub>1 \<cong> (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
proof -
have "r\<^sub>0 \<cong> Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>)"
using Maps.REP_CLS composable isomorphic_symmetric r\<^sub>0s\<^sub>1.k_is_map by fastforce
hence 3: "isomorphic r\<^sub>0\<^sup>* (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>*"
using \<rho>.leg0_is_map
by (simp add: isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint)
moreover have 4: "s\<^sub>1 \<cong> Maps.REP (Maps.MkArr (src s\<^sub>0) (trg s) \<lbrakk>s\<^sub>1\<rbrakk>)"
using Maps.REP_CLS isomorphic_symmetric r\<^sub>0s\<^sub>1.h_is_map by fastforce
ultimately show ?thesis
proof -
have 1: "src r\<^sub>0\<^sup>* = trg s\<^sub>1"
using \<rho>.T0.antipar(2) r\<^sub>0s\<^sub>1.cospan by argo
have 2: "ide s\<^sub>1"
by simp
have "src (Maps.REP (Maps.MkArr (src r\<^sub>0) (trg s) \<lbrakk>r\<^sub>0\<rbrakk>))\<^sup>* = trg s\<^sub>1"
by (metis 3 \<rho>.T0.antipar(2) isomorphic_implies_hpar(3) r\<^sub>0s\<^sub>1.cospan)
thus ?thesis
using 1 2
by (meson 3 4 hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_implies_ide(2)
isomorphic_transitive)
qed
qed
ultimately have 1: "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using r\<^sub>0s\<^sub>1.isomorphic_implies_same_tab by simp
show "Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>0 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using 1 by simp
show "Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>r\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>s\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>prj\<^sub>1 s\<^sub>1 r\<^sub>0\<rbrakk>\<rbrakk>"
using 1 by simp
qed
end
context identity_in_bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
notation isomorphic (infix "\<cong>" 50)
text \<open>
A 1-cell \<open>r\<close> in a bicategory of spans is a map if and only if the ``input leg''
\<open>tab\<^sub>0 r\<close> of the chosen tabulation of \<open>r\<close> is an equivalence map.
Since a tabulation of \<open>r\<close> is unique up to equivalence, and equivalence maps compose,
the result actually holds if ``chosen tabulation'' is replaced by ``any tabulation''.
\<close>
lemma is_map_iff_tab\<^sub>0_is_equivalence:
shows "is_left_adjoint r \<longleftrightarrow> equivalence_map (tab\<^sub>0 r)"
proof
assume 1: "equivalence_map (tab\<^sub>0 r)"
have 2: "quasi_inverses (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*"
proof -
obtain g' \<eta>' \<epsilon>' where \<eta>'\<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) g' \<eta>' \<epsilon>'"
using 1 equivalence_map_def by auto
have "adjoint_pair (tab\<^sub>0 r) g'"
using \<eta>'\<epsilon>' quasi_inverses_def quasi_inverses_are_adjoint_pair by blast
moreover have "adjoint_pair (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>*"
using T0.adjunction_in_bicategory_axioms adjoint_pair_def by auto
ultimately have "g' \<cong> (tab\<^sub>0 r)\<^sup>*"
using left_adjoint_determines_right_up_to_iso by simp
thus ?thesis
using \<eta>'\<epsilon>' quasi_inverses_def quasi_inverses_isomorphic_right by blast
qed
obtain \<eta>' \<epsilon>' where \<eta>'\<epsilon>': "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \<eta>' \<epsilon>'"
using 2 quasi_inverses_def by auto
interpret \<eta>'\<epsilon>': equivalence_in_bicategory V H \<a> \<i> src trg \<open>tab\<^sub>0 r\<close> \<open>(tab\<^sub>0 r)\<^sup>*\<close> \<eta>' \<epsilon>'
using \<eta>'\<epsilon>' by auto
have "is_left_adjoint (tab\<^sub>0 r)\<^sup>*"
using 2 quasi_inverses_are_adjoint_pair quasi_inverses_symmetric by blast
hence "is_left_adjoint (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using left_adjoints_compose by simp
thus "is_left_adjoint r"
using yields_isomorphic_representation isomorphic_def left_adjoint_preserved_by_iso'
by meson
next
assume 1: "is_left_adjoint r"
have 2: "is_left_adjoint (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using 1 yields_isomorphic_representation left_adjoint_preserved_by_iso'
isomorphic_symmetric isomorphic_def
by meson
hence "is_left_adjoint (tab\<^sub>0 r)\<^sup>*"
using is_ide BS4 [of "tab\<^sub>1 r" "(tab\<^sub>0 r)\<^sup>*"] by auto
hence "is_left_adjoint ((tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>0 r) \<and> is_left_adjoint (tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using left_adjoints_compose T0.antipar by simp
hence 3: "iso \<eta> \<and> iso \<epsilon>"
using BS3 [of "src (tab\<^sub>0 r)" "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>0 r" \<eta> \<eta>]
BS3 [of "tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*" "trg (tab\<^sub>0 r)" \<epsilon> \<epsilon>]
T0.unit_in_hom T0.counit_in_hom obj_is_self_adjoint
by auto
hence "equivalence_in_bicategory V H \<a> \<i> src trg (tab\<^sub>0 r) (tab\<^sub>0 r)\<^sup>* \<eta> \<epsilon>"
apply unfold_locales by auto
thus "equivalence_map (tab\<^sub>0 r)"
using equivalence_map_def by blast
qed
text \<open>
The chosen tabulation (and indeed, any other tabulation, which is equivalent)
of an object is symmetric in the sense that its two legs are isomorphic.
\<close>
lemma obj_has_symmetric_tab:
assumes "obj r"
shows "tab\<^sub>0 r \<cong> tab\<^sub>1 r"
proof -
have "tab\<^sub>0 r \<cong> r \<star> tab\<^sub>0 r"
proof -
have "trg (tab\<^sub>0 r) = r"
using assms by auto
moreover have "\<guillemotleft>\<l>\<^sup>-\<^sup>1[tab\<^sub>0 r] : tab\<^sub>0 r \<Rightarrow> trg (tab\<^sub>0 r) \<star> tab\<^sub>0 r\<guillemotright> \<and> iso \<l>\<^sup>-\<^sup>1[tab\<^sub>0 r]"
using assms by simp
ultimately show ?thesis
unfolding isomorphic_def by metis
qed
also have "... \<cong> tab\<^sub>1 r"
proof -
have "\<guillemotleft>tab : tab\<^sub>1 r \<Rightarrow> r \<star> tab\<^sub>0 r\<guillemotright>"
using tab_in_hom by simp
moreover have "is_left_adjoint (r \<star> tab\<^sub>0 r)"
using assms left_adjoints_compose obj_is_self_adjoint by simp
ultimately show ?thesis
using BS3 [of "tab\<^sub>1 r" "r \<star> tab\<^sub>0 r" tab tab] isomorphic_symmetric isomorphic_def
by auto
qed
finally show ?thesis by simp
qed
text \<open>
The chosen tabulation of \<open>r\<close> determines a span in \<open>Maps(B)\<close>.
\<close>
lemma determines_span:
assumes "ide r"
shows "span_in_category Maps.comp \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>\<rparr>"
using assms Maps.CLS_in_hom [of "tab\<^sub>0 r"] Maps.CLS_in_hom [of "tab\<^sub>1 r"]
tab\<^sub>0_in_hom tab\<^sub>1_in_hom
apply unfold_locales by fastforce
end
subsection "Arrows of Tabulations in Maps"
text \<open>
Here we consider the situation of two tabulations: a tabulation \<open>\<rho>\<close> of \<open>r\<close>
and a tabulation \<open>\<sigma>\<close> of \<open>s\<close>, both ``legs'' of each tabulation being maps,
together with an arbitrary 2-cell \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close>.
The 2-cell \<open>\<mu>\<close> at the base composes with the tabulation \<open>\<rho>\<close> to yield a 2-cell
\<open>\<Delta> = (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>\<close> ``over'' s. By property \<open>T1\<close> of tabulation \<open>\<sigma>\<close>, this induces a map
from the apex of \<open>\<rho>\<close> to the apex of \<open>\<sigma>\<close>, which together with the other data
forms a triangular prism whose sides commute up to (unique) isomorphism.
\<close>
text \<open>
$$
\xymatrix{
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \dtwocell\omit{^<-1>\sigma} & \\
&{\rm trg}~s && {\rm src}~s \ar[ll]^{s} \\
& \rrtwocell\omit{^\mu} &&\\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar@ {.>}[uuur]^<>(0.3){{\rm chine}} \dtwocell\omit{^\rho}& \\
{\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur]
}
$$
\<close>
locale arrow_of_tabulations_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and \<mu> :: 'a +
assumes in_hom: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
begin
abbreviation (input) \<Delta>
where "\<Delta> \<equiv> (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>"
lemma \<Delta>_in_hom [intro]:
shows "\<guillemotleft>\<Delta> : src \<rho> \<rightarrow> trg \<sigma>\<guillemotright>"
and "\<guillemotleft>\<Delta> : r\<^sub>1 \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>"
proof -
show "\<guillemotleft>\<Delta> : r\<^sub>1 \<Rightarrow> s \<star> r\<^sub>0\<guillemotright>"
using in_hom \<rho>.leg0_in_hom(2) \<rho>.tab_in_vhom' by auto
thus "\<guillemotleft>\<Delta> : src \<rho> \<rightarrow> trg \<sigma>\<guillemotright>"
by (metis \<rho>.tab_simps(3) \<rho>.base_in_hom(2) \<sigma>.tab_simps(3) \<sigma>.base_in_hom(2) arrI in_hom
seqI' vcomp_in_hhom vseq_implies_hpar(1-2))
qed
lemma \<Delta>_simps [simp]:
shows "arr \<Delta>"
and "src \<Delta> = src \<rho>" and "trg \<Delta> = trg \<sigma>"
and "dom \<Delta> = r\<^sub>1" and "cod \<Delta> = s \<star> r\<^sub>0"
using \<Delta>_in_hom by auto
abbreviation is_induced_map
where "is_induced_map w \<equiv> \<sigma>.is_induced_by_cell w r\<^sub>0 \<Delta>"
text \<open>
The following is an equivalent restatement, in elementary terms, of the conditions
for being an induced map.
\<close>
abbreviation (input) is_induced_map'
where "is_induced_map' w \<equiv>
ide w \<and>
(\<exists>\<nu> \<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>)"
lemma is_induced_map_iff:
shows "is_induced_map w \<longleftrightarrow> is_induced_map' w"
proof
assume w: "is_induced_map' w"
show "is_induced_map w"
proof
have 1: "dom \<Delta> = r\<^sub>1"
by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg
r\<^sub>0 \<open>dom \<Delta>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w
proof -
have "arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w"
using w apply unfold_locales by auto
thus "arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s \<sigma> s\<^sub>0 s\<^sub>1 w"
using 1 by simp
qed
show "arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s\<^sub>0 s\<^sub>1 w"
using w.arrow_of_spans_of_maps_axioms by auto
show "\<sigma>.composite_cell w w.the_\<theta> \<cdot> w.the_\<nu> = \<Delta>"
proof -
obtain \<theta> \<nu>
where \<theta>\<nu>: "\<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>"
using w w.the_\<theta>_props(1) by auto
have "(s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu> = \<Delta>"
using \<theta>\<nu> by argo
moreover have "\<theta> = w.the_\<theta> \<and> \<nu> = w.the_\<nu>"
using \<theta>\<nu> 1 w.the_\<nu>_props(1) w.leg0_uniquely_isomorphic w.leg1_uniquely_isomorphic
by auto
ultimately show ?thesis
using comp_assoc by simp
qed
qed
next
assume w: "is_induced_map w"
show "is_induced_map' w"
proof -
interpret w: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w
using w in_hom by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w
..
have "dom \<Delta> = r\<^sub>1" by auto
thus ?thesis
using w comp_assoc w.the_\<nu>_props(1) w.the_\<nu>_props(2) w.uw\<theta> by metis
qed
qed
lemma exists_induced_map:
shows "\<exists>w. is_induced_map w"
proof -
obtain w \<theta> \<nu>
where w\<theta>\<nu>: "ide w \<and> \<guillemotleft>\<theta> : s\<^sub>0 \<star> w \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<Delta> = (s \<star> \<theta>) \<cdot> \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> \<nu>"
using \<Delta>_in_hom \<rho>.ide_leg0 \<sigma>.T1 comp_assoc
by (metis in_homE)
thus ?thesis
using is_induced_map_iff by blast
qed
lemma induced_map_unique:
assumes "is_induced_map w" and "is_induced_map w'"
shows "w \<cong> w'"
using assms \<sigma>.induced_map_unique by blast
definition chine
where "chine \<equiv> SOME w. is_induced_map w"
lemma chine_is_induced_map:
shows "is_induced_map chine"
unfolding chine_def
using exists_induced_map someI_ex [of is_induced_map] by simp
lemma chine_in_hom [intro]:
shows "\<guillemotleft>chine : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>"
and "\<guillemotleft>chine: chine \<Rightarrow> chine\<guillemotright>"
proof -
show "\<guillemotleft>chine : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright>"
using chine_is_induced_map
by (metis \<Delta>_simps(1) \<Delta>_simps(4) \<rho>.leg1_simps(3) \<sigma>.ide_base \<sigma>.ide_leg0 \<sigma>.leg0_simps(3)
\<sigma>.tab_simps(2) arrow_of_spans_of_maps.is_ide arrow_of_spans_of_maps.the_\<nu>_simps(2)
assoc_simps(2) hseqE in_hhom_def seqE src_vcomp vseq_implies_hpar(1))
show "\<guillemotleft>chine: chine \<Rightarrow> chine\<guillemotright>"
using chine_is_induced_map
by (meson arrow_of_spans_of_maps.is_ide ide_in_hom(2))
qed
lemma chine_simps [simp]:
shows "arr chine" and "ide chine"
and "src chine = src r\<^sub>0" and "trg chine = src s\<^sub>0"
and "dom chine = chine" and "cod chine = chine"
using chine_in_hom apply auto
by (meson arrow_of_spans_of_maps.is_ide chine_is_induced_map)
end
sublocale arrow_of_tabulations_in_maps \<subseteq>
arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 chine
using chine_is_induced_map is_induced_map_iff
by unfold_locales auto
sublocale arrow_of_tabulations_in_maps \<subseteq>
arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 chine
..
context arrow_of_tabulations_in_maps
begin
text \<open>
The two factorizations of the composite 2-cell \<open>\<Delta>\<close> amount to a naturality condition.
\<close>
lemma \<Delta>_naturality:
shows "(\<mu> \<star> r\<^sub>0) \<cdot> \<rho> = (s \<star> the_\<theta>) \<cdot> \<a>[s, s\<^sub>0, chine] \<cdot> (\<sigma> \<star> chine) \<cdot> the_\<nu>"
using chine_is_induced_map is_induced_map_iff
by (metis leg0_uniquely_isomorphic(2) leg1_uniquely_isomorphic(2) the_\<nu>_props(1) uw\<theta>)
lemma induced_map_preserved_by_iso:
assumes "is_induced_map w" and "isomorphic w w'"
shows "is_induced_map w'"
proof -
interpret w: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s\<^sub>0 s\<^sub>1 w
using assms in_hom by auto
interpret w: arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 w
..
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow> w'\<guillemotright> \<and> iso \<phi>"
using assms(2) isomorphic_def by auto
show ?thesis
proof
interpret w': arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 \<open>dom \<Delta>\<close> s\<^sub>0 s\<^sub>1 w'
proof
show "is_left_adjoint r\<^sub>0"
by (simp add: \<rho>.satisfies_T0)
show "is_left_adjoint (dom \<Delta>)"
by (simp add: \<rho>.leg1_is_map)
show "ide w'" using assms by force
show "\<exists>\<theta>. \<guillemotleft>\<theta> : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using \<phi> w.the_\<theta>_props \<sigma>.leg0_in_hom(2) assms(2) comp_in_homI hcomp_in_vhom
inv_in_hom isomorphic_implies_hpar(4) w.the_\<theta>_simps(4) w.w_simps(4)
by metis
have "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
proof (intro conjI)
show "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using \<phi> w.the_\<nu>_props
by (intro comp_in_homI, auto)
thus "iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
using \<phi> w.the_\<nu>_props
by (meson \<sigma>.ide_leg1 arrI iso_hcomp hseqE ide_is_iso isos_compose seqE)
qed
hence "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>"
by auto
thus "\<exists>\<nu>. \<guillemotleft>\<nu> : dom \<Delta> \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>"
by auto
qed
interpret w': arrow_of_spans_of_maps_to_tabulation V H \<a> \<i> src trg
r\<^sub>0 \<open>dom \<Delta>\<close> s \<sigma> s\<^sub>0 s\<^sub>1 w'
..
show "arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) s\<^sub>0 s\<^sub>1 w'"
using w'.arrow_of_spans_of_maps_axioms by auto
show "\<sigma>.composite_cell w' w'.the_\<theta> \<cdot> w'.the_\<nu> = \<Delta>"
proof -
have 1: "w'.the_\<theta> = w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)"
proof -
have "\<guillemotleft>w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>) : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using w.the_\<theta>_props \<phi>
by (intro comp_in_homI, auto)
moreover have "\<guillemotleft>w'.the_\<theta> : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright>"
using w'.the_\<theta>_props by simp
ultimately show ?thesis
using w'.leg0_uniquely_isomorphic(2) by blast
qed
moreover have "w'.the_\<nu> = (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
proof -
have "\<guillemotleft>(s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu> : dom \<Delta> \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright>"
using w.the_\<nu>_props \<phi>
by (intro comp_in_homI, auto)
moreover have "iso ((s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>)"
using w.the_\<nu>_props \<phi> iso_hcomp
by (meson \<sigma>.ide_leg1 arrI calculation hseqE ide_is_iso isos_compose seqE)
ultimately show ?thesis
using w'.the_\<nu>_props w'.leg1_uniquely_isomorphic(2) by blast
qed
ultimately have "\<sigma>.composite_cell w' w'.the_\<theta> \<cdot> w'.the_\<nu> =
\<sigma>.composite_cell w' (w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)) \<cdot> (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
by simp
also have "... = (s \<star> w.the_\<theta> \<cdot> (s\<^sub>0 \<star> inv \<phi>)) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot>
(\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) \<cdot> w.the_\<nu>"
using comp_assoc by presburger
also have "... = (s \<star> w.the_\<theta>) \<cdot> ((s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot>
(\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>)) \<cdot> w.the_\<nu>"
using 1 comp_assoc w'.the_\<theta>_simps(1) whisker_left
by auto
also have "... = (s \<star> w.the_\<theta>) \<cdot> (\<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)) \<cdot> w.the_\<nu>"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) =
\<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>) =
\<a>[s, s\<^sub>0, w] \<cdot> ((s \<star> s\<^sub>0) \<star> inv \<phi>) \<cdot> (\<sigma> \<star> w') \<cdot> (s\<^sub>1 \<star> \<phi>)"
proof -
have "(s \<star> s\<^sub>0 \<star> inv \<phi>) \<cdot> \<a>[s, s\<^sub>0, w'] = \<a>[s, s\<^sub>0, w] \<cdot> ((s \<star> s\<^sub>0) \<star> inv \<phi>)"
using assms \<phi> assoc_naturality [of s s\<^sub>0 "inv \<phi>"] w.w_simps(4)
by (metis \<sigma>.leg0_simps(2-5) \<sigma>.base_simps(2-4) arr_inv cod_inv dom_inv
in_homE trg_cod)
thus ?thesis using comp_assoc by metis
qed
also have "... = \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>)"
proof -
have "((s \<star> s\<^sub>0) \<star> inv \<phi>) \<cdot> (\<sigma> \<star> w') = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>)"
using \<phi> comp_arr_dom comp_cod_arr in_hhom_def
interchange [of "s \<star> s\<^sub>0" \<sigma> "inv \<phi>" w']
interchange [of \<sigma> s\<^sub>1 w "inv \<phi>"]
by auto
thus ?thesis
using comp_assoc by metis
qed
also have "... = \<a>[s, s\<^sub>0, w] \<cdot> (\<sigma> \<star> w)"
proof -
have "(\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>) = \<sigma> \<star> w"
proof -
have "(\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi>) \<cdot> (s\<^sub>1 \<star> \<phi>) = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> inv \<phi> \<cdot> \<phi>)"
using \<phi> whisker_left in_hhom_def by auto
also have "... = (\<sigma> \<star> w) \<cdot> (s\<^sub>1 \<star> w)"
using \<phi> comp_inv_arr' by auto
also have "... = \<sigma> \<star> w"
using whisker_right [of w \<sigma> s\<^sub>1] comp_arr_dom in_hhom_def by auto
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = \<Delta>"
using assms(1) comp_assoc w.is_ide w.the_\<nu>_props(1) w.the_\<theta>_props(1) by simp
finally show ?thesis
using comp_assoc by auto
qed
qed
qed
end
text \<open>
In the special case that \<open>\<mu>\<close> is an identity 2-cell, the induced map from the apex of \<open>\<rho>\<close>
to the apex of \<open>\<sigma>\<close> is an equivalence map.
\<close>
locale identity_arrow_of_tabulations_in_maps =
arrow_of_tabulations_in_maps +
assumes is_ide: "ide \<mu>"
begin
lemma r_eq_s:
shows "r = s"
using is_ide by (metis ide_char in_hom in_homE)
lemma \<Delta>_eq_\<rho>:
shows "\<Delta> = \<rho>"
by (meson \<Delta>_simps(1) comp_ide_arr ide_hcomp hseq_char' ide_u is_ide seqE
seq_if_composable)
lemma chine_is_equivalence:
shows "equivalence_map chine"
proof -
obtain w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where e: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src s\<^sub>0 \<rightarrow> src r\<^sub>0\<guillemotright> \<and> \<guillemotleft>w' : src r\<^sub>0 \<rightarrow> src s\<^sub>0\<guillemotright> \<and>
\<guillemotleft>\<theta> : r\<^sub>0 \<star> w \<Rightarrow> s\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu> : s\<^sub>1 \<Rightarrow> r\<^sub>1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
\<sigma> = (s \<star> \<theta>) \<cdot> \<a>[s, r\<^sub>0, w] \<cdot> (\<rho> \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : s\<^sub>0 \<star> w' \<Rightarrow> r\<^sub>0\<guillemotright> \<and> \<guillemotleft>\<nu>' : r\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
\<rho> = (s \<star> \<theta>') \<cdot> \<a>[s, s\<^sub>0, w'] \<cdot> (\<sigma> \<star> w') \<cdot> \<nu>'"
using r_eq_s \<sigma>.apex_unique_up_to_equivalence [of \<rho> r\<^sub>0 r\<^sub>1] \<rho>.tabulation_axioms by blast
have w': "equivalence_map w'"
using e equivalence_map_def by auto
hence "is_induced_map w'"
using e r_eq_s \<Delta>_eq_\<rho> is_induced_map_iff comp_assoc equivalence_map_is_ide by metis
hence "isomorphic chine w'"
using induced_map_unique chine_is_induced_map by simp
thus ?thesis
using w' equivalence_map_preserved_by_iso isomorphic_symmetric by blast
qed
end
text \<open>
The following gives an interpretation of @{locale arrow_of_tabulations_in_maps}
in the special case that the tabulations are those that we have chosen for the
domain and codomain of the underlying 2-cell \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close>.
In this case, we can recover \<open>\<mu>\<close> from \<open>\<Delta>\<close> via adjoint transpose.
\<close>
locale arrow_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r +
s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and s :: 'a
and \<mu> :: 'a +
assumes in_hom: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
begin
abbreviation (input) r\<^sub>0 where "r\<^sub>0 \<equiv> tab\<^sub>0 r"
abbreviation (input) r\<^sub>1 where "r\<^sub>1 \<equiv> tab\<^sub>1 r"
abbreviation (input) s\<^sub>0 where "s\<^sub>0 \<equiv> tab\<^sub>0 s"
abbreviation (input) s\<^sub>1 where "s\<^sub>1 \<equiv> tab\<^sub>1 s"
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \<mu>"
using in_hom by unfold_locales auto
end
sublocale identity_in_bicategory_of_spans \<subseteq> arrow_in_bicategory_of_spans V H \<a> \<i> src trg r r r
apply unfold_locales using is_ide by auto
context arrow_in_bicategory_of_spans
begin
interpretation arrow_of_tabulations_in_maps V H \<a> \<i> src trg r r.tab r\<^sub>0 r\<^sub>1 s s.tab s\<^sub>0 s\<^sub>1 \<mu>
using is_arrow_of_tabulations_in_maps by simp
interpretation r: arrow_of_tabulations_in_maps V H \<a> \<i> src trg r r.tab r\<^sub>0 r\<^sub>1 r r.tab r\<^sub>0 r\<^sub>1 r
using r.is_arrow_of_tabulations_in_maps by simp
lemma \<mu>_in_terms_of_\<Delta>:
shows "\<mu> = r.T0.trnr\<^sub>\<epsilon> (cod \<mu>) \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
proof -
have \<mu>: "arr \<mu>"
using in_hom by auto
have "\<mu> \<cdot> r.T0.trnr\<^sub>\<epsilon> r r.tab = r.T0.trnr\<^sub>\<epsilon> s \<Delta>"
proof -
have "\<mu> \<cdot> r.T0.trnr\<^sub>\<epsilon> r r.tab =
(\<mu> \<cdot> \<r>[r]) \<cdot> (r \<star> r.\<epsilon>) \<cdot> \<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
unfolding r.T0.trnr\<^sub>\<epsilon>_def using comp_assoc by presburger
also have "... = \<r>[s] \<cdot> ((\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>)) \<cdot>
\<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> runit_naturality comp_assoc
by (metis in_hom in_homE)
also have "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> ((\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot>
\<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*]) \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "(\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>) = \<mu> \<star> r.\<epsilon>"
using \<mu> interchange comp_arr_dom comp_cod_arr
by (metis in_hom in_homE r.T0.counit_simps(1) r.T0.counit_simps(3) r.u_simps(3)
src_dom)
also have "... = (s \<star> r.\<epsilon>) \<cdot> (\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
using in_hom interchange [of s \<mu> r.\<epsilon> "tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*"]
comp_arr_dom comp_cod_arr r.T0.counit_simps(1) r.T0.counit_simps(2)
by auto
finally have "(\<mu> \<star> src \<mu>) \<cdot> (r \<star> r.\<epsilon>) = (s \<star> r.\<epsilon>) \<cdot> (\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*)"
by blast
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> \<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot>
((\<mu> \<star> tab\<^sub>0 r) \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot> (r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
proof -
have "(\<mu> \<star> tab\<^sub>0 r \<star> (tab\<^sub>0 r)\<^sup>*) \<cdot> \<a>[r, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] =
\<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot> ((\<mu> \<star> tab\<^sub>0 r) \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> assoc_naturality [of \<mu> "tab\<^sub>0 r" "(tab\<^sub>0 r)\<^sup>*"]
by (metis ide_char in_hom in_homE r.T0.antipar(1) r.T0.ide_right r.u_simps(3)
src_dom u_simps(2) u_simps(4-5))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<r>[s] \<cdot> (s \<star> r.\<epsilon>) \<cdot> \<a>[s, tab\<^sub>0 r, (tab\<^sub>0 r)\<^sup>*] \<cdot>
((\<mu> \<star> tab\<^sub>0 r) \<cdot> r.tab \<star> (tab\<^sub>0 r)\<^sup>*)"
using \<mu> whisker_right \<Delta>_simps(1) by auto
also have "... = r.T0.trnr\<^sub>\<epsilon> s \<Delta>"
unfolding r.T0.trnr\<^sub>\<epsilon>_def by simp
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> r.yields_isomorphic_representation invert_side_of_triangle(2)
by (metis in_hom in_homE seqI')
qed
end
subsubsection "Vertical Composite"
locale vertical_composite_of_arrows_of_tabulations_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>: tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 +
\<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 \<mu> +
\<pi>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<pi>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and t :: 'a
and \<tau> :: 'a
and t\<^sub>0 :: 'a
and t\<^sub>1 :: 'a
and \<mu> :: 'a
and \<pi> :: 'a
begin
text \<open>
$$
\xymatrix{
&&& {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^{t_0} \dtwocell\omit{^<-1>\tau} & \\
&&{\rm trg}~t && {\rm src}~t \ar[ll]^{s} \\
&& \rrtwocell\omit{^\pi} && \\
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.3){\pi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\
&{\rm trg}~s \ar@ {=}[uuur] && {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\
& \rrtwocell\omit{^\mu} &&\\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.3){\mu.{\rm chine}} \dtwocell\omit{^\rho} & \\
{\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r \ar[ll]^{r} \ar@ {=}[uuur]
}
$$
\<close>
notation isomorphic (infix "\<cong>" 50)
interpretation arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<open>\<pi> \<cdot> \<mu>\<close>
using \<mu>.in_hom \<pi>.in_hom by (unfold_locales, blast)
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 (\<pi> \<cdot> \<mu>)"
..
lemma chine_char:
shows "chine \<cong> \<pi>.chine \<star> \<mu>.chine"
proof -
have "is_induced_map (\<pi>.chine \<star> \<mu>.chine)"
proof -
let ?f = "\<mu>.chine"
have f: "\<guillemotleft>?f : src \<rho> \<rightarrow> src \<sigma>\<guillemotright> \<and> is_left_adjoint ?f \<and> ide ?f \<and> \<mu>.is_induced_map ?f"
using \<mu>.chine_is_induced_map \<mu>.is_map by auto
let ?g = "\<pi>.chine"
have g: "\<guillemotleft>?g : src \<sigma> \<rightarrow> src \<tau>\<guillemotright> \<and> is_left_adjoint ?g \<and> ide ?g \<and> \<pi>.is_induced_map ?g"
using \<pi>.chine_is_induced_map \<pi>.is_map by auto
let ?\<theta> = "\<mu>.the_\<theta> \<cdot> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
let ?\<nu> = "\<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
have \<theta>: "\<guillemotleft>?\<theta> : t\<^sub>0 \<star> ?g \<star> ?f \<Rightarrow> r\<^sub>0\<guillemotright>"
using f g \<pi>.the_\<theta>_props \<mu>.the_\<theta>_props
by (intro comp_in_homI hcomp_in_vhom, auto+)
have \<nu>: "\<guillemotleft>?\<nu> : r\<^sub>1 \<Rightarrow> t\<^sub>1 \<star> ?g \<star> ?f\<guillemotright>"
using f g \<pi>.the_\<theta>_props \<mu>.the_\<theta>_props
by (intro comp_in_homI hcomp_in_vhom, auto)
interpret gf: arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 r\<^sub>1 t\<^sub>0 t\<^sub>1 \<open>?g \<star> ?f\<close>
proof
show "ide (?g \<star> ?f)" by simp
show "\<exists>\<theta>. \<guillemotleft>\<theta> : t\<^sub>0 \<star> ?g \<star> ?f \<Rightarrow> r\<^sub>0\<guillemotright>"
using \<theta> by blast
show "\<exists>\<nu>. \<guillemotleft>\<nu> : r\<^sub>1 \<Rightarrow> t\<^sub>1 \<star> ?g \<star> ?f\<guillemotright> \<and> iso \<nu>"
using \<nu> \<mu>.the_\<nu>_props \<mu>.the_\<theta>_props \<pi>.the_\<nu>_props \<pi>.the_\<theta>_props
isos_compose [of "\<mu>.the_\<nu>" "\<pi>.the_\<nu>"] \<mu>.is_ide \<nu> \<open>ide (\<pi>.chine \<star> \<mu>.chine)\<close>
\<pi>.uw\<theta> \<pi>.w_simps(4) \<tau>.ide_leg1 \<tau>.leg1_simps(3) arrI hseq_char ideD(1)
ide_is_iso iso_assoc iso_hcomp isos_compose seqE
by metis
qed
show ?thesis
proof (intro conjI)
have \<theta>_eq: "?\<theta> = gf.the_\<theta>"
using \<theta> gf.the_\<theta>_props gf.leg0_uniquely_isomorphic by auto
have \<nu>_eq: "?\<nu> = gf.the_\<nu>"
using \<nu> gf.the_\<nu>_props gf.leg1_uniquely_isomorphic by auto
have A: "src ?g = trg ?f"
using f g by fastforce
show "arrow_of_spans_of_maps V H \<a> \<i> src trg r\<^sub>0 (dom \<Delta>) t\<^sub>0 t\<^sub>1 (?g \<star> ?f)"
using gf.arrow_of_spans_of_maps_axioms by simp
have "((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot> gf.the_\<nu> = \<Delta>"
proof -
have "\<Delta> = (\<pi> \<star> r\<^sub>0) \<cdot> (\<mu> \<star> r\<^sub>0) \<cdot> \<rho>"
using whisker_right comp_assoc
by (metis \<Delta>_simps(1) hseqE ide_u seqE)
also have "... = ((\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>)) \<cdot> \<a>[s, s\<^sub>0, ?f] \<cdot> (\<sigma> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using \<mu>.\<Delta>_naturality comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> ((\<pi> \<star> s\<^sub>0 \<star> ?f) \<cdot> \<a>[s, s\<^sub>0, ?f]) \<cdot> (\<sigma> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have "(\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>) = \<pi> \<star> \<mu>.the_\<theta>"
using f comp_arr_dom comp_cod_arr \<mu>.the_\<theta>_props \<pi>.in_hom
interchange [of \<pi> s r\<^sub>0 \<mu>.the_\<theta>]
by (metis in_homE)
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<pi> \<star> s\<^sub>0 \<star> ?f)"
using f comp_arr_dom comp_cod_arr \<mu>.the_\<theta>_props \<pi>.in_hom
interchange [of t \<pi> \<mu>.the_\<theta> "s\<^sub>0 \<star> ?f"]
by (metis in_homE)
finally have "(\<pi> \<star> r\<^sub>0) \<cdot> (s \<star> \<mu>.the_\<theta>) = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<pi> \<star> s\<^sub>0 \<star> ?f)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot> (((\<pi> \<star> s\<^sub>0) \<star> ?f) \<cdot> (\<sigma> \<star> ?f)) \<cdot> \<mu>.the_\<nu>"
proof -
have "(\<pi> \<star> s\<^sub>0 \<star> ?f) \<cdot> \<a>[s, s\<^sub>0, ?f] = \<a>[t, s\<^sub>0, ?f] \<cdot> ((\<pi> \<star> s\<^sub>0) \<star> ?f)"
using f assoc_naturality [of \<pi> s\<^sub>0 ?f] \<pi>.in_hom by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot> (\<pi>.\<Delta> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using whisker_right comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot>
((t \<star> \<pi>.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g] \<cdot> (\<tau> \<star> ?g) \<cdot> \<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using \<pi>.\<Delta>_naturality by simp
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> \<a>[t, s\<^sub>0, ?f] \<cdot>
(((t \<star> \<pi>.the_\<theta>) \<star> ?f) \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) \<cdot> ((\<tau> \<star> ?g) \<star> ?f) \<cdot>
(\<pi>.the_\<nu> \<star> ?f)) \<cdot> \<mu>.the_\<nu>"
using f g \<pi>.the_\<theta>_props \<pi>.the_\<nu>_props whisker_right
by (metis \<pi>.\<Delta>_simps(1) \<pi>.\<Delta>_naturality seqE)
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (\<a>[t, s\<^sub>0, ?f] \<cdot>
((t \<star> \<pi>.the_\<theta>) \<star> ?f)) \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) \<cdot> ((\<tau> \<star> ?g) \<star> ?f) \<cdot>
(\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using comp_assoc by presburger
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot>
(\<a>[t, t\<^sub>0 \<star> ?g, ?f] \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f)) \<cdot>
((\<tau> \<star> ?g) \<star> ?f) \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using f \<pi>.the_\<theta>_props assoc_naturality [of t "\<pi>.the_\<theta>" ?f] \<pi>.\<theta>_simps(3) comp_assoc
by auto
also have "... = (t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<a>[t \<star> t\<^sub>0, ?g, ?f] \<cdot>
((\<tau> \<star> ?g) \<star> ?f)) \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have "seq \<a>[t, t\<^sub>0, ?g \<star> ?f] \<a>[t \<star> t\<^sub>0, ?g, ?f]"
using f g by fastforce
moreover have "inv (t \<star> \<a>[t\<^sub>0, ?g, ?f]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
using f g by simp
moreover have "iso (t \<star> \<a>[t\<^sub>0, ?g, ?f])"
using f g by simp
have "\<a>[t, t\<^sub>0 \<star> ?g, ?f] \<cdot> (\<a>[t, t\<^sub>0, ?g] \<star> ?f) =
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> \<a>[t \<star> t\<^sub>0, ?g, ?f]"
proof -
have "seq \<a>[t, t\<^sub>0, ?g \<star> ?f] \<a>[t \<star> t\<^sub>0, ?g, ?f]"
using f g by fastforce
moreover have "inv (t \<star> \<a>[t\<^sub>0, ?g, ?f]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]"
using f g by simp
moreover have "iso (t \<star> \<a>[t\<^sub>0, ?g, ?f])"
using f g by simp
ultimately show ?thesis
using A f g pentagon invert_side_of_triangle(1)
by (metis \<pi>.w_simps(4) \<tau>.ide_base \<tau>.ide_leg0 \<tau>.leg0_simps(3))
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> \<mu>.the_\<theta>) \<cdot> (t \<star> \<pi>.the_\<theta> \<star> ?f)) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f) \<cdot>
\<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
using f g assoc_naturality [of \<tau> ?g ?f] comp_assoc by simp
also have "... = (t \<star> \<mu>.the_\<theta> \<cdot> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f]) \<cdot>
\<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot>
(\<tau> \<star> ?g \<star> ?f) \<cdot> \<a>[t\<^sub>1, ?g, ?f] \<cdot> (\<pi>.the_\<nu> \<star> ?f) \<cdot> \<mu>.the_\<nu>"
proof -
have 1: "seq \<mu>.the_\<theta> ((\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])"
using \<theta>_eq by auto
hence "t \<star> (\<pi>.the_\<theta> \<star> ?f) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f] =
(t \<star> \<pi>.the_\<theta> \<star> ?f) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, ?g, ?f])"
using whisker_left \<tau>.ide_base by blast
thus ?thesis
using 1 whisker_left \<tau>.ide_base comp_assoc by presburger
qed
also have "... = ((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot> gf.the_\<nu>"
using \<theta>_eq \<nu>_eq by (simp add: comp_assoc)
finally show ?thesis
using comp_assoc by presburger
qed
thus "((t \<star> gf.the_\<theta>) \<cdot> \<a>[t, t\<^sub>0, ?g \<star> ?f] \<cdot> (\<tau> \<star> ?g \<star> ?f)) \<cdot>
arrow_of_spans_of_maps.the_\<nu> (\<cdot>) (\<star>) (dom ((\<pi> \<cdot> \<mu> \<star> r\<^sub>0) \<cdot> \<rho>)) t\<^sub>1 (?g \<star> ?f) =
\<Delta>"
by simp
qed
qed
thus ?thesis
using chine_is_induced_map induced_map_unique by simp
qed
end
sublocale vertical_composite_of_arrows_of_tabulations_in_maps \<subseteq>
arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 "\<pi> \<cdot> \<mu>"
using is_arrow_of_tabulations_in_maps by simp
subsubsection "Horizontal Composite"
locale horizontal_composite_of_arrows_of_tabulations_in_maps =
bicategory_of_spans V H \<a> \<i> src trg +
\<rho>: tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 +
\<sigma>: tabulation_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>: tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 +
\<mu>: tabulation_in_maps V H \<a> \<i> src trg u \<mu> u\<^sub>0 u\<^sub>1 +
\<rho>\<sigma>: composite_tabulation_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 s \<sigma> s\<^sub>0 s\<^sub>1 +
\<tau>\<mu>: composite_tabulation_in_maps V H \<a> \<i> src trg t \<tau> t\<^sub>0 t\<^sub>1 u \<mu> u\<^sub>0 u\<^sub>1 +
\<omega>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg r \<rho> r\<^sub>0 r\<^sub>1 t \<tau> t\<^sub>0 t\<^sub>1 \<omega> +
\<chi>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg s \<sigma> s\<^sub>0 s\<^sub>1 u \<mu> u\<^sub>0 u\<^sub>1 \<chi>
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and \<rho> :: 'a
and r\<^sub>0 :: 'a
and r\<^sub>1 :: 'a
and s :: 'a
and \<sigma> :: 'a
and s\<^sub>0 :: 'a
and s\<^sub>1 :: 'a
and t :: 'a
and \<tau> :: 'a
and t\<^sub>0 :: 'a
and t\<^sub>1 :: 'a
and u :: 'a
and \<mu> :: 'a
and u\<^sub>0 :: 'a
and u\<^sub>1 :: 'a
and \<omega> :: 'a
and \<chi> :: 'a
begin
text \<open>
$$
\xymatrix{
&&& {\rm src}~t_0u_1.\phi \ar[dl]_{\tau\mu.p_1} \ar[dr]^{\tau\mu.p_0} \ddtwocell\omit{^{t_0u_1.\phi}} \\
&& {\rm src}~\tau \ar[dl]_{t_1} \ar[dr]^<>(0.4){t_0} \dtwocell\omit{^<-1>\tau}
&& {\rm src}~\mu \ar[dl]_{u_1} \ar[dr]^{u_0} \dtwocell\omit{^<-1>\mu} & \\
& {\rm trg}~t && {\rm src}~t = {\rm trg}~u \ar[ll]^{t}
&& {\rm src}~u \ar[ll]^{u} \\
& \xtwocell[r]{}\omit{^\omega}
& {\rm src}~r_0s_1.\phi \ar[uuur]_<>(0.2){{\rm chine}}
\ar[dl]^{\rho\sigma.p_1} \ar[dr]_{\rho\sigma.p_0\hspace{20pt}} \ddtwocell\omit{^{r_0s_1.\phi}}
& \rrtwocell\omit{^\chi} && \\
& {\rm src}~\rho \ar[dl]_{r_1} \ar[dr]^{r_0} \ar[uuur]^<>(0.4){\omega.{\rm chine}} \dtwocell\omit{^\rho}
&& {\rm src}~\sigma \ar[dl]_{s_1} \ar[dr]^{s_0} \ar[uuur]^<>(0.4){\chi.{\rm chine}} \dtwocell\omit{^<-1>\sigma} & \\
{\rm trg}~r \ar@ {=}[uuur] && {\rm src}~r = {\rm trg}~s \ar[ll]^{r} \ar@ {=}[uuur]
&& {\rm src}~s \ar[ll]^{s} \ar@ {=}[uuur] \\
}
$$
\<close>
notation isomorphic (infix "\<cong>" 50)
interpretation arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>t \<star> u\<close> \<tau>\<mu>.tab \<open>u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0\<close> \<open>t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1\<close> \<open>\<omega> \<star> \<chi>\<close>
using \<rho>\<sigma>.composable \<omega>.in_hom \<chi>.in_hom
by unfold_locales auto
lemma is_arrow_of_tabulations_in_maps:
shows "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(r \<star> s) \<rho>\<sigma>.tab (s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1)
(t \<star> u) \<tau>\<mu>.tab (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) (\<omega> \<star> \<chi>)"
..
sublocale arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>t \<star> u\<close> \<tau>\<mu>.tab \<open>u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0\<close> \<open>t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1\<close> \<open>\<omega> \<star> \<chi>\<close>
using is_arrow_of_tabulations_in_maps by simp
interpretation Maps: maps_category V H \<a> \<i> src trg ..
notation Maps.comp (infixr "\<odot>" 55)
interpretation r\<^sub>0s\<^sub>1: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg s\<^sub>1 r\<^sub>0
using \<rho>.leg0_is_map \<sigma>.leg1_is_map \<rho>\<sigma>.composable apply unfold_locales by auto
interpretation r\<^sub>0s\<^sub>1: arrow_of_tabulations_in_maps \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close> r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close> r\<^sub>0s\<^sub>1.tab r\<^sub>0s\<^sub>1.p\<^sub>0 r\<^sub>0s\<^sub>1.p\<^sub>1
\<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>
using r\<^sub>0s\<^sub>1.is_arrow_of_tabulations_in_maps by simp
interpretation t\<^sub>0u\<^sub>1: cospan_of_maps_in_bicategory_of_spans \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg u\<^sub>1 t\<^sub>0
using \<tau>.leg0_is_map \<mu>.leg1_is_map \<tau>\<mu>.composable apply unfold_locales by auto
interpretation t\<^sub>0u\<^sub>1: arrow_of_tabulations_in_maps \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close> t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close> t\<^sub>0u\<^sub>1.tab t\<^sub>0u\<^sub>1.p\<^sub>0 t\<^sub>0u\<^sub>1.p\<^sub>1
\<open>t\<^sub>0\<^sup>* \<star> u\<^sub>1\<close>
using t\<^sub>0u\<^sub>1.is_arrow_of_tabulations_in_maps by simp
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
no_notation in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text \<open>
The following lemma states that the rectangular faces of the ``top prism'' commute
up to isomorphism. This was not already proved in @{locale composite_tabulation_in_maps},
because there we did not consider any composite structure of the ``source'' 2-cell.
There are common elements, though to the proof that the composite of tabulations is
a tabulation and the present lemma.
The proof idea is to use property \<open>T2\<close> of the ``base'' tabulations to establish the
existence of the desired isomorphisms. The proofs have to be carried out in
sequence, starting from the ``output'' side, because the arrow \<open>\<beta>\<close> required in the
hypotheses of \<open>T2\<close> depends, for the ``input'' tabulation, on the isomorphism constructed
for the ``output'' tabulation.
\<close>
lemma prj_chine:
shows "\<tau>\<mu>.p\<^sub>0 \<star> chine \<cong> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
and "\<tau>\<mu>.p\<^sub>1 \<star> chine \<cong> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
proof -
have 1: "arrow_of_spans_of_maps V H \<a> \<i> src trg
(s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1) (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) chine \<and>
(((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<tau>\<mu>.tab \<star> chine)) \<cdot> the_\<nu> =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab"
using chine_is_induced_map by simp
let ?u\<^sub>\<tau> = "u \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
let ?w\<^sub>\<tau> = "\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
let ?w\<^sub>\<tau>' = "\<tau>\<mu>.p\<^sub>1 \<star> chine"
have u\<^sub>\<tau>: "ide ?u\<^sub>\<tau>"
using \<chi>.u_simps(3) by auto
have w\<^sub>\<tau>: "ide ?w\<^sub>\<tau> \<and> is_left_adjoint ?w\<^sub>\<tau>"
by (simp add: \<omega>.is_map \<rho>.T0.antipar(1) left_adjoints_compose)
have w\<^sub>\<tau>': "ide ?w\<^sub>\<tau>' \<and> is_left_adjoint ?w\<^sub>\<tau>'"
by (simp add: is_map left_adjoints_compose)
let ?\<theta>\<^sub>\<tau> = "\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot> (\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
let ?\<theta>\<^sub>\<tau>' = "(u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
let ?\<beta>\<^sub>\<tau> = "\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
have \<theta>\<^sub>\<tau>: "\<guillemotleft>?\<theta>\<^sub>\<tau> : t\<^sub>0 \<star> ?w\<^sub>\<tau> \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1) \<omega>.the_\<theta>_in_hom \<chi>.u_simps(3)
by (intro comp_in_homI, auto)
have \<theta>\<^sub>\<tau>': "\<guillemotleft>?\<theta>\<^sub>\<tau>' : t\<^sub>0 \<star> ?w\<^sub>\<tau>' \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] : t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine \<Rightarrow> (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc'_in_hom by simp
show "\<guillemotleft>t\<^sub>0u\<^sub>1.\<phi> \<star> chine : (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
using \<tau>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine : (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine : ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> (u \<star> u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
using assoc_in_hom by auto
show "\<guillemotleft>\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] : (u \<star> u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> u \<star> (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by auto
show "\<guillemotleft>u \<star> the_\<theta> : u \<star> (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> ?u\<^sub>\<tau>\<guillemotright>"
by (intro hcomp_in_vhom, auto)
qed
have \<beta>\<^sub>\<tau>: "\<guillemotleft>?\<beta>\<^sub>\<tau> : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> (t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1 : (t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<omega>.the_\<nu>_props \<rho>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>the_\<nu> : r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using the_\<nu>_in_hom(2) by simp
show "\<guillemotleft>\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] : (t\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using t\<^sub>0u\<^sub>1.p\<^sub>1_simps assoc_in_hom by simp
qed
define LHS where "LHS = (t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)"
have LHS: "\<guillemotleft>LHS : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof (unfold LHS_def, intro comp_in_homI)
show "\<guillemotleft>\<tau> \<star> ?w\<^sub>\<tau> : t\<^sub>1 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t \<star> t\<^sub>0) \<star> ?w\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] : (t \<star> t\<^sub>0) \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>t \<star> ?\<theta>\<^sub>\<tau> : t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof -
have "src t = trg (t\<^sub>0 \<star> \<omega>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>1)"
by (metis \<chi>.u_simps(3) \<mu>.ide_base \<sigma>.ide_leg0 \<sigma>.leg1_simps(3) \<tau>\<mu>.composable
\<theta>\<^sub>\<tau> arrI assoc_simps(3) r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps trg_vcomp vconn_implies_hpar(2))
thus ?thesis
using \<theta>\<^sub>\<tau> by blast
qed
qed
define RHS where "RHS = ((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> ?\<beta>\<^sub>\<tau>"
have RHS: "\<guillemotleft>RHS : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
unfolding RHS_def
proof
show "\<guillemotleft>?\<beta>\<^sub>\<tau> : t\<^sub>1 \<star> ?w\<^sub>\<tau> \<Rightarrow> t\<^sub>1 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using \<beta>\<^sub>\<tau> by simp
show "\<guillemotleft>(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') : t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
proof
show "\<guillemotleft>\<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') : t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>'\<guillemotright>"
using \<tau>.T0.antipar(1) by fastforce
show "\<guillemotleft>t \<star> ?\<theta>\<^sub>\<tau>' : t \<star> t\<^sub>0 \<star> ?w\<^sub>\<tau>' \<Rightarrow> t \<star> ?u\<^sub>\<tau>\<guillemotright>"
using w\<^sub>\<tau>' \<theta>\<^sub>\<tau>' \<tau>.leg0_simps(2) \<tau>.leg0_simps(3) hseqI' ideD(1) t\<^sub>0u\<^sub>1.p\<^sub>1_simps
trg_hcomp \<tau>.base_in_hom(2) hcomp_in_vhom
by presburger
qed
qed
have eq: "LHS = RHS"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) = \<Delta>"
proof -
text \<open>
Here we use \<open>\<omega>.\<Delta>_naturality\<close> to replace @{term \<omega>.chine}
in favor of @{term \<omega>}.
We have to bring @{term \<omega>.the_\<nu>}, @{term \<tau>}, and @{term \<omega>.the_\<theta>} together,
with @{term \<rho>\<sigma>.p\<^sub>1} on the right.
\<close>
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
unfolding LHS_def
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1])"
using whisker_left \<tau>.ide_base \<theta>\<^sub>\<tau> arrI seqE
by (metis (full_types))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(\<tau> \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality
by (metis \<omega>.w_simps(2-6) \<rho>.leg1_simps(3) \<rho>\<sigma>.leg1_simps(2) \<tau>.tab_simps(1)
\<tau>.tab_simps(2,4-5) hseqE r\<^sub>0s\<^sub>1.leg1_simps(5) r\<^sub>0s\<^sub>1.leg1_simps(6))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
((t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot>
(\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot>
\<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "seq \<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
by (simp add: \<rho>.T0.antipar(1))
moreover have "inv (t \<star> \<a>[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) = t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using \<rho>.T0.antipar(1) by simp
ultimately show ?thesis
using pentagon \<rho>.T0.antipar(1) iso_hcomp
invert_side_of_triangle(1)
[of "\<a>[t, t\<^sub>0, \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1] \<cdot> \<a>[t \<star> t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
"t \<star> \<a>[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
"\<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1)"]
by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (((t \<star> \<omega>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
(\<a>[t, t\<^sub>0, \<omega>.chine] \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> ((\<tau> \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot>
(\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>[t, t\<^sub>0 \<star> \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] =
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((t \<star> \<omega>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality [of t \<omega>.the_\<theta> \<rho>\<sigma>.p\<^sub>1] \<omega>.\<theta>_simps(3) \<rho>\<sigma>.leg1_simps(2) hseq_char
by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right \<rho>.T0.antipar(1) \<omega>.\<Delta>_simps(1) \<omega>.\<Delta>_naturality comp_assoc
by fastforce
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> ((t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "t \<star> (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0 = (t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)"
using whisker_left whisker_right \<rho>.T0.antipar(1)
by (metis (full_types) \<chi>.\<Delta>_simps(1) \<tau>.ide_base \<theta>\<^sub>\<tau> arrI r\<^sub>0s\<^sub>1.ide_u seqE)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<cdot> \<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> (\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0) =
t \<star> \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>0)"
using \<chi>.in_hom whisker_left by auto
also have "... = t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]"
using assoc_naturality [of \<chi> s\<^sub>0 \<rho>\<sigma>.p\<^sub>0] \<chi>.in_hom by auto
also have "... = (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0])"
proof -
have "seq (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]"
using \<chi>.in_hom
apply (intro seqI hseqI)
apply auto
proof -
show "\<guillemotleft>\<chi> : src u \<rightarrow> trg \<chi>\<guillemotright>"
by (metis \<chi>.\<Delta>_simps(1) \<chi>.u_simps(3) hseqE in_hhom_def seqE)
show "dom (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) = s \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
by (metis \<Delta>_simps(1) \<chi>.in_hom hcomp_simps(1,3) hseq_char in_homE seqE
u_simps(4))
qed
thus ?thesis
using whisker_left by simp
qed
finally show ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(\<a>[t, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> ((\<omega> \<star> r\<^sub>0) \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot>
(\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1)) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using assoc_naturality [of \<omega> r\<^sub>0 \<rho>\<sigma>.p\<^sub>1] \<omega>.in_hom \<rho>.T0.antipar(1) comp_assoc
by fastforce
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> ((t \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(t \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (t \<star> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> (\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1) = \<omega> \<star> r\<^sub>0s\<^sub>1.\<phi>"
using comp_cod_arr comp_arr_dom \<omega>.in_hom interchange comp_ide_arr
by (metis \<tau>.base_in_hom(2) \<tau>.ide_base r\<^sub>0s\<^sub>1.\<phi>_simps(1) r\<^sub>0s\<^sub>1.\<phi>_simps(4) seqI')
also have "... = (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>)"
using r\<^sub>0s\<^sub>1.\<phi>_in_hom comp_cod_arr comp_arr_dom \<omega>.in_hom interchange
by (metis in_homE)
finally have "(t \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> (\<omega> \<star> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1) = (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> r\<^sub>0s\<^sub>1.\<phi>)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
((t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_left \<rho>.T0.antipar(1) \<rho>\<sigma>.composable \<chi>.in_hom comp_assoc by auto
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "(t \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (\<omega> \<star> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0) =
\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)"
proof -
have "\<guillemotleft>(\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) : s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> u \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
using \<omega>.in_hom \<chi>.in_hom by force
thus ?thesis
by (metis (no_types) \<omega>.in_hom comp_arr_dom comp_cod_arr in_homE
interchange)
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "\<omega> \<star> (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) =
(\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0))"
proof -
have "seq (\<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) (\<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0))"
using \<chi>.in_hom by force
thus ?thesis
using comp_arr_dom comp_cod_arr \<omega>.in_hom \<chi>.in_hom interchange
by (metis in_homE)
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (r \<star> \<a>[s, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<sigma> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
(r \<star> r\<^sub>0s\<^sub>1.\<phi>) \<cdot> \<a>[r, r\<^sub>0, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<rho> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<omega> \<star> \<chi> \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[r, s, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0]"
using assoc_naturality \<omega>.in_hom \<chi>.in_hom
by (metis \<rho>\<sigma>.leg0_simps(3) assoc'_naturality hcomp_in_vhomE in_hom in_homE
u_simps(2) u_simps(4) u_simps(5))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<Delta>"
using whisker_left \<rho>\<sigma>.tab_def comp_assoc by simp
finally show ?thesis by auto
qed
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
proof -
text \<open>Now cancel @{term \<omega>.the_\<nu>} and its inverse.\<close>
have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
(\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
unfolding RHS_def
using comp_assoc by presburger
also have "... = \<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
(\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu>"
proof -
have "the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
((t\<^sub>1 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using comp_inv_arr \<rho>.T0.antipar(1) comp_assoc_assoc' by simp
also have "... = the_\<nu> \<cdot> (inv \<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using comp_cod_arr \<rho>.T0.antipar(1) by simp
also have "... = the_\<nu> \<cdot> (r\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>1)"
using whisker_right [of \<rho>\<sigma>.p\<^sub>1] r\<^sub>0s\<^sub>1.ide_leg1 \<omega>.the_\<nu>_props(2) \<omega>.the_\<nu>_simps(4)
\<rho>.leg1_simps(2) comp_inv_arr'
by metis
also have "... = the_\<nu>"
using comp_arr_dom by simp
finally show ?thesis
using comp_assoc by simp
qed
text \<open>
Now reassociate to move @{term the_\<theta>} to the left and get other terms composed
with @{term chine}, where they can be reduced to @{term \<tau>\<mu>.tab}.
\<close>
also have "... = (\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot>
(t \<star> u \<star> the_\<theta>)) \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> (\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot>
\<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> the_\<nu>"
proof -
have "arr ((u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using \<theta>\<^sub>\<tau>' by blast
moreover have "arr (\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr ((\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr (((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
moreover have "arr ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using calculation by blast
ultimately
have "t \<star> (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> (t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] =
(t \<star> u \<star> the_\<theta>) \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine])"
using whisker_left \<rho>.T0.antipar(1) \<tau>.ide_base by presburger
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
((\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot> \<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> the_\<nu>"
using assoc'_naturality [of t u the_\<theta>] \<tau>\<mu>.composable \<theta>_simps(3) comp_assoc by auto
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(\<tau> \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine) \<cdot> \<a>[t\<^sub>1, \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine)"
using assoc_naturality
by (metis \<tau>.leg1_simps(3) \<tau>.tab_simps(1,2,4) \<tau>.tab_simps(5) \<tau>\<mu>.leg0_simps(2)
\<tau>\<mu>.leg1_simps(2) hseqE src_hcomp t\<^sub>0u\<^sub>1.leg1_simps(3,5-6) w_simps(2)
w_simps(4-6))
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot>
((t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot>
\<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine)"
using pentagon t\<^sub>0u\<^sub>1.p\<^sub>1_simps uw\<theta> \<tau>.T0.antipar(1) iso_hcomp
comp_assoc_assoc'
invert_side_of_triangle(1)
[of "\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1 \<star> chine] \<cdot> \<a>[t \<star> t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
"t \<star> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]"
"\<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine)"]
by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>[t, t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1, chine] =
\<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine)"
using assoc_naturality [of t t\<^sub>0u\<^sub>1.\<phi> chine] t\<^sub>0u\<^sub>1.cospan by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>[t, u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0, chine] =
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> ((t \<star> (\<mu> \<star> \<tau>\<mu>.p\<^sub>0)) \<star> chine)"
using assoc_naturality [of t "\<mu> \<star> \<tau>\<mu>.p\<^sub>0" chine]
by (simp add: \<tau>\<mu>.composable)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "(((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) =
((t \<star> ((u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0)) \<star> chine) \<cdot> ((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine)"
using whisker_right whisker_left [of t "\<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]" "\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]"]
\<tau>\<mu>.composable comp_assoc_assoc'
by simp
also have "... = (t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine"
using comp_cod_arr \<tau>\<mu>.composable by simp
finally have "(((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) =
(t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine"
by simp
thus ?thesis
using comp_assoc by metis
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (((\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
(\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot> the_\<nu>"
proof -
have "((\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> (\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot>
((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) =
((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine)"
using comp_inv_arr' comp_cod_arr \<tau>\<mu>.composable comp_assoc_assoc'
whisker_right [of chine "\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0]" "\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine)) \<cdot> the_\<nu>"
using comp_assoc by presburger
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot>
(\<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot>
(t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine)) \<cdot>
(\<tau>\<mu>.tab \<star> chine) \<cdot> the_\<nu>"
proof -
have "(\<a>\<^sup>-\<^sup>1[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> ((t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot>
((t \<star> \<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t \<star> t\<^sub>0u\<^sub>1.\<phi>) \<star> chine) \<cdot>
(\<a>[t, t\<^sub>0, \<tau>\<mu>.p\<^sub>1] \<star> chine) \<cdot> ((\<tau> \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) =
\<tau>\<mu>.tab \<star> chine"
using uw\<theta> whisker_right [of chine]
by (metis \<tau>\<mu>.tab_def \<tau>\<mu>.tab_in_vhom' arrI seqE)
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((t \<star> u) \<star> the_\<theta>) \<cdot> \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<tau>\<mu>.tab \<star> chine) \<cdot> the_\<nu>"
proof -
have "\<a>\<^sup>-\<^sup>1[t, u, (u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine] \<cdot> (t \<star> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]) \<cdot>
(t \<star> \<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot> \<a>[t, (u \<star> u\<^sub>0) \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot>
((t \<star> \<a>\<^sup>-\<^sup>1[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0]) \<star> chine) \<cdot> (\<a>[t, u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0] \<star> chine) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, \<^bold>\<langle>u\<^bold>\<rangle>, (\<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot> (\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>) \<^bold>\<cdot> \<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, (\<^bold>\<langle>u\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle>, \<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>chine\<^bold>\<rangle>)\<rbrace>"
using \<a>'_def \<alpha>_def \<tau>\<mu>.composable by simp
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>t\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>, \<^bold>\<langle>u\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>\<tau>\<mu>.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<tau>\<mu>.composable
apply (intro E.eval_eqI) by simp_all
also have "... = \<a>[t \<star> u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine]"
using \<a>'_def \<alpha>_def \<tau>\<mu>.composable by simp
finally show ?thesis by simp
qed
also have "... = \<Delta>"
using \<Delta>_naturality by simp
finally show ?thesis by simp
qed
finally have "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
by blast
(*
* TODO: This is common enough that there should be "cancel_iso_left" and
* "cancel_iso_right" rules for doing it.
*)
hence "(LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1) =
(RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (\<omega>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>1)"
using u\<^sub>\<tau> r\<^sub>0s\<^sub>1.ide_u LHS RHS iso_is_section [of "\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0]"] section_is_mono
monoE \<tau>\<mu>.composable comp_assoc
by (metis (no_types, lifting) \<Delta>_simps(1) \<mu>.ide_base
\<open>\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0] \<cdot> LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> r\<^sub>0s\<^sub>1.p\<^sub>1) =
((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab\<close>
\<tau>.ide_base hseq_char ideD(1) ide_u iso_assoc')
hence 1: "LHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] = RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using epiE LHS RHS iso_is_retraction retraction_is_epi \<tau>\<mu>.composable
\<omega>.the_\<nu>_props iso_hcomp
by (metis \<Delta>_simps(1) \<omega>.the_\<nu>_simps(2)
\<open>((\<omega> \<star> \<chi>) \<star> s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0) \<cdot> \<rho>\<sigma>.tab =
\<a>\<^sup>-\<^sup>1[t, u, s\<^sub>0 \<star> r\<^sub>0s\<^sub>1.p\<^sub>0] \<cdot> RHS \<cdot> \<a>[t\<^sub>1, \<omega>.chine, r\<^sub>0s\<^sub>1.p\<^sub>1] \<cdot> (\<omega>.the_\<nu> \<star> r\<^sub>0s\<^sub>1.p\<^sub>1)\<close>
\<rho>.leg1_simps(3) ide_is_iso local.comp_assoc r\<^sub>0s\<^sub>1.ide_leg1 r\<^sub>0s\<^sub>1.p\<^sub>1_simps seqE)
show "LHS = RHS"
proof -
have "epi \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using iso_is_retraction retraction_is_epi \<rho>.T0.antipar(1) by simp
moreover have "seq LHS \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using LHS \<rho>.T0.antipar(1) by auto
moreover have "seq RHS \<a>[t\<^sub>1, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]"
using RHS \<rho>.T0.antipar(1) by auto
ultimately show ?thesis
using epiE 1 by blast
qed
qed
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)"
using LHS_def RHS_def u\<^sub>\<tau> w\<^sub>\<tau> w\<^sub>\<tau>' \<beta>\<^sub>\<tau> \<theta>\<^sub>\<tau> \<theta>\<^sub>\<tau>' eq \<tau>.T2 [of ?w\<^sub>\<tau> ?w\<^sub>\<tau>' ?\<theta>\<^sub>\<tau> ?u\<^sub>\<tau> ?\<theta>\<^sub>\<tau>' ?\<beta>\<^sub>\<tau>]
by fastforce
obtain \<gamma>\<^sub>\<tau> where \<gamma>\<^sub>\<tau>: "\<guillemotleft>\<gamma>\<^sub>\<tau> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using 1 by auto
text \<open>
At this point we could show that @{term \<gamma>\<^sub>\<tau>} is invertible using \<open>BS3\<close>,
but we want to avoid using \<open>BS3\<close> if possible and we also want to
establish a characterization of @{term "inv \<gamma>\<^sub>\<tau>"}. So we show the invertibility of
@{term \<gamma>\<^sub>\<tau>} directly, using a few more applications of \<open>T2\<close>.
\<close>
have iso_\<beta>\<^sub>\<tau>: "iso ?\<beta>\<^sub>\<tau>"
using uw\<theta> \<beta>\<^sub>\<tau> the_\<nu>_props \<omega>.the_\<nu>_props hseqI' iso_assoc' \<omega>.hseq_leg\<^sub>0
apply (intro isos_compose)
apply (metis \<omega>.is_ide \<rho>\<sigma>.leg1_simps(2) \<tau>.ide_leg1 \<tau>.leg1_simps(2)
\<tau>.leg1_simps(3) hseqE r\<^sub>0s\<^sub>1.ide_leg1 hcomp_simps(1) vconn_implies_hpar(3))
apply (metis \<rho>\<sigma>.leg1_simps(2) hseqE ide_is_iso r\<^sub>0s\<^sub>1.ide_leg1 src_inv iso_inv_iso
iso_hcomp vconn_implies_hpar(1))
apply blast
apply blast
apply blast
apply (metis \<tau>.ide_leg1 \<tau>.leg1_simps(3) hseqE ide_char iso_assoc t\<^sub>0u\<^sub>1.ide_leg1
t\<^sub>0u\<^sub>1.p\<^sub>1_simps w\<^sub>\<tau>')
by blast
hence eq': "((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) =
((t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)) \<cdot> inv ?\<beta>\<^sub>\<tau>"
proof -
have "seq ((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) ?\<beta>\<^sub>\<tau>"
using LHS RHS_def eq by blast
hence "(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') =
(((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> ?\<beta>\<^sub>\<tau>) \<cdot> inv ?\<beta>\<^sub>\<tau>"
by (meson invert_side_of_triangle(2) iso_\<beta>\<^sub>\<tau>)
thus ?thesis
using LHS_def RHS_def eq by argo
qed
have 2: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> inv ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>)"
using u\<^sub>\<tau> w\<^sub>\<tau> w\<^sub>\<tau>' \<beta>\<^sub>\<tau> \<theta>\<^sub>\<tau> \<theta>\<^sub>\<tau>' eq' \<tau>.T2 [of ?w\<^sub>\<tau>' ?w\<^sub>\<tau> ?\<theta>\<^sub>\<tau>'?u\<^sub>\<tau> ?\<theta>\<^sub>\<tau> "inv ?\<beta>\<^sub>\<tau>"] iso_\<beta>\<^sub>\<tau> comp_assoc
by blast
obtain \<gamma>\<^sub>\<tau>' where
\<gamma>\<^sub>\<tau>': "\<guillemotleft>\<gamma>\<^sub>\<tau>' : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> inv ?\<beta>\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using 2 by auto
have "inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'"
proof
have "\<guillemotleft>\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' by auto
moreover have "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> = t\<^sub>1 \<star> ?w\<^sub>\<tau>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left \<beta>\<^sub>\<tau> iso_\<beta>\<^sub>\<tau> comp_inv_arr'
by (metis (no_types, lifting) \<tau>.ide_leg1 calculation in_homE)
moreover have "?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
proof -
have "?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau> by simp
also have "... = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>') \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau>' comp_assoc by simp
also have "... = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left
by (metis (full_types) \<tau>.ide_leg0 seqI')
finally show ?thesis by simp
qed
moreover have
"\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>) \<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>"
proof -
have "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> ?w\<^sub>\<tau> = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>)"
proof -
have "(t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>) =
((t \<star> ?\<theta>\<^sub>\<tau>) \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>)) \<cdot> (t\<^sub>1 \<star> ?w\<^sub>\<tau>)"
by (metis LHS LHS_def comp_arr_dom in_homE)
thus ?thesis
using w\<^sub>\<tau> \<theta>\<^sub>\<tau> \<omega>.w_simps(4) \<tau>.leg1_in_hom(2) \<tau>.leg1_simps(3) hcomp_in_vhom ideD(1)
trg_hcomp ide_in_hom(2) \<tau>.T2
by presburger
qed
thus "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau> \<Rightarrow> ?w\<^sub>\<tau>\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau> \<and> ?\<theta>\<^sub>\<tau> = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>) \<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>"
by (metis \<theta>\<^sub>\<tau> comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\<tau>)
qed
ultimately have "\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau> = ?w\<^sub>\<tau>"
by simp
thus "ide (\<gamma>\<^sub>\<tau>' \<cdot> \<gamma>\<^sub>\<tau>)"
using w\<^sub>\<tau> by simp
have "\<guillemotleft>\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright>"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' by auto
moreover have "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = t\<^sub>1 \<star> ?w\<^sub>\<tau>'"
by (metis \<beta>\<^sub>\<tau> \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' \<tau>.ide_leg1 calculation comp_arr_inv' in_homE iso_\<beta>\<^sub>\<tau> whisker_left)
moreover have "?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
proof -
have "?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau>' by simp
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>) \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau> comp_assoc by simp
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
using \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' whisker_left \<tau>.ide_leg0 by fastforce
finally show ?thesis by simp
qed
moreover have "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)
\<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>'"
proof -
have "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> ?w\<^sub>\<tau>' = t\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)"
proof -
have "(t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>') =
((t \<star> ?\<theta>\<^sub>\<tau>') \<cdot> \<a>[t, t\<^sub>0, ?w\<^sub>\<tau>'] \<cdot> (\<tau> \<star> ?w\<^sub>\<tau>')) \<cdot> (t\<^sub>1 \<star> ?w\<^sub>\<tau>')"
proof -
have 1: "t\<^sub>1 \<star> \<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>')"
by (meson \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' \<tau>.ide_leg1 seqI' whisker_left)
have "((LHS \<cdot> inv ?\<beta>\<^sub>\<tau>) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>)) \<cdot> (t\<^sub>1 \<star> \<gamma>\<^sub>\<tau>') = LHS \<cdot> inv ?\<beta>\<^sub>\<tau>"
using LHS_def RHS_def \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>' eq eq' by argo
thus ?thesis
unfolding LHS_def
using 1 by (simp add: calculation(2) eq' comp_assoc)
qed
thus ?thesis
using w\<^sub>\<tau>' \<theta>\<^sub>\<tau>' \<omega>.w_simps(4) \<tau>.leg1_in_hom(2) \<tau>.leg1_simps(3) hcomp_in_vhom ideD(1)
trg_hcomp ide_in_hom(2) \<tau>.T2 \<tau>.T0.antipar(1) t\<^sub>0u\<^sub>1.base_simps(2)
t\<^sub>0u\<^sub>1.leg1_simps(4)
by presburger
qed
thus "\<And>\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<tau>' \<Rightarrow> ?w\<^sub>\<tau>'\<guillemotright> \<and> t\<^sub>1 \<star> \<gamma> = t\<^sub>1 \<star> ?w\<^sub>\<tau>' \<and> ?\<theta>\<^sub>\<tau>' = ?\<theta>\<^sub>\<tau>' \<cdot> (t\<^sub>0 \<star> \<gamma>)
\<Longrightarrow> \<gamma> = ?w\<^sub>\<tau>'"
by (metis \<theta>\<^sub>\<tau>' comp_arr_dom ide_in_hom(2) in_homE w\<^sub>\<tau>')
qed
ultimately have "\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>' = ?w\<^sub>\<tau>'"
by simp
thus "ide (\<gamma>\<^sub>\<tau> \<cdot> \<gamma>\<^sub>\<tau>')"
using w\<^sub>\<tau>' by simp
qed
thus "\<tau>\<mu>.p\<^sub>1 \<star> chine \<cong> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
using w\<^sub>\<tau> w\<^sub>\<tau>' \<gamma>\<^sub>\<tau> isomorphic_symmetric isomorphic_def by blast
have iso_\<gamma>\<^sub>\<tau>: "iso \<gamma>\<^sub>\<tau>"
using \<open>inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'\<close> by auto
have \<gamma>\<^sub>\<tau>'_eq: "\<gamma>\<^sub>\<tau>' = inv \<gamma>\<^sub>\<tau>"
using \<open>inverse_arrows \<gamma>\<^sub>\<tau> \<gamma>\<^sub>\<tau>'\<close> inverse_unique by blast
let ?w\<^sub>\<mu> = "\<tau>\<mu>.p\<^sub>0 \<star> chine"
let ?w\<^sub>\<mu>' = "\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
let ?u\<^sub>\<mu> = "s\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>0"
let ?\<theta>\<^sub>\<mu> = "the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"
let ?\<theta>\<^sub>\<mu>' = "(\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
let ?\<beta>\<^sub>\<mu> = "\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot> (\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
have w\<^sub>\<mu>: "ide ?w\<^sub>\<mu> \<and> is_left_adjoint ?w\<^sub>\<mu>"
using is_map left_adjoints_compose by simp
have w\<^sub>\<mu>': "ide ?w\<^sub>\<mu>' \<and> is_left_adjoint ?w\<^sub>\<mu>'"
using \<chi>.is_map left_adjoints_compose
by (simp add: is_map left_adjoints_compose)
have 1: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>)"
proof -
have \<theta>\<^sub>\<mu>: "\<guillemotleft>?\<theta>\<^sub>\<mu> : u\<^sub>0 \<star> ?w\<^sub>\<mu> \<Rightarrow> ?u\<^sub>\<mu>\<guillemotright>"
by auto
have \<theta>\<^sub>\<mu>': "\<guillemotleft>?\<theta>\<^sub>\<mu>' : u\<^sub>0 \<star> ?w\<^sub>\<mu>' \<Rightarrow> ?u\<^sub>\<mu>\<guillemotright>"
by fastforce
have \<beta>\<^sub>\<mu>: "\<guillemotleft>?\<beta>\<^sub>\<mu> : u\<^sub>1 \<star> ?w\<^sub>\<mu> \<Rightarrow> u\<^sub>1 \<star> ?w\<^sub>\<mu>'\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] : u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0 \<star> chine \<Rightarrow> (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine\<guillemotright>"
by auto
show "\<guillemotleft>inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine : (u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine \<Rightarrow> (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine\<guillemotright>"
using t\<^sub>0u\<^sub>1.\<phi>_uniqueness(2)
by (simp add: t\<^sub>0u\<^sub>1.\<phi>_in_hom(2) w_in_hom(2))
show "\<guillemotleft>\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] : (t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine \<Rightarrow> t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine\<guillemotright>"
using \<tau>.T0.antipar(1) by auto
show "\<guillemotleft>t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau> : t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1 \<star> chine \<Rightarrow> t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<gamma>\<^sub>\<tau> iso_\<gamma>\<^sub>\<tau> using \<tau>.T0.antipar(1) by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] : t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1 : (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1\<guillemotright>"
using \<rho>.T0.antipar(1) by auto
show "\<guillemotleft>r\<^sub>0s\<^sub>1.\<phi> : r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1 \<Rightarrow> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
show "\<guillemotleft>\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0 : s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
show "\<guillemotleft>\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] : (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0 \<Rightarrow> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<guillemotright>"
by auto
qed
text \<open>
The proof of the equation below needs to make use of the equation
\<open>\<theta>\<^sub>\<tau>' = \<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')\<close> from the previous section. So the overall strategy is to
work toward an expression of the form \<open>\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> \<gamma>\<^sub>\<tau>')\<close> and perform the replacement
to eliminate \<open>\<gamma>\<^sub>\<tau>'\<close>.
\<close>
have eq\<^sub>\<mu>: "(u \<star> ?\<theta>\<^sub>\<mu>) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>) =
((u \<star> ?\<theta>\<^sub>\<mu>') \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>'] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>')) \<cdot> ?\<beta>\<^sub>\<mu>"
proof -
let ?LHS = "(u \<star> ?\<theta>\<^sub>\<mu>) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
let ?RHS = "((u \<star> ?\<theta>\<^sub>\<mu>') \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>'] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>')) \<cdot> ?\<beta>\<^sub>\<mu>"
have "?RHS = (u \<star> (\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = (u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0]) \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "u \<star> (\<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] =
(u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])"
using whisker_left \<mu>.ide_base \<theta>\<^sub>\<mu>' by blast
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "seq (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])
(\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0))"
by auto
moreover have "src u = trg \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
moreover have "inv (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) = u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
moreover have "iso (u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0])"
by simp
moreover have "iso \<a>[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
by simp
ultimately have "(u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0] =
\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
using pentagon hseqI' comp_assoc
invert_opposite_sides_of_square
[of "u \<star> \<a>[u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
"\<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0)"
"\<a>[u, u\<^sub>0, \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0]" "\<a>[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"]
inv_hcomp \<chi>.is_ide \<chi>.w_simps(3) \<chi>.w_simps(4) \<mu>.base_simps(2) \<mu>.ide_base
\<mu>.ide_leg0 \<mu>.leg0_simps(2) \<mu>.leg0_simps(3) \<sigma>.leg1_simps(3)
assoc'_eq_inv_assoc ide_hcomp r\<^sub>0s\<^sub>1.ide_u r\<^sub>0s\<^sub>1.p\<^sub>0_simps hcomp_simps(1)
by presburger
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot>
\<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "(u \<star> \<chi>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>[u, u\<^sub>0 \<star> \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] =
\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0)"
using assoc_naturality [of u \<chi>.the_\<theta> \<rho>\<sigma>.p\<^sub>0] \<chi>.\<theta>_simps(3) by auto
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> (\<mu> \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0) =
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]"
using assoc'_naturality [of \<mu> \<chi>.chine \<rho>\<sigma>.p\<^sub>0] by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by metis
also have "... = \<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> (((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> ((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
(\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0)) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1] \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "(\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0] \<cdot> \<a>[u\<^sub>1, \<chi>.chine, \<rho>\<sigma>.p\<^sub>0]) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0"
using comp_inv_arr' comp_cod_arr by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<a>[u, s\<^sub>0, \<rho>\<sigma>.p\<^sub>0] \<cdot> ((\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> r\<^sub>0s\<^sub>1.\<phi> \<cdot>
(\<omega>.the_\<theta> \<star> \<rho>\<sigma>.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[t\<^sub>0, \<omega>.chine, \<rho>\<sigma>.p\<^sub>1]) \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "arr ((u \<star> \<chi>.the_\<theta>) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine] \<cdot> (\<mu> \<star> \<chi>.chine) \<cdot> \<chi>.the_\<nu>)"
using \<chi>.\<theta>_simps(3) by simp
hence "((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
(u \<star> \<chi>.the_\<theta>) \<cdot> \<a>[u, u\<^sub>0, \<chi>.chine] \<cdot> (\<mu> \<star> \<chi>.chine) \<cdot> \<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0"
using whisker_right by simp
also have "... = (\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0"
using \<chi>.\<Delta>_naturality by simp
finally have "((u \<star> \<chi>.the_\<theta>) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<a>[u, u\<^sub>0, \<chi>.chine] \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot>
((\<mu> \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0) \<cdot> (\<chi>.the_\<nu> \<star> \<rho>\<sigma>.p\<^sub>0) =
(\<chi> \<star> s\<^sub>0) \<cdot> \<sigma> \<star> \<rho>\<sigma>.p\<^sub>0"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (?\<theta>\<^sub>\<tau> \<cdot> (t\<^sub>0 \<star> inv \<gamma>\<^sub>\<tau>)) \<cdot>
\<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = ?\<theta>\<^sub>\<tau>' \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine] \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using \<gamma>\<^sub>\<tau>' \<gamma>\<^sub>\<tau>'_eq by simp
also have "... = (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_assoc by presburger
also have "... = (u \<star> the_\<theta>) \<cdot> \<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
((\<mu> \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] =
\<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
proof -
have "((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((\<a>\<^sup>-\<^sup>1[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot> \<a>[t\<^sub>0, \<tau>\<mu>.p\<^sub>1, chine]) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine] =
((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> ((t\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>1) \<star> chine) \<cdot>
(inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_inv_arr' \<tau>.T0.antipar(1) by auto
also have "... = ((t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> (inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine)) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_cod_arr t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = (t\<^sub>0u\<^sub>1.\<phi> \<cdot> inv t\<^sub>0u\<^sub>1.\<phi> \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using whisker_right t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = ((u\<^sub>1 \<star> \<tau>\<mu>.p\<^sub>0) \<star> chine) \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_arr_inv' \<tau>.T0.antipar(1) t\<^sub>0u\<^sub>1.\<phi>_uniqueness by simp
also have "... = \<a>\<^sup>-\<^sup>1[u\<^sub>1, \<tau>\<mu>.p\<^sub>0, chine]"
using comp_cod_arr \<tau>.T0.antipar(1) by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = (u \<star> the_\<theta>) \<cdot> (\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using assoc'_naturality [of \<mu> \<tau>\<mu>.p\<^sub>0 chine] comp_assoc by auto
also have "... = ((u \<star> the_\<theta>) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine])) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using uw\<theta> pentagon comp_assoc
invert_opposite_sides_of_square
[of "u \<star> \<a>[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"
"\<a>[u, u\<^sub>0 \<star> \<tau>\<mu>.p\<^sub>0, chine] \<cdot> (\<a>[u, u\<^sub>0, \<tau>\<mu>.p\<^sub>0] \<star> chine)" "\<a>[u, u\<^sub>0, ?w\<^sub>\<mu>]"
"\<a>[u \<star> u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]"]
\<mu>.base_simps(2) \<mu>.ide_base \<mu>.ide_leg0 \<mu>.leg0_simps(2) assoc'_eq_inv_assoc
ide_hcomp hcomp_simps(1) t\<^sub>0u\<^sub>1.ide_u
by force
also have "... = (u \<star> the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[u\<^sub>0, \<tau>\<mu>.p\<^sub>0, chine]) \<cdot> \<a>[u, u\<^sub>0, ?w\<^sub>\<mu>] \<cdot> (\<mu> \<star> ?w\<^sub>\<mu>)"
using whisker_left comp_assoc by simp
finally show ?thesis by simp
qed
show "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>)"
using w\<^sub>\<mu> w\<^sub>\<mu>' \<theta>\<^sub>\<mu> \<theta>\<^sub>\<mu>' \<beta>\<^sub>\<mu> eq\<^sub>\<mu> \<mu>.T2 [of ?w\<^sub>\<mu> ?w\<^sub>\<mu>' ?\<theta>\<^sub>\<mu> ?u\<^sub>\<mu> ?\<theta>\<^sub>\<mu>' ?\<beta>\<^sub>\<mu>] by fast
qed
obtain \<gamma>\<^sub>\<mu> where \<gamma>\<^sub>\<mu>: "\<guillemotleft>\<gamma>\<^sub>\<mu> : ?w\<^sub>\<mu> \<Rightarrow> ?w\<^sub>\<mu>'\<guillemotright> \<and> ?\<beta>\<^sub>\<mu> = u\<^sub>1 \<star> \<gamma>\<^sub>\<mu> \<and> ?\<theta>\<^sub>\<mu> = ?\<theta>\<^sub>\<mu>' \<cdot> (u\<^sub>0 \<star> \<gamma>\<^sub>\<mu>)"
using 1 by auto
show "?w\<^sub>\<mu> \<cong> ?w\<^sub>\<mu>'"
using w\<^sub>\<mu> w\<^sub>\<mu>' \<gamma>\<^sub>\<mu> BS3 [of ?w\<^sub>\<mu> ?w\<^sub>\<mu>' \<gamma>\<^sub>\<mu> \<gamma>\<^sub>\<mu>] isomorphic_def by auto
qed
lemma comp_L:
shows "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> =
Maps.MkArr (src (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)) (src t) (Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>)"
proof -
show "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)"
using \<omega>.is_map r\<^sub>0s\<^sub>1.leg1_is_map left_adjoints_compose r\<^sub>0s\<^sub>1.p\<^sub>1_simps by auto
thus ?thesis
using Maps.CLS_in_hom r\<^sub>0s\<^sub>1.leg1_is_map
apply (intro Maps.seqI')
apply blast
using Maps.CLS_in_hom [of t\<^sub>0] \<tau>.leg0_is_map \<rho>\<sigma>.leg1_in_hom by auto
qed
thus "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> =
Maps.MkArr (src (\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1)) (src t) (Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>)"
using Maps.comp_char by auto
qed
lemma comp_R:
shows "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> =
Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>0\<rbrakk>)"
proof -
show "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0)"
using \<chi>.is_map r\<^sub>0s\<^sub>1.leg0_is_map left_adjoints_compose [of \<chi>.chine \<rho>\<sigma>.p\<^sub>0] by simp
thus ?thesis
using Maps.CLS_in_hom \<mu>.leg1_is_map
apply (intro Maps.seqI')
apply blast
using Maps.CLS_in_hom [of u\<^sub>1] \<mu>.leg1_is_map by simp
qed
thus "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> =
Maps.MkArr (src r\<^sub>0s\<^sub>1.p\<^sub>0) (trg u) (Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> r\<^sub>0s\<^sub>1.p\<^sub>0\<rbrakk>)"
using Maps.comp_char by auto
qed
lemma comp_L_eq_comp_R:
shows "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof (intro Maps.arr_eqI)
show "Maps.seq \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using comp_L(1) by simp
show "Maps.seq \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using comp_R(1) by simp
show "Maps.Dom (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Dom (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
by (metis (no_types, lifting) Maps.Dom.simps(1) \<omega>.w_simps(2) \<omega>.w_simps(3)
\<rho>.leg1_simps(3) \<rho>\<sigma>.leg1_in_hom(2) comp_L(2) comp_R(2) hcomp_in_vhomE hseqI'
r\<^sub>0s\<^sub>1.leg1_simps(3) hcomp_simps(1))
show "Maps.Cod (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Cod (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
by (metis Maps.Cod.simps(1) \<tau>\<mu>.composable comp_L(2) comp_R(2))
have A: "Maps.Map (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using comp_L(1) Maps.comp_char by auto
have B: "Maps.Map (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>) = Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using comp_R(1) Maps.comp_char by auto
have C: "Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk> = Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof (intro Maps.Comp_eqI)
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<in> Maps.Comp \<lbrakk>t\<^sub>0\<rbrakk> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
proof (intro Maps.in_CompI)
show "is_iso_class \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using prj_chine(2) is_iso_classI isomorphic_implies_hpar(2) by blast
show "is_iso_class \<lbrakk>t\<^sub>0\<rbrakk>"
using is_iso_classI by auto
show "\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<in> \<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using ide_in_iso_class prj_chine(2) isomorphic_implies_hpar(2) by blast
show "t\<^sub>0 \<in> \<lbrakk>t\<^sub>0\<rbrakk>"
using ide_in_iso_class by simp
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1"
using isomorphic_reflexive prj_chine(2) isomorphic_implies_hpar(2) by auto
qed
show "u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<in> Maps.Comp \<lbrakk>u\<^sub>1\<rbrakk> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof (intro Maps.in_CompI)
show "is_iso_class \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using is_iso_classI by simp
show "is_iso_class \<lbrakk>u\<^sub>1\<rbrakk>"
using is_iso_classI by simp
show "\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<in> \<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using ide_in_iso_class by simp
show "u\<^sub>1 \<in> iso_class u\<^sub>1"
using ide_in_iso_class by simp
show "u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0 \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
using isomorphic_reflexive isomorphic_implies_hpar(2) by auto
qed
show "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
proof -
have "t\<^sub>0 \<star> \<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1 \<cong> (t\<^sub>0 \<star> \<omega>.chine) \<star> \<rho>\<sigma>.p\<^sub>1"
using assoc'_in_hom [of t\<^sub>0 \<omega>.chine \<rho>\<sigma>.p\<^sub>1] iso_assoc' isomorphic_def r\<^sub>0s\<^sub>1.p\<^sub>1_simps
by auto
also have "... \<cong> r\<^sub>0 \<star> \<rho>\<sigma>.p\<^sub>1"
using \<omega>.leg0_uniquely_isomorphic hcomp_isomorphic_ide
by (simp add: \<rho>.T0.antipar(1))
also have "... \<cong> s\<^sub>1 \<star> \<rho>\<sigma>.p\<^sub>0"
using isomorphic_def r\<^sub>0s\<^sub>1.\<phi>_uniqueness(2) by blast
also have "... \<cong> (u\<^sub>1 \<star> \<chi>.chine) \<star> \<rho>\<sigma>.p\<^sub>0"
using \<chi>.leg1_uniquely_isomorphic hcomp_isomorphic_ide by auto
also have "... \<cong> u\<^sub>1 \<star> \<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0"
using assoc_in_hom [of u\<^sub>1 \<chi>.chine \<rho>\<sigma>.p\<^sub>0] iso_assoc isomorphic_def by auto
finally show ?thesis by simp
qed
qed
show "Maps.Map (\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>) = Maps.Map (\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>)"
using A B C by simp
qed
lemma csq:
shows "Maps.commutative_square \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof
show "Maps.cospan \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>"
using comp_L(1) comp_R(1) comp_L_eq_comp_R
by (metis (no_types, lifting) Maps.cod_comp Maps.seq_char)
show "Maps.span \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using comp_L(1) comp_R(1) comp_L_eq_comp_R
by (metis (no_types, lifting) Maps.dom_comp Maps.seq_char)
show "Maps.dom \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> = Maps.cod \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using comp_L(1) by auto
show "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using comp_L_eq_comp_R by simp
qed
lemma CLS_chine:
shows "\<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = Maps.tuple \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
let ?T = "Maps.tuple \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
have "\<exists>!l. \<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> l = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and> \<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> l = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using csq \<tau>\<mu>.prj_char
Maps.universal [of "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"]
by simp
moreover have "\<lbrakk>\<lbrakk>\<tau>\<mu>.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> ?T = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and>
\<lbrakk>\<lbrakk>\<tau>\<mu>.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> ?T = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using csq \<tau>\<mu>.prj_char
Maps.prj_tuple [of "\<lbrakk>\<lbrakk>t\<^sub>0\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>u\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>" "\<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"]
by simp
moreover have "\<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>\<omega>.chine \<star> \<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk> \<and>
\<lbrakk>\<lbrakk>t\<^sub>0u\<^sub>1.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>\<chi>.chine \<star> \<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using prj_chine \<tau>\<mu>.leg0_is_map \<tau>\<mu>.leg1_is_map is_map t\<^sub>0u\<^sub>1.leg1_is_map
t\<^sub>0u\<^sub>1.satisfies_T0 Maps.comp_CLS
by blast
ultimately show "\<lbrakk>\<lbrakk>chine\<rbrakk>\<rbrakk> = ?T" by auto
qed
end
subsection "Equivalence of B and Span(Maps(B))"
subsubsection "The Functor SPN"
text \<open>
We now define a function \<open>SPN\<close> on arrows and will ultimately show that it extends to a
biequivalence from the underlying bicategory \<open>B\<close> to \<open>Span(Maps(B))\<close>.
The idea is that \<open>SPN\<close> takes \<open>\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>\<close> to the isomorphism class of an induced arrow
of spans from the chosen tabulation of \<open>r\<close> to the chosen tabulation of \<open>s\<close>.
To obtain this, we first use isomorphisms \<open>r.tab\<^sub>1 \<star> r.tab\<^sub>0\<^sup>* \<cong> r\<close> and \<open>s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>* \<cong> s\<close>
to transform \<open>\<mu>\<close> to \<open>\<guillemotleft>\<mu>' : r.tab\<^sub>1 \<star> r.tab\<^sub>0\<^sup>* \<Rightarrow> s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>*\<guillemotright>\<close>.
We then take the adjoint transpose of \<open>\<mu>'\<close> to obtain
\<open>\<guillemotleft>\<omega> : r.tab\<^sub>1 \<Rightarrow> (s.tab\<^sub>1 \<star> s.tab\<^sub>0\<^sup>*) \<star> r.tab\<^sub>0\<guillemotright>\<close>. The 2-cell \<open>\<omega>\<close> induces a map \<open>w\<close>
which is an arrow of spans from \<open>(r.tab\<^sub>0, r.tab\<^sub>1)\<close> to \<open>(s.tab\<^sub>0, s.tab\<^sub>1)\<close>.
We take the arrow of \<open>Span(Maps(B))\<close> defined by \<open>w\<close> as the value of \<open>SPN \<mu>\<close>.
Ensuring that \<open>SPN\<close> is functorial is a somewhat delicate point, which requires that all
the underlying definitions that have been set up are ``just so'', with no extra choices
other than those that are forced, and with the tabulation assigned to each 1-cell \<open>r\<close> in
the proper relationship with the canonical tabulation assigned to its chosen factorization
\<open>r = g \<star> f\<^sup>*\<close>.
\<close>
context bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 50)
definition spn
where "spn \<mu> \<equiv>
arrow_of_tabulations_in_maps.chine V H \<a> \<i> src trg
(tab_of_ide (dom \<mu>)) (tab\<^sub>0 (dom \<mu>)) (cod \<mu>)
(tab_of_ide (cod \<mu>)) (tab\<^sub>0 (cod \<mu>)) (tab\<^sub>1 (cod \<mu>)) \<mu>"
lemma is_induced_map_spn:
assumes "arr \<mu>"
shows "arrow_of_tabulations_in_maps.is_induced_map V H \<a> \<i> src trg
(tab_of_ide (dom \<mu>)) (tab\<^sub>0 (dom \<mu>)) (cod \<mu>)
(tab_of_ide (cod \<mu>)) (tab\<^sub>0 (cod \<mu>)) (tab\<^sub>1 (cod \<mu>))
\<mu> (spn \<mu>)"
proof -
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using assms by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<mu>.s.tab \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
show ?thesis
unfolding spn_def
using \<mu>.chine_is_induced_map by blast
qed
lemma spn_props:
assumes "arr \<mu>"
shows "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
and "is_left_adjoint (spn \<mu>)"
and "tab\<^sub>0 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>0 (dom \<mu>)"
and "tab\<^sub>1 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>1 (dom \<mu>)"
proof -
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using assms by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<mu>.s.tab \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
show "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
using spn_def by simp
show "is_left_adjoint (spn \<mu>)"
using spn_def by (simp add: \<mu>.is_map)
show "tab\<^sub>0 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>0 (dom \<mu>)"
using spn_def isomorphic_def \<mu>.leg0_uniquely_isomorphic(1) by auto
show "tab\<^sub>1 (cod \<mu>) \<star> spn \<mu> \<cong> tab\<^sub>1 (dom \<mu>)"
using spn_def isomorphic_def isomorphic_symmetric
\<mu>.leg1_uniquely_isomorphic(1)
by auto
qed
lemma spn_in_hom [intro]:
assumes "arr \<mu>"
shows "\<guillemotleft>spn \<mu> : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src (tab\<^sub>0 (cod \<mu>))\<guillemotright>"
and "\<guillemotleft>spn \<mu> : spn \<mu> \<Rightarrow> spn \<mu>\<guillemotright>"
using assms spn_props left_adjoint_is_ide by auto
lemma spn_simps [simp]:
assumes "arr \<mu>"
shows "is_left_adjoint (spn \<mu>)"
and "ide (spn \<mu>)"
and "src (spn \<mu>) = src (tab\<^sub>0 (dom \<mu>))"
and "trg (spn \<mu>) = src (tab\<^sub>0 (cod \<mu>))"
using assms spn_props left_adjoint_is_ide by auto
text \<open>
We need the next result to show that \<open>SPN\<close> is functorial; in particular,
that it takes \<open>\<guillemotleft>r : r \<Rightarrow> r\<guillemotright>\<close> in the underlying bicategory to a 1-cell
in \<open>Span(Maps(B))\<close>. The 1-cells in \<open>Span(Maps(B))\<close> have objects of \<open>Maps(B)\<close>
as their chines, and objects of \<open>Maps(B)\<close> are isomorphism classes of objects in the
underlying bicategory \<open>B\<close>. So we need the induced map associated with \<open>r\<close> to be isomorphic
to an object.
\<close>
lemma spn_ide:
assumes "ide r"
shows "spn r \<cong> src (tab\<^sub>0 r)"
proof -
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r
using assms by unfold_locales auto
interpret r: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> r
using r.is_arrow_of_tabulations_in_maps by simp
interpret tab: tabulation V H \<a> \<i> src trg r \<open>r.tab\<close> \<open>tab\<^sub>0 r\<close> \<open>dom r.tab\<close>
using assms r.tab_is_tabulation by simp
interpret tab: tabulation_in_maps V H \<a> \<i> src trg r \<open>r.tab\<close> \<open>tab\<^sub>0 r\<close> \<open>dom r.tab\<close>
by (unfold_locales, simp_all)
have "tab.is_induced_by_cell (spn r) (tab\<^sub>0 r) r.tab"
using spn_def comp_ide_arr r.chine_is_induced_map by auto
thus ?thesis
using tab.induced_map_unique [of "tab\<^sub>0 r" "r.tab" "spn r" "src r.s\<^sub>0"]
tab.apex_is_induced_by_cell
by (simp add: comp_assoc)
qed
text \<open>
The other key result we need to show that \<open>SPN\<close> is functorial is to show
that the induced map of a composite is isomorphic to the composite of
induced maps.
\<close>
lemma spn_hcomp:
assumes "seq \<tau> \<mu>" and "g \<cong> spn \<tau>" and "f \<cong> spn \<mu>"
shows "spn (\<tau> \<cdot> \<mu>) \<cong> g \<star> f"
proof -
interpret \<tau>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<tau>\<close> \<open>cod \<tau>\<close> \<tau>
using assms by unfold_locales auto
interpret \<tau>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<open>cod \<tau>\<close> \<tau>.s.tab \<open>tab\<^sub>0 (cod \<tau>)\<close> \<open>tab\<^sub>1 (cod \<tau>)\<close>
\<tau>
using \<tau>.is_arrow_of_tabulations_in_maps by simp
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>dom \<tau>\<close> \<mu>
using assms apply unfold_locales
apply auto[1]
by (elim seqE, auto)
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
interpret \<tau>\<mu>: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<mu>.r.tab \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>dom \<tau>\<close> \<tau>.r.tab \<open>tab\<^sub>0 (dom \<tau>)\<close> \<open>tab\<^sub>1 (dom \<tau>)\<close>
\<open>cod \<tau>\<close> \<tau>.s.tab \<open>tab\<^sub>0 (cod \<tau>)\<close> \<open>tab\<^sub>1 (cod \<tau>)\<close>
\<mu> \<tau>
..
have "g \<cong> \<tau>.chine"
using assms(2) spn_def by auto
moreover have "f \<cong> \<mu>.chine"
using assms(1) assms(3) spn_def by auto
moreover have "src g = trg f"
using calculation(1-2) isomorphic_implies_hpar(3-4) by auto
moreover have "src g = trg \<mu>.chine"
using calculation(1) isomorphic_implies_hpar(3) by auto
ultimately have "g \<star> f \<cong> \<tau>.chine \<star> \<mu>.chine"
using hcomp_ide_isomorphic hcomp_isomorphic_ide isomorphic_transitive
by (meson \<mu>.is_ide isomorphic_implies_ide(1))
also have "... \<cong> spn (\<tau> \<cdot> \<mu>)"
using spn_def \<tau>\<mu>.chine_char isomorphic_symmetric
by (metis \<tau>\<mu>.in_hom in_homE)
finally show ?thesis
using isomorphic_symmetric by simp
qed
abbreviation (input) SPN\<^sub>0
where "SPN\<^sub>0 r \<equiv> Span.mkIde \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
definition SPN
where "SPN \<mu> \<equiv> if arr \<mu> then
\<lparr>Chn = \<lbrakk>\<lbrakk>spn \<mu>\<rbrakk>\<rbrakk>,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<rparr>
else Span.null"
lemma Dom_SPN [simp]:
assumes "arr \<mu>"
shows "Dom (SPN \<mu>) = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>"
using assms SPN_def by simp
lemma Cod_SPN [simp]:
assumes "arr \<mu>"
shows "Cod (SPN \<mu>) = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>"
using assms SPN_def by simp
text \<open>Now we have to show this does the right thing for us.\<close>
lemma SPN_in_hom:
assumes "arr \<mu>"
shows "Span.in_hom (SPN \<mu>) (SPN\<^sub>0 (dom \<mu>)) (SPN\<^sub>0 (cod \<mu>))"
proof
interpret Dom: span_in_category Maps.comp \<open>Dom (SPN \<mu>)\<close>
proof
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close>
using assms by unfold_locales auto
show "Maps.span (Leg0 (Dom (SPN \<mu>))) (Leg1 (Dom (SPN \<mu>)))"
using assms Maps.CLS_in_hom SPN_def
by (metis (no_types, lifting) Maps.in_homE bicategory_of_spans.Dom_SPN
bicategory_of_spans_axioms r.leg1_is_map r.leg1_simps(3) r.satisfies_T0
span_data.simps(1) span_data.simps(2))
qed
interpret Dom': span_in_category Maps.comp
\<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<close>
using assms Dom.span_in_category_axioms SPN_def by simp
interpret Cod: span_in_category Maps.comp \<open>Cod (SPN \<mu>)\<close>
proof
interpret s: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod \<mu>\<close>
using assms by unfold_locales auto
show "Maps.span (Leg0 (Cod (SPN \<mu>))) (Leg1 (Cod (SPN \<mu>)))"
using assms Maps.CLS_in_hom SPN_def
by (metis (no_types, lifting) bicategory_of_spans.Cod_SPN bicategory_of_spans_axioms
ide_dom s.base_simps(2) s.base_simps(3) s.determines_span span_in_category.is_span)
qed
interpret Cod': span_in_category Maps.comp
\<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (cod \<mu>)\<rbrakk>\<rbrakk>\<rparr>\<close>
using assms Cod.span_in_category_axioms SPN_def by simp
show 1: "Span.arr (SPN \<mu>)"
proof (unfold Span.arr_char)
show "arrow_of_spans Maps.comp (SPN \<mu>)"
proof (unfold_locales)
show "Maps.in_hom (Chn (SPN \<mu>)) Dom.apex Cod.apex"
unfolding SPN_def Maps.in_hom_char
using assms Dom'.apex_def Cod'.apex_def Dom'.is_span Cod'.is_span Maps.arr_char
by auto
show "Cod.leg0 \<odot> Chn (SPN \<mu>) = Dom.leg0"
unfolding SPN_def
using assms spn_props [of \<mu>] Maps.comp_CLS [of "tab\<^sub>0 (cod \<mu>)" "spn \<mu>"] by simp
show "Cod.leg1 \<odot> Chn (SPN \<mu>) = Dom.leg1"
unfolding SPN_def
using assms spn_props [of \<mu>] Maps.comp_CLS [of "tab\<^sub>1 (cod \<mu>)" "spn \<mu>"] by simp
qed
qed
show "Span.dom (SPN \<mu>) = SPN\<^sub>0 (dom \<mu>)"
using assms 1 Span.dom_char Dom'.apex_def SPN_def by simp
show "Span.cod (SPN \<mu>) = SPN\<^sub>0 (cod \<mu>)"
using assms 1 Span.cod_char Cod'.apex_def SPN_def by simp
qed
interpretation SPN: "functor" V Span.vcomp SPN
proof
show "\<And>\<mu>. \<not> arr \<mu> \<Longrightarrow> SPN \<mu> = Span.null"
unfolding SPN_def by simp
show 1: "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.arr (SPN \<mu>)"
using SPN_in_hom by auto
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.dom (SPN \<mu>) = SPN (dom \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>)"
proof -
have "src (tab\<^sub>0 (dom \<mu>)) \<in> Collect obj"
using \<mu> by simp
moreover have "src \<mu> \<in> Collect obj"
using \<mu> by simp
moreover have "\<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (local.dom \<mu>))) (src \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>0 (dom \<mu>) : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src \<mu>\<guillemotright>"
using \<mu> by simp
moreover have "is_left_adjoint (tab\<^sub>0 (dom \<mu>))"
using \<mu> tab\<^sub>0_simps [of "dom \<mu>"] by auto
ultimately show ?thesis by auto
qed
ultimately show ?thesis by simp
qed
have "\<lbrakk>spn (dom \<mu>)\<rbrakk> = \<lbrakk>src (tab\<^sub>0 (dom \<mu>))\<rbrakk>"
using \<mu> spn_ide iso_class_eqI by auto
hence "SPN (dom \<mu>) = SPN\<^sub>0 (dom \<mu>)"
unfolding SPN_def
using \<mu> 1 Maps.dom_char by simp
thus "Span.dom (SPN \<mu>) = SPN (dom \<mu>)"
using \<mu> SPN_in_hom by auto
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.cod (SPN \<mu>) = SPN (cod \<mu>)"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
have 1: "Maps.arr (Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk>)"
proof -
have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using \<mu> by simp
moreover have "src \<mu> \<in> Collect obj"
using \<mu> by simp
moreover have "\<lbrakk>tab\<^sub>0 (cod \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (cod \<mu>))) (src \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>0 (cod \<mu>) : src (tab\<^sub>0 (cod \<mu>)) \<rightarrow> src \<mu>\<guillemotright>"
using \<mu> by simp
moreover have "is_left_adjoint (tab\<^sub>0 (cod \<mu>))"
using \<mu> by simp
ultimately show ?thesis by auto
qed
ultimately show ?thesis by simp
qed
have "\<lbrakk>spn (cod \<mu>)\<rbrakk> = \<lbrakk>src (tab\<^sub>0 (cod \<mu>))\<rbrakk>"
using \<mu> spn_ide iso_class_eqI by auto
hence "SPN (cod \<mu>) = SPN\<^sub>0 (cod \<mu>)"
unfolding SPN_def
using \<mu> 1 Maps.dom_char by simp
thus "Span.cod (SPN \<mu>) = SPN (cod \<mu>)"
using \<mu> SPN_in_hom by auto
qed
show "\<And>\<nu> \<mu>. seq \<nu> \<mu> \<Longrightarrow> SPN (\<nu> \<cdot> \<mu>) = SPN \<nu> \<bullet> SPN \<mu>"
proof -
fix \<mu> \<nu>
assume seq: "seq \<nu> \<mu>"
have "Dom (SPN (\<nu> \<cdot> \<mu>)) = Dom (SPN \<nu> \<bullet> SPN \<mu>)"
using seq 1 Span.vcomp_def Span.arr_char
by (elim seqE, simp)
moreover have "Cod (SPN (\<nu> \<cdot> \<mu>)) = Cod (SPN \<nu> \<bullet> SPN \<mu>)"
using seq 1 Span.vcomp_def Span.arr_char
by (elim seqE, simp)
moreover have "Chn (SPN (\<nu> \<cdot> \<mu>)) = Chn (SPN \<nu> \<bullet> SPN \<mu>)"
proof -
have *: "\<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk> = Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof
show "\<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk> \<subseteq> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof
fix h
assume h: "h \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
show "h \<in> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
proof -
have 1: "spn \<nu> \<in> \<lbrakk>spn \<nu>\<rbrakk>"
using seq ide_in_iso_class by auto
moreover have 2: "spn \<mu> \<in> \<lbrakk>spn \<mu>\<rbrakk>"
using seq ide_in_iso_class by auto
moreover have "spn \<nu> \<star> spn \<mu> \<cong> h"
proof -
have "spn \<nu> \<star> spn \<mu> \<cong> spn (\<nu> \<cdot> \<mu>)"
using seq spn_hcomp 1 2 iso_class_def isomorphic_reflexive
isomorphic_symmetric
by simp
thus ?thesis
using h isomorphic_transitive iso_class_def by simp
qed
ultimately show ?thesis
unfolding Maps.Comp_def
by (metis (mono_tags, lifting) is_iso_classI spn_simps(2)
mem_Collect_eq seq seqE)
qed
qed
show "Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk> \<subseteq> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
proof
fix h
assume h: "h \<in> Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>"
show "h \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
proof -
obtain f g where 1: "g \<in> \<lbrakk>spn \<nu>\<rbrakk> \<and> f \<in> \<lbrakk>spn \<mu>\<rbrakk> \<and> g \<star> f \<cong> h"
using h Maps.Comp_def [of "iso_class (spn \<nu>)" "iso_class (spn \<mu>)"]
iso_class_def iso_class_elems_isomorphic
by blast
have fg: "g \<cong> spn \<nu> \<and> f \<cong> spn \<mu> \<and> g \<star> f \<cong> h"
proof -
have "spn \<nu> \<in> \<lbrakk>spn \<nu>\<rbrakk> \<and> spn \<mu> \<in> \<lbrakk>spn \<mu>\<rbrakk>"
using seq ide_in_iso_class by auto
thus ?thesis
using 1 iso_class_elems_isomorphic isomorphic_symmetric is_iso_classI
by (meson spn_simps(2) seq seqE)
qed
have "g \<star> f \<in> \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
using seq fg 1 spn_hcomp iso_class_def isomorphic_symmetric by simp
thus ?thesis
using fg isomorphic_transitive iso_class_def by blast
qed
qed
qed
have "Chn (SPN \<nu> \<bullet> SPN \<mu>) =
Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk> \<odot>
Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>"
using 1 seq SPN_def Span.vcomp_def Span.arr_char
apply (elim seqE)
apply simp
by (metis (no_types, lifting) seq vseq_implies_hpar(1) vseq_implies_hpar(2))
also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<nu>)))
(Maps.Comp \<lbrakk>spn \<nu>\<rbrakk> \<lbrakk>spn \<mu>\<rbrakk>)"
proof -
have "Maps.seq (Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>)"
proof
show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (local.dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>)))
\<lbrakk>spn \<mu>\<rbrakk>)
(Maps.MkIde (src (tab\<^sub>0 (dom \<mu>))))
(Maps.MkIde (src (tab\<^sub>0 (cod \<mu>))))"
proof -
have "src (tab\<^sub>0 (dom \<mu>)) \<in> Collect obj"
using in_hhom_def seq by auto
moreover have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using seq by auto
moreover have "\<lbrakk>spn \<mu>\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>)))"
using spn_props
by (metis (mono_tags, lifting) mem_Collect_eq seq seqE)
ultimately show ?thesis
using Maps.MkArr_in_hom by simp
qed
show "Maps.in_hom (Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>)
(Maps.MkIde (src (tab\<^sub>0 (cod \<mu>))))
(Maps.MkIde (src (tab\<^sub>0 (cod \<nu>))))"
proof -
have "src (tab\<^sub>0 (cod \<mu>)) \<in> Collect obj"
using in_hhom_def seq by auto
moreover have "src (tab\<^sub>0 (cod \<nu>)) \<in> Collect obj"
using seq by auto
moreover have "\<lbrakk>spn \<nu>\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>)))"
using spn_props
by (metis (mono_tags, lifting) mem_Collect_eq seq seqE)
ultimately show ?thesis
using Maps.MkArr_in_hom by simp
qed
qed
thus ?thesis
using Maps.comp_char
[of "Maps.MkArr (src (tab\<^sub>0 (cod \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn \<nu>\<rbrakk>"
"Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<mu>))) \<lbrakk>spn \<mu>\<rbrakk>"]
by simp
qed
also have "... = Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src (tab\<^sub>0 (cod \<nu>))) \<lbrakk>spn (\<nu> \<cdot> \<mu>)\<rbrakk>"
using * by simp
also have "... = Chn (SPN (\<nu> \<cdot> \<mu>))"
using seq SPN_def Span.vcomp_def
by (elim seqE, simp)
finally show ?thesis by simp
qed
ultimately show "SPN (\<nu> \<cdot> \<mu>) = SPN \<nu> \<bullet> SPN \<mu>" by simp
qed
qed
lemma SPN_is_functor:
shows "functor V Span.vcomp SPN"
..
interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
proof
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.isomorphic (SPN (src \<mu>)) (Span.src (SPN \<mu>))"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
let ?src = "Maps.MkIde (src \<mu>)"
have src: "Maps.ide ?src"
using \<mu> by simp
interpret src: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>src \<mu>\<close>
using \<mu> by unfold_locales auto
interpret src: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>src \<mu>\<close> src.tab \<open>tab\<^sub>0 (src \<mu>)\<close> \<open>tab\<^sub>1 (src \<mu>)\<close>
\<open>src \<mu>\<close> src.tab \<open>tab\<^sub>0 (src \<mu>)\<close> \<open>tab\<^sub>1 (src \<mu>)\<close>
\<open>src \<mu>\<close>
using src.is_arrow_of_tabulations_in_maps by simp
interpret src: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>\<close>
using src by (unfold_locales, simp)
let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
have tab\<^sub>0_src: "\<guillemotleft>tab\<^sub>0 (src \<mu>) : src (tab\<^sub>0 (src \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (src \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>0: "Maps.arr ?tab\<^sub>0"
using \<mu> Maps.arr_MkArr tab\<^sub>0_src by blast
let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
have tab\<^sub>1_src: "\<guillemotleft>tab\<^sub>1 (src \<mu>) : src (tab\<^sub>0 (src \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (src \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>1: "Maps.arr ?tab\<^sub>1"
using \<mu> Maps.arr_MkArr tab\<^sub>1_src by blast
interpret tab: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<close>
using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)
have "src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>0 (src \<mu>)"
using \<mu> iso_lunit isomorphic_def
by (metis lunit_in_hom(2) src.ide_u src.u_simps(3) src_src)
hence "src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>1 (src \<mu>)"
using \<mu> src.obj_has_symmetric_tab isomorphic_transitive by blast
have "\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)"
using \<mu> tab\<^sub>0_src by blast
have "\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)"
proof -
have "\<guillemotleft>src \<mu> : src \<mu> \<rightarrow> src \<mu>\<guillemotright> \<and> is_left_adjoint (src \<mu>) \<and> \<lbrakk>src \<mu>\<rbrakk> = \<lbrakk>src \<mu>\<rbrakk>"
using \<mu> obj_is_self_adjoint by simp
thus ?thesis by auto
qed
interpret SPN_src: arrow_of_spans Maps.comp \<open>SPN (src \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have SPN_src: "SPN (src \<mu>) =
\<lparr>Chn = Maps.MkArr (src (tab\<^sub>0 (src \<mu>))) (src (tab\<^sub>0 (src \<mu>))) \<lbrakk>spn (src \<mu>)\<rbrakk>,
Dom = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<rparr>"
unfolding SPN_def using \<mu> by simp
interpret src_SPN: arrow_of_spans Maps.comp \<open>Span.src (SPN \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have src_SPN: "Span.src (SPN \<mu>) =
\<lparr>Chn = ?src,
Dom = \<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>,
Cod = \<lparr>Leg0 = ?src, Leg1 = ?src\<rparr>\<rparr>"
proof -
let ?tab\<^sub>0_dom = "Maps.MkArr (src (tab\<^sub>0 (dom \<mu>))) (src \<mu>) \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>"
have "Maps.arr ?tab\<^sub>0_dom"
proof -
have "\<guillemotleft>tab\<^sub>0 (dom \<mu>) : src (tab\<^sub>0 (dom \<mu>)) \<rightarrow> src \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (dom \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (dom \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis
using \<mu> Maps.arr_MkArr by blast
qed
thus ?thesis
using \<mu> Maps.cod_char Span.src_def by simp
qed
text \<open>
The idea of the proof is that @{term "iso_class (tab\<^sub>0 (src \<mu>))"} is invertible
in \<open>Maps(B)\<close> and determines an invertible arrow of spans from @{term "SPN (src \<mu>)"}
to @{term "Span.src (SPN \<mu>)"}.
\<close>
let ?\<phi> = "\<lparr>Chn = ?tab\<^sub>0, Dom = Dom (SPN (src \<mu>)), Cod = Cod (Span.src (SPN \<mu>))\<rparr>"
interpret \<phi>: arrow_of_spans Maps.comp ?\<phi>
apply (unfold_locales, simp_all)
proof -
show "Maps.in_hom ?tab\<^sub>0 SPN_src.dom.apex src_SPN.cod.apex"
using \<mu> tab\<^sub>0 Maps.dom_char Maps.cod_char SPN_src src_SPN
tab.apex_def src_SPN.cod.apex_def
apply (intro Maps.in_homI) by simp_all
show "src_SPN.cod.leg0 \<odot> ?tab\<^sub>0 = SPN_src.dom.leg0"
proof -
have "Maps.seq src_SPN.cod.leg0 ?tab\<^sub>0"
using \<mu> src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>0 (src \<mu>) \<in> Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "src \<mu>"]
ide_in_iso_class [of "tab\<^sub>0 (src \<mu>)"] \<open>src \<mu> \<star> tab\<^sub>0 (src \<mu>) \<cong> tab\<^sub>0 (src \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)\<close>
\<open>\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of src_SPN.cod.leg0 ?tab\<^sub>0] src_SPN by simp
qed
show "src_SPN.cod.leg1 \<odot> ?tab\<^sub>0 = SPN_src.dom.leg1"
proof -
have "Maps.seq src_SPN.cod.leg1 ?tab\<^sub>0"
using \<mu> src_SPN tab\<^sub>0 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (src \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (src \<mu>) \<in> Maps.Comp \<lbrakk>src \<mu>\<rbrakk> \<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "src \<mu>"]
ide_in_iso_class [of "tab\<^sub>0 (src \<mu>)"]
\<open>isomorphic (src \<mu> \<star> tab\<^sub>0 (src \<mu>)) (tab\<^sub>1 (src \<mu>))\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>0 (src \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (src \<mu>))) (src \<mu>)\<close>
\<open>\<lbrakk>src \<mu>\<rbrakk> \<in> Maps.Hom (src \<mu>) (src \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of src_SPN.cod.leg1 ?tab\<^sub>0] src_SPN by simp
qed
qed
have "Span.in_hom ?\<phi> (SPN (src \<mu>)) (Span.src (SPN \<mu>))"
using \<mu> tab\<^sub>0 spn_ide [of "src \<mu>"] iso_class_eqI
Span.arr_char Span.dom_char Span.cod_char \<phi>.arrow_of_spans_axioms
SPN_src src_SPN src.apex_def tab.apex_def Maps.dom_char
apply (intro Span.in_homI) by auto
(* The preceding cannot be written "by (intro Span.in_homI, auto)", why? *)
moreover have "Maps.iso ?tab\<^sub>0"
using \<mu> tab\<^sub>0 ide_in_iso_class src.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint
Maps.iso_char' [of ?tab\<^sub>0]
by auto
ultimately show "Span.isomorphic (SPN (src \<mu>)) (Span.src (SPN \<mu>))"
using Span.isomorphic_def Span.iso_char by auto
qed
show "\<And>\<mu>. arr \<mu> \<Longrightarrow> Span.isomorphic (SPN (trg \<mu>)) (Span.trg (SPN \<mu>))"
proof -
fix \<mu>
assume \<mu>: "arr \<mu>"
let ?trg = "Maps.MkIde (trg \<mu>)"
have trg: "Maps.ide ?trg"
using \<mu> by simp
interpret trg: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>trg \<mu>\<close>
using \<mu> by unfold_locales auto
interpret trg: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>\<close>
using trg by (unfold_locales, simp)
let ?tab\<^sub>0 = "Maps.MkArr (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
have tab\<^sub>0_trg: "\<guillemotleft>tab\<^sub>0 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>0 (trg \<mu>)) \<and> \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>0: "Maps.arr ?tab\<^sub>0"
using \<mu> Maps.arr_MkArr tab\<^sub>0_trg by blast
let ?tab\<^sub>1 = "Maps.MkArr (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
have tab\<^sub>1_trg: "\<guillemotleft>tab\<^sub>1 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (trg \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
have tab\<^sub>1: "Maps.arr ?tab\<^sub>1"
using \<mu> Maps.arr_MkArr tab\<^sub>1_trg by blast
interpret tab: span_in_category Maps.comp \<open>\<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<close>
using tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char by (unfold_locales, simp)
have "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>0 (trg \<mu>)"
proof -
have "\<guillemotleft>\<l>[tab\<^sub>1 (trg \<mu>)] : trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<Rightarrow> tab\<^sub>1 (trg \<mu>)\<guillemotright>"
using \<mu> by simp
moreover have "iso \<l>[tab\<^sub>1 (trg \<mu>)]"
using \<mu> iso_lunit by simp
ultimately have "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)"
using isomorphic_def by auto
also have "tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>0 (trg \<mu>)"
using \<mu> trg.obj_has_symmetric_tab isomorphic_symmetric by auto
finally show ?thesis by blast
qed
hence "trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)"
using \<mu> trg.obj_has_symmetric_tab isomorphic_transitive by blast
have "\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)"
proof -
have "\<guillemotleft>tab\<^sub>1 (trg \<mu>) : src (tab\<^sub>0 (trg \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and> is_left_adjoint (tab\<^sub>0 (trg \<mu>)) \<and>
\<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis by auto
qed
have "\<lbrakk>trg \<mu>\<rbrakk> \<in> Maps.Hom (trg \<mu>) (trg \<mu>)"
proof -
have "\<guillemotleft>trg \<mu> : trg \<mu> \<rightarrow> trg \<mu>\<guillemotright> \<and> is_left_adjoint (trg \<mu>) \<and> \<lbrakk>trg \<mu>\<rbrakk> = \<lbrakk>trg \<mu>\<rbrakk>"
using \<mu> obj_is_self_adjoint by simp
thus ?thesis by auto
qed
interpret SPN_trg: arrow_of_spans Maps.comp \<open>SPN (trg \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have SPN_trg: "SPN (trg \<mu>) =
\<lparr>Chn = Maps.MkArr (src (tab\<^sub>1 (trg \<mu>))) (src (tab\<^sub>1 (trg \<mu>))) \<lbrakk>spn (trg \<mu>)\<rbrakk>,
Dom = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>,
Cod = \<lparr>Leg0 = ?tab\<^sub>0, Leg1 = ?tab\<^sub>1\<rparr>\<rparr>"
unfolding SPN_def using \<mu> by simp
interpret trg_SPN: arrow_of_spans Maps.comp \<open>Span.trg (SPN \<mu>)\<close>
using \<mu> SPN.preserves_reflects_arr Span.arr_char by blast
have trg_SPN: "Span.trg (SPN \<mu>) = \<lparr>Chn = ?trg,
Dom = \<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>,
Cod = \<lparr>Leg0 = ?trg, Leg1 = ?trg\<rparr>\<rparr>"
proof -
let ?tab\<^sub>1_dom = "Maps.MkArr (src (tab\<^sub>1 (dom \<mu>))) (trg \<mu>) \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>"
have "Maps.arr ?tab\<^sub>1_dom"
proof -
have "\<guillemotleft>tab\<^sub>1 (dom \<mu>) : src (tab\<^sub>1 (dom \<mu>)) \<rightarrow> trg \<mu>\<guillemotright> \<and>
is_left_adjoint (tab\<^sub>1 (dom \<mu>)) \<and> \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (dom \<mu>)\<rbrakk>"
using \<mu> by simp
thus ?thesis
using \<mu> Maps.arr_MkArr by blast
qed
thus ?thesis
using \<mu> Maps.cod_char Span.trg_def by simp
qed
let ?\<phi> = "\<lparr>Chn = ?tab\<^sub>1, Dom = Dom (SPN (trg \<mu>)), Cod = Cod (Span.trg (SPN \<mu>))\<rparr>"
interpret \<phi>: arrow_of_spans Maps.comp ?\<phi>
apply (unfold_locales, simp_all)
proof -
show "Maps.in_hom ?tab\<^sub>1 SPN_trg.dom.apex trg_SPN.cod.apex"
using \<mu> tab\<^sub>0 tab\<^sub>1 Maps.dom_char Maps.cod_char SPN_trg trg_SPN
tab.apex_def trg_SPN.cod.apex_def
apply (intro Maps.in_homI) by simp_all
(* The preceding cannot be written "by (intro Maps.in_homI, simp_all)", why? *)
show "Maps.comp trg_SPN.cod.leg0 ?tab\<^sub>1 = SPN_trg.dom.leg0"
proof -
have "Maps.seq trg_SPN.cod.leg0 ?tab\<^sub>1"
using \<mu> trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (trg \<mu>) \<in> Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "trg \<mu>"]
ide_in_iso_class [of "tab\<^sub>1 (trg \<mu>)"] \<open>trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>iso_class (tab\<^sub>1 (trg \<mu>)) \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)\<close>
\<open>iso_class (trg \<mu>) \<in> Maps.Hom (trg \<mu>) (trg \<mu>)\<close>
by meson
qed
moreover have "\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk>"
using \<mu> iso_class_eqI trg.obj_has_symmetric_tab by auto
ultimately show ?thesis
using \<mu> Maps.comp_char [of trg_SPN.cod.leg0 ?tab\<^sub>1] trg_SPN
by simp
qed
show "trg_SPN.cod.leg1 \<odot> ?tab\<^sub>1 = SPN_trg.dom.leg1"
proof -
have "Maps.seq trg_SPN.cod.leg1 ?tab\<^sub>1"
using \<mu> trg_SPN tab\<^sub>1 Maps.dom_char Maps.cod_char
by (intro Maps.seqI, auto)
moreover have "Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
proof -
have "tab\<^sub>1 (trg \<mu>) \<in> Maps.Comp \<lbrakk>trg \<mu>\<rbrakk> \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> is_iso_classI ide_in_iso_class [of "trg \<mu>"]
ide_in_iso_class [of "tab\<^sub>1 (trg \<mu>)"] \<open>trg \<mu> \<star> tab\<^sub>1 (trg \<mu>) \<cong> tab\<^sub>1 (trg \<mu>)\<close>
by auto
thus ?thesis
using Maps.Comp_eq_iso_class_memb
\<open>\<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk> \<in> Maps.Hom (src (tab\<^sub>0 (trg \<mu>))) (trg \<mu>)\<close>
\<open>\<lbrakk>trg \<mu>\<rbrakk> \<in> Maps.Hom (trg \<mu>) (trg \<mu>)\<close>
by meson
qed
ultimately show ?thesis
using \<mu> Maps.comp_char [of trg_SPN.cod.leg1 ?tab\<^sub>1] trg_SPN by simp
qed
qed
have \<phi>: "Span.in_hom ?\<phi> (SPN (trg \<mu>)) (Span.trg (SPN \<mu>))"
using \<mu> tab\<^sub>0 spn_ide [of "trg \<mu>"] iso_class_eqI
Span.arr_char Span.dom_char Span.cod_char \<phi>.arrow_of_spans_axioms
SPN_trg trg_SPN trg.apex_def tab.apex_def Maps.dom_char
apply (intro Span.in_homI) by auto
have "Maps.iso ?tab\<^sub>1"
proof -
have "Maps.iso ?tab\<^sub>0"
using \<mu> tab\<^sub>0 ide_in_iso_class trg.is_map_iff_tab\<^sub>0_is_equivalence obj_is_self_adjoint
Maps.iso_char' [of ?tab\<^sub>0]
by auto
moreover have "?tab\<^sub>0 = ?tab\<^sub>1"
proof -
have "\<lbrakk>tab\<^sub>0 (trg \<mu>)\<rbrakk> = \<lbrakk>tab\<^sub>1 (trg \<mu>)\<rbrakk>"
using \<mu> iso_class_eqI trg.obj_has_symmetric_tab by auto
thus ?thesis by simp
qed
ultimately show ?thesis by simp
qed
thus "Span.isomorphic (SPN (trg \<mu>)) (Span.trg (SPN \<mu>))"
using \<phi> Span.isomorphic_def Span.iso_char by auto
qed
qed
lemma SPN_is_weak_arrow_of_homs:
shows "weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN"
..
end
subsubsection "Compositors"
text \<open>
To complete the proof that \<open>SPN\<close> is a pseudofunctor, we need to obtain a natural
isomorphism \<open>\<Phi>\<close>, whose component at \<open>(r, s)\<close> is an isomorphism \<open>\<Phi> (r, s)\<close>
from the horizontal composite \<open>SPN r \<circ> SPN s\<close> to \<open>SPN (r \<star> s)\<close> in \<open>Span(Maps(B))\<close>,
and we need to prove that the coherence conditions are satisfied.
We have shown that the tabulations of \<open>r\<close> and \<open>s\<close> compose to yield a tabulation of \<open>r \<star> s\<close>.
Since tabulations of the same arrow are equivalent, this tabulation must be equivalent
to the chosen tabulation of \<open>r \<star> s\<close>. We therefore obtain an equivalence map from the
apex of \<open>SPN r \<circ> SPN s\<close> to the apex of \<open>SPN (r \<star> s)\<close> which commutes with the
legs of these spans up to isomorphism. This equivalence map determines an invertible
arrow in \<open>Span(Maps(B))\<close>. Moreover, by property \<open>T2\<close>, any two such equivalence maps are
connected by a unique 2-cell, which is consequently an isomorphism. This shows that
the arrow in \<open>Span(Maps(B))\<close> is uniquely determined, which fact we can exploit to establish
the required coherence conditions.
\<close>
locale two_composable_identities_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r +
s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and r :: 'a
and s :: 'a +
assumes composable: "src r = trg s"
begin
notation isomorphic (infix "\<cong>" 50)
interpretation r: arrow_in_bicategory_of_spans V H \<a> \<i> src trg r r r
by unfold_locales auto
interpretation r: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close>
r
using r.is_arrow_of_tabulations_in_maps by simp
interpretation s: arrow_in_bicategory_of_spans V H \<a> \<i> src trg s s s
by unfold_locales auto
interpretation s: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
s
using s.is_arrow_of_tabulations_in_maps by simp
sublocale identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>r \<star> s\<close>
apply unfold_locales by (simp add: composable)
sublocale horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close>
r s
using composable by unfold_locales auto
abbreviation p\<^sub>0 where "p\<^sub>0 \<equiv> \<rho>\<sigma>.p\<^sub>0"
abbreviation p\<^sub>1 where "p\<^sub>1 \<equiv> \<rho>\<sigma>.p\<^sub>1"
text \<open>
We will take as the composition isomorphism from \<open>SPN r \<circ> SPN s\<close> to \<open>SPN (r \<star> s)\<close>
the arrow of tabulations, induced by the identity \<open>r \<star> s\<close>, from the composite of
the chosen tabulations of \<open>r\<close> and \<open>s\<close> to the chosen tabulation of \<open>r \<star> s\<close>.
As this arrow of tabulations is induced by an identity, it is an equivalence map.
\<close>
interpretation cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> \<rho>\<sigma>.tab \<open>tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 r \<star> \<rho>\<sigma>.p\<^sub>1\<close>
\<open>r \<star> s\<close> tab \<open>tab\<^sub>0 (r \<star> s)\<close> \<open>tab\<^sub>1 (r \<star> s)\<close>
\<open>r \<star> s\<close>
using composable
by unfold_locales auto
lemma cmp_interpretation:
shows "identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(r \<star> s) \<rho>\<sigma>.tab (tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0) (tab\<^sub>1 r \<star> \<rho>\<sigma>.p\<^sub>1)
(r \<star> s) tab (tab\<^sub>0 (r \<star> s)) (tab\<^sub>1 (r \<star> s))
(r \<star> s)"
..
definition cmp
where "cmp = cmp.chine"
lemma cmp_props:
shows "\<guillemotleft>cmp : src \<rho>\<sigma>.tab \<rightarrow> src tab\<guillemotright>"
and "\<guillemotleft>cmp : cmp \<Rightarrow> cmp\<guillemotright>"
and "equivalence_map cmp"
and "tab\<^sub>0 (r \<star> s) \<star> cmp \<cong> tab\<^sub>0 s \<star> \<rho>\<sigma>.p\<^sub>0"
and "tab\<^sub>1 (r \<star> s) \<star> cmp \<cong> tab\<^sub>1 r \<star> \<rho>\<sigma>.p\<^sub>1"
using cmp_def cmp.leg0_uniquely_isomorphic(1) cmp.leg1_uniquely_isomorphic(1)
isomorphic_symmetric cmp.chine_is_equivalence
by auto
lemma cmp_in_hom [intro]:
shows "\<guillemotleft>cmp : src \<rho>\<sigma>.tab \<rightarrow> src tab\<guillemotright>"
and "\<guillemotleft>cmp : cmp \<Rightarrow> cmp\<guillemotright>"
using cmp_props by auto
lemma cmp_simps [simp]:
shows "arr cmp" and "ide cmp"
and "src cmp = src \<rho>\<sigma>.tab" and "trg cmp = src tab"
and "dom cmp = cmp" and "cod cmp = cmp"
using cmp_props equivalence_map_is_ide by auto
text \<open>
Now we have to use the above properties of the underlying bicategory to
exhibit the composition isomorphisms as actual arrows in \<open>Span(Maps(B))\<close>
and to prove the required naturality and coherence conditions.
\<close>
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
interpretation SPN: "functor" V Span.vcomp SPN
using SPN_is_functor by simp
interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
using SPN_is_weak_arrow_of_homs by simp
interpretation SPN_r_SPN_s: arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using composable Span.ide_char [of "SPN r \<circ> SPN s"] by simp
interpretation SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using composable Span.ide_char [of "SPN r \<circ> SPN s"]
by (unfold_locales, simp)
interpretation SPN_rs: arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using composable Span.arr_char r.base_simps(2) s.base_simps(2) by blast
interpretation SPN_rs: identity_arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using composable Span.ide_char SPN.preserves_ide r.is_ide s.is_ide
by (unfold_locales, simp)
text \<open>
The following are the legs (as arrows of \<open>Maps\<close>) of the spans \<open>SPN r\<close> and \<open>SPN s\<close>.
\<close>
definition R\<^sub>0 where "R\<^sub>0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>"
definition R\<^sub>1 where "R\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
definition S\<^sub>0 where "S\<^sub>0 = \<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>"
definition S\<^sub>1 where "S\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 s\<rbrakk>\<rbrakk>"
lemma span_legs_eq:
shows "Leg0 (Dom (SPN r)) = R\<^sub>0" and "Leg1 (Dom (SPN r)) = R\<^sub>1"
and "Leg0 (Dom (SPN s)) = S\<^sub>0" and "Leg1 (Dom (SPN s)) = S\<^sub>1"
using SPN_def R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_def composable by auto
lemma R\<^sub>0_in_hom [intro]:
shows "Maps.in_hom R\<^sub>0 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (src r))"
by (simp add: Maps.MkArr_in_hom' R\<^sub>0_def)
lemma R\<^sub>1_in_hom [intro]:
shows "Maps.in_hom R\<^sub>1 (Maps.MkIde (src r.s\<^sub>0)) (Maps.MkIde (trg r))"
by (simp add: Maps.MkArr_in_hom' R\<^sub>1_def)
lemma S\<^sub>0_in_hom [intro]:
shows "Maps.in_hom S\<^sub>0 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (src s))"
by (simp add: Maps.MkArr_in_hom' S\<^sub>0_def)
lemma S\<^sub>1_in_hom [intro]:
shows "Maps.in_hom S\<^sub>1 (Maps.MkIde (src s.s\<^sub>0)) (Maps.MkIde (trg s))"
by (simp add: Maps.MkArr_in_hom' S\<^sub>1_def)
lemma RS_simps [simp]:
shows "Maps.arr R\<^sub>0" and "Maps.dom R\<^sub>0 = Maps.MkIde (src r.s\<^sub>0)"
and "Maps.cod R\<^sub>0 = Maps.MkIde (src r)"
and "Maps.Dom R\<^sub>0 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>0 = src r"
and "Maps.arr R\<^sub>1" and "Maps.dom R\<^sub>1 = Maps.MkIde (src r.s\<^sub>0)"
and "Maps.cod R\<^sub>1 = Maps.MkIde (trg r)"
and "Maps.Dom R\<^sub>1 = src r.s\<^sub>0" and "Maps.Cod R\<^sub>1 = trg r"
and "Maps.arr S\<^sub>0" and "Maps.dom S\<^sub>0 = Maps.MkIde (src s.s\<^sub>0)"
and "Maps.cod S\<^sub>0 = Maps.MkIde (src s)"
and "Maps.Dom S\<^sub>0 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>0 = src s"
and "Maps.arr S\<^sub>1" and "Maps.dom S\<^sub>1 = Maps.MkIde (src s.s\<^sub>0)"
and "Maps.cod S\<^sub>1 = Maps.MkIde (trg s)"
and "Maps.Dom S\<^sub>1 = src s.s\<^sub>0" and "Maps.Cod S\<^sub>1 = trg s"
using R\<^sub>0_in_hom R\<^sub>1_in_hom S\<^sub>0_in_hom S\<^sub>1_in_hom composable
by (auto simp add: R\<^sub>0_def R\<^sub>1_def S\<^sub>0_def S\<^sub>1_def)
text \<open>
The apex of the composite span @{term "SPN r \<circ> SPN s"} (defined in terms of pullback)
coincides with the apex of the composite tabulation \<open>\<rho>\<sigma>\<close> (defined using
the chosen tabulation of \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s)\<close>). We need this to be true in order
to define the compositor of a pseudofunctor from the underlying bicategory \<open>B\<close>
to \<open>Span(Maps(B))\<close>. It is only true if we have carefully chosen pullbacks in \<open>Maps(B)\<close>
in order to ensure the relationship with the chosen tabulations.
\<close>
lemma SPN_r_SPN_s_apex_eq:
shows "SPN_r_SPN_s.apex = Maps.MkIde (src \<rho>\<sigma>.tab)"
proof -
have "obj (Maps.Cod SPN_r_SPN_s.leg0)"
using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
moreover have "obj (Maps.Dom SPN_r_SPN_s.leg0)"
using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
moreover have "SPN_r_SPN_s.leg0 \<noteq> Maps.Null"
using Maps.arr_char [of "SPN_r_SPN_s.leg0"] by simp
moreover have "Maps.Dom SPN_r_SPN_s.leg0 = src \<rho>\<sigma>.tab"
proof -
interpret REP_S\<^sub>1: map_in_bicategory V H \<a> \<i> src trg \<open>Maps.REP S\<^sub>1\<close>
using Maps.REP_in_Map Maps.arr_char Maps.in_HomD S\<^sub>1_def
apply unfold_locales
by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom)
interpret REP_R\<^sub>0: map_in_bicategory V H \<a> \<i> src trg \<open>Maps.REP R\<^sub>0\<close>
using Maps.REP_in_Map Maps.arr_char Maps.in_HomD R\<^sub>0_def
apply unfold_locales
by (meson Maps.REP_in_hhom(2) R\<^sub>0_in_hom)
have "Maps.Dom SPN_r_SPN_s.leg0 = Maps.Dom (S\<^sub>0 \<odot> Maps.PRJ\<^sub>0 R\<^sub>0 S\<^sub>1)"
using composable Span.hcomp_def S\<^sub>0_def R\<^sub>0_def S\<^sub>1_def by simp
also have "... = Maps.Dom \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
proof -
have "ide ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)"
proof -
have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)"
using REP_R\<^sub>0.is_map REP_S\<^sub>1.is_map left_adjoint_is_ide R\<^sub>0_def S\<^sub>1_def
by (metis (no_types, lifting) Maps.REP_CLS REP_R\<^sub>0.antipar(2)
isomorphic_implies_hpar(4) composable r.leg0_simps(3)
r.satisfies_T0 s.leg1_is_map s.leg1_simps(3) s.leg1_simps(4))
thus ?thesis by simp
qed
thus ?thesis by simp
qed
moreover have "Maps.Dom (S\<^sub>0 \<odot> \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>) =
src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
proof -
have "Maps.arr (\<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk>)"
using Maps.CLS_in_hom Maps.prj0_simps(1) Maps.PRJ\<^sub>0_def composable by fastforce
moreover have "Maps.Dom S\<^sub>0 = Maps.Cod \<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk>"
proof -
have "Maps.Cod \<lbrakk>\<lbrakk>prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)\<rbrakk>\<rbrakk> =
trg (prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0))"
by simp
also have "... = src (Maps.REP S\<^sub>1)"
proof -
have "ide (Maps.REP S\<^sub>1)"
by simp
moreover have "is_left_adjoint (Maps.REP R\<^sub>0)"
by auto
moreover have "trg (Maps.REP S\<^sub>1) = trg (Maps.REP R\<^sub>0)"
by (simp add: composable)
ultimately show ?thesis
using S\<^sub>1_def Maps.REP_CLS r.leg0_is_map s.leg1_is_map by simp
qed
also have "... = src (tab\<^sub>0 s)"
using tab\<^sub>0_in_hom(1) by simp
also have "... = Maps.Dom S\<^sub>0"
using S\<^sub>0_def by simp
finally show ?thesis by simp
qed
ultimately have
"Maps.Dom (S\<^sub>0 \<odot> \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>) =
Maps.Dom \<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)\<rbrakk>\<rbrakk>"
using Maps.CLS_in_hom by simp
thus ?thesis by simp
qed
ultimately show ?thesis
using Maps.PRJ\<^sub>0_def composable Maps.Dom.simps(1) RS_simps(1) RS_simps(16)
RS_simps(18) RS_simps(3)
by presburger
qed
also have "... = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
by simp
finally have
"Maps.Dom SPN_r_SPN_s.leg0 = src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1))"
by simp
also have "... = src (tab\<^sub>0 (r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1))"
proof -
interpret r\<^sub>0's\<^sub>1: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1\<close>
using composable by (unfold_locales, simp)
have "(Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1 \<cong> r.s\<^sub>0\<^sup>* \<star> s.s\<^sub>1"
proof -
have "(Maps.REP R\<^sub>0)\<^sup>* \<cong> r.s\<^sub>0\<^sup>*"
proof -
have 1: "adjoint_pair (Maps.REP R\<^sub>0) (Maps.REP R\<^sub>0)\<^sup>*"
using REP_R\<^sub>0.is_map left_adjoint_extends_to_adjoint_pair by blast
moreover have "adjoint_pair r.s\<^sub>0 (Maps.REP R\<^sub>0)\<^sup>*"
proof -
have "Maps.REP R\<^sub>0 \<cong> r.s\<^sub>0"
unfolding R\<^sub>0_def
using Maps.REP_CLS r.leg0_is_map composable by force
thus ?thesis
using 1 adjoint_pair_preserved_by_iso isomorphic_def
REP_R\<^sub>0.triangle_in_hom(4) REP_R\<^sub>0.triangle_right'
by auto
qed
ultimately show ?thesis
using r.leg0_is_map left_adjoint_determines_right_up_to_iso
left_adjoint_extends_to_adjoint_pair
by auto
qed
moreover have "Maps.REP S\<^sub>1 \<cong> s.s\<^sub>1"
unfolding S\<^sub>1_def
using Maps.REP_CLS s.leg1_is_map composable by force
moreover have "\<exists>a. a \<cong> (tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<and> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1 \<cong> a"
using calculation composable isomorphic_implies_hpar(3)
hcomp_ide_isomorphic hcomp_isomorphic_ide [of "(Maps.REP R\<^sub>0)\<^sup>*" "r.s\<^sub>0\<^sup>*" s.s\<^sub>1]
by auto
ultimately show ?thesis
using isomorphic_transitive by blast
qed
thus ?thesis
using r\<^sub>0's\<^sub>1.isomorphic_implies_same_tab isomorphic_symmetric by metis
qed
also have "... = src \<rho>\<sigma>.tab"
using VV.ide_char VV.arr_char composable Span.hcomp_def \<rho>\<sigma>.tab_simps(2) by auto
finally show ?thesis by simp
qed
ultimately show ?thesis
using composable Maps.arr_char Maps.dom_char SPN_r_SPN_s.dom.apex_def
apply auto
by (metis (no_types, lifting) Maps.not_arr_null SPN_r_SPN_s.chine_eq_apex
SPN_r_SPN_s.chine_simps(1))
qed
text \<open>
We will be taking the arrow @{term "CLS cmp"} of \<open>Maps\<close> as the composition isomorphism from
@{term "SPN r \<circ> SPN s"} to @{term "SPN (r \<star> s)"}. The following result shows that it
has the right domain and codomain for that purpose.
\<close>
lemma iso_class_cmp_in_hom:
shows "Maps.in_hom (Maps.MkArr (src \<rho>\<sigma>.tab) (src tab) \<lbrakk>cmp\<rbrakk>)
SPN_r_SPN_s.apex SPN_rs.apex"
and "Maps.in_hom \<lbrakk>\<lbrakk>cmp\<rbrakk>\<rbrakk> SPN_r_SPN_s.apex SPN_rs.apex"
proof -
show "Maps.in_hom (Maps.MkArr (src \<rho>\<sigma>.tab) (src tab) \<lbrakk>cmp\<rbrakk>)
SPN_r_SPN_s.apex SPN_rs.apex"
proof -
have "obj (src \<rho>\<sigma>.tab)"
using obj_src \<rho>\<sigma>.tab_in_hom by blast
moreover have "obj (src tab)"
using obj_src by simp
moreover have "\<lbrakk>cmp\<rbrakk> \<in> Maps.Hom (src \<rho>\<sigma>.tab) (src tab)"
by (metis (mono_tags, lifting) cmp.is_map cmp_def cmp_props(1) mem_Collect_eq)
moreover have "SPN_r_SPN_s.apex = Maps.MkIde (src \<rho>\<sigma>.tab)"
using SPN_r_SPN_s_apex_eq by simp
moreover have "SPN_rs.apex = Maps.MkIde (src tab)"
using SPN_def composable SPN_rs.cod.apex_def Maps.arr_char Maps.dom_char
SPN_rs.dom.leg_simps(1)
by fastforce
ultimately show ?thesis
using Maps.MkArr_in_hom by simp
qed
thus "Maps.in_hom \<lbrakk>\<lbrakk>cmp\<rbrakk>\<rbrakk> SPN_r_SPN_s.apex SPN_rs.apex" by simp
qed
interpretation r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans
V H \<a> \<i> src trg \<open>(Maps.REP R\<^sub>0)\<^sup>*\<close> \<open>Maps.REP S\<^sub>1\<close>
proof
show "ide (Maps.REP S\<^sub>1)"
using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
by (meson Maps.REP_in_hhom(2) S\<^sub>1_in_hom)
show "ide (Maps.REP R\<^sub>0)\<^sup>*"
using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
Maps.REP_in_hhom(2) R\<^sub>0_in_hom by auto
show "src (Maps.REP R\<^sub>0)\<^sup>* = trg (Maps.REP S\<^sub>1)"
using Maps.REP_in_hhom(2) R\<^sub>0_in_hom composable by auto
qed
interpretation R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s\<close>
by (unfold_locales, simp add: composable)
lemma prj_tab_agreement:
shows "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<cong> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1"
and "\<rho>\<sigma>.p\<^sub>0 \<cong> prj\<^sub>0 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)"
and "\<rho>\<sigma>.p\<^sub>1 \<cong> prj\<^sub>1 (Maps.REP S\<^sub>1) (Maps.REP R\<^sub>0)"
proof -
show 1: "(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s \<cong> (Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1"
proof -
have "(tab\<^sub>0 r)\<^sup>* \<cong> (Maps.REP R\<^sub>0)\<^sup>*"
using Maps.REP_CLS isomorphic_symmetric R\<^sub>0_def composable r.satisfies_T0
isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint
by fastforce
moreover have "tab\<^sub>1 s \<cong> Maps.REP S\<^sub>1"
by (metis Maps.REP_CLS isomorphic_symmetric S\<^sub>1_def s.leg1_is_map s.leg1_simps(3-4))
moreover have "src (Maps.REP R\<^sub>0)\<^sup>* = trg (tab\<^sub>1 s)"
using composable r.T0.antipar right_adjoint_simps(2) by fastforce
ultimately show ?thesis
using hcomp_isomorphic_ide [of "(tab\<^sub>0 r)\<^sup>*" "(Maps.REP R\<^sub>0)\<^sup>*" "tab\<^sub>1 s"]
hcomp_ide_isomorphic isomorphic_transitive composable
by auto
qed
show "\<rho>\<sigma>.p\<^sub>0 \<cong> tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)"
using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto
show "\<rho>\<sigma>.p\<^sub>1 \<cong> tab\<^sub>1 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)"
using 1 R\<^sub>0'S\<^sub>1.isomorphic_implies_same_tab isomorphic_reflexive by auto
qed
lemma chine_hcomp_SPN_SPN:
shows "Span.chine_hcomp (SPN r) (SPN s) = Maps.MkIde (src \<rho>\<sigma>.p\<^sub>0)"
proof -
have "Span.chine_hcomp (SPN r) (SPN s) =
Maps.MkIde (src (tab\<^sub>0 ((Maps.REP R\<^sub>0)\<^sup>* \<star> Maps.REP S\<^sub>1)))"
using Span.chine_hcomp_ide_ide [of "SPN r" "SPN s"] composable
Maps.pbdom_def Maps.PRJ\<^sub>0_def Maps.CLS_in_hom Maps.dom_char R\<^sub>0_def S\<^sub>1_def
apply simp
using Maps.prj0_simps(1) RS_simps(1) RS_simps(16) RS_simps(18) RS_simps(3)
by presburger
also have "... = Maps.MkIde (src \<rho>\<sigma>.p\<^sub>0)"
using prj_tab_agreement isomorphic_implies_hpar(3) by force
finally show ?thesis by simp
qed
end
text \<open>
The development above focused on two specific composable 1-cells in bicategory \<open>B\<close>.
Here we reformulate those results as statements about the entire bicategory.
\<close>
context bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 50)
interpretation SPN: "functor" V Span.vcomp SPN
using SPN_is_functor by simp
interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
using SPN_is_weak_arrow_of_homs by simp
interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
SPN.FF \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<circ> snd \<mu>\<nu>\<close>
..
interpretation SPNoH: composite_functor VV.comp V
Span.vcomp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
text \<open>
Given arbitrary composable 1-cells \<open>r\<close> and \<open>s\<close>, obtain an arrow of spans in \<open>Maps\<close>
having the isomorphism class of \<open>rs.cmp\<close> as its chine.
\<close>
definition CMP
where "CMP r s \<equiv>
\<lparr>Chn = \<lbrakk>\<lbrakk>two_composable_identities_in_bicategory_of_spans.cmp V H \<a> \<i> src trg r s\<rbrakk>\<rbrakk>,
Dom = Dom (SPN r \<circ> SPN s), Cod = Dom (SPN (r \<star> s))\<rparr>"
lemma compositor_in_hom [intro]:
assumes "ide r" and "ide s" and "src r = trg s"
shows "Span.in_hhom (CMP r s) (SPN.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (trg r))"
and "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))"
proof -
have rs: "VV.ide (r, s)"
using assms VV.ide_char VV.arr_char by simp
interpret rs: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg r s
using rs VV.ide_char VV.arr_char by unfold_locales auto
interpret cmp: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>r \<star> s\<close> rs.\<rho>\<sigma>.tab \<open>tab\<^sub>0 s \<star> rs.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 r \<star> rs.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>r \<star> s\<close> rs.tab \<open>tab\<^sub>0 (r \<star> s)\<close> \<open>tab\<^sub>1 (r \<star> s)\<close>
\<open>r \<star> s\<close>
by unfold_locales auto
have "rs.cmp = cmp.chine"
using rs.cmp_def by simp
interpret SPN_r_SPN_s: arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using rs.composable Span.ide_char [of "SPN r \<circ> SPN s"] by simp
interpret SPN_r_SPN_s: identity_arrow_of_spans Maps.comp \<open>SPN r \<circ> SPN s\<close>
using rs.composable Span.ide_char [of "SPN r \<circ> SPN s"]
by (unfold_locales, simp)
interpret SPN_rs: arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using Span.arr_char rs.is_ide SPN.preserves_arr by blast
interpret SPN_rs: identity_arrow_of_spans Maps.comp \<open>SPN (r \<star> s)\<close>
using Span.ide_char rs.is_ide SPN.preserves_ide
by (unfold_locales, simp)
interpret Dom: span_in_category Maps.comp \<open>Dom (CMP r s)\<close>
by (unfold_locales, simp add: CMP_def)
interpret Cod: span_in_category Maps.comp \<open>Cod (CMP r s)\<close>
proof -
(* TODO: I don't understand what makes this so difficult. *)
have "\<guillemotleft>tab\<^sub>0 (r \<star> s) : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> src s\<guillemotright> \<and> is_left_adjoint (tab\<^sub>0 (r \<star> s)) \<and>
\<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> = \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk>"
by simp
hence "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> src s\<guillemotright> \<and> is_left_adjoint f \<and> \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> = \<lbrakk>f\<rbrakk>"
by blast
moreover have "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (r \<star> s)) \<rightarrow> trg r\<guillemotright> \<and> is_left_adjoint f \<and>
\<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> = \<lbrakk>f\<rbrakk>"
by (metis rs.base_simps(2) rs.leg1_in_hom(1) rs.leg1_is_map trg_hcomp)
ultimately show "span_in_category Maps.comp (Cod (CMP r s))"
using assms Maps.arr_char Maps.dom_char CMP_def
by unfold_locales auto
qed
interpret r\<^sub>0's\<^sub>1: two_composable_identities_in_bicategory_of_spans
V H \<a> \<i> src trg \<open>(Maps.REP rs.R\<^sub>0)\<^sup>*\<close> \<open>Maps.REP rs.S\<^sub>1\<close>
proof
show "ide (Maps.REP rs.S\<^sub>1)"
using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
by (meson Maps.REP_in_hhom(2) rs.S\<^sub>1_in_hom)
show "ide (Maps.REP rs.R\<^sub>0)\<^sup>*"
using Maps.REP_in_Map Maps.arr_char left_adjoint_is_ide
Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom by auto
show "src (Maps.REP rs.R\<^sub>0)\<^sup>* = trg (Maps.REP rs.S\<^sub>1)"
using Maps.REP_in_hhom(2) rs.R\<^sub>0_in_hom rs.composable by auto
qed
interpret R\<^sub>0'S\<^sub>1: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>(tab\<^sub>0 r)\<^sup>* \<star> tab\<^sub>1 s\<close>
by (unfold_locales, simp add: rs.composable)
text \<open>
Here we obtain explicit formulas for the legs and apex of \<open>SPN_r_SPN_s\<close> and \<open>SPN_rs\<close>.
\<close>
have SPN_r_SPN_s_leg0_eq:
"SPN_r_SPN_s.leg0 = Maps.comp rs.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)"
using rs.composable Span.hcomp_def rs.S\<^sub>0_def rs.R\<^sub>0_def rs.S\<^sub>1_def by simp
have SPN_r_SPN_s_leg1_eq:
"SPN_r_SPN_s.leg1 = Maps.comp rs.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)"
using rs.composable Span.hcomp_def rs.R\<^sub>1_def rs.R\<^sub>0_def rs.S\<^sub>1_def by simp
have "SPN_r_SPN_s.apex = Maps.MkIde (src rs.\<rho>\<sigma>.tab)"
using rs.SPN_r_SPN_s_apex_eq by auto
have SPN_rs_leg0_eq: "SPN_rs.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk>\<rbrakk>"
unfolding SPN_def using rs by simp
have SPN_rs_leg1_eq: "SPN_rs.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk>\<rbrakk>"
unfolding SPN_def using rs by simp
have "SPN_rs.apex = Maps.MkIde (src (tab_of_ide (r \<star> s)))"
using SPN_rs.dom.apex_def Maps.dom_char SPN_rs_leg0_eq SPN_rs.dom.leg_simps(1)
by simp
text \<open>
The composition isomorphism @{term "CMP r s"} is an arrow of spans in \<open>Maps(B)\<close>.
\<close>
interpret arrow_of_spans Maps.comp \<open>CMP r s\<close>
proof
show 1: "Maps.in_hom (Chn (CMP r s)) Dom.apex Cod.apex"
using rs.iso_class_cmp_in_hom rs.composable CMP_def by simp
show "Cod.leg0 \<odot> Chn (CMP r s) = Dom.leg0"
proof (intro Maps.arr_eqI)
show 2: "Maps.seq Cod.leg0 (Chn (CMP r s))"
using 1 Maps.dom_char Maps.cod_char by blast
show 3: "Maps.arr Dom.leg0" by simp
show "Maps.Dom (Cod.leg0 \<odot> Chn (CMP r s)) = Maps.Dom Dom.leg0"
using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char
Dom.leg_in_hom Maps.in_hom_char Maps.seq_char
by auto
show "Maps.Cod (Cod.leg0 \<odot> Chn (CMP r s)) = Maps.Cod Dom.leg0"
using 2 3 Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"]
Dom.leg_simps Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg0_eq
Maps.comp_char [of rs.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"] CMP_def
by simp
show "Maps.Map (Cod.leg0 \<odot> Chn (CMP r s)) = Maps.Map Dom.leg0"
proof -
have "Maps.Map (Cod.leg0 \<odot> Chn (CMP r s)) = Maps.Comp \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk>"
using 1 2 Maps.dom_char Maps.cod_char
Maps.comp_char [of Cod.leg0 "Chn (CMP r s)"] CMP_def
by simp
also have "... = Maps.Comp \<lbrakk>tab\<^sub>0 s\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
proof -
have "Maps.Comp \<lbrakk>tab\<^sub>0 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk> = \<lbrakk>tab\<^sub>0 (r \<star> s) \<star> rs.cmp\<rbrakk>"
using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
by auto
also have "... = Maps.Comp \<lbrakk>tab\<^sub>0 s\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp iso_class_eqI rs.cmp_props(4)
by auto
finally show ?thesis by simp
qed
also have "... = Maps.Map Dom.leg0"
proof -
have "Maps.seq rs.S\<^sub>0 (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1)"
by (intro Maps.seqI, simp_all add: rs.composable)
moreover have "\<lbrakk>prj\<^sub>0 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\<rbrakk> = \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>0\<rbrakk>"
using "rs.prj_tab_agreement" iso_class_eqI by auto
moreover have "Maps.Dom (Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\<rho>\<sigma>.p\<^sub>0"
using rs.prj_tab_agreement Maps.PRJ\<^sub>0_def rs.composable
isomorphic_implies_hpar(3)
by auto
ultimately show ?thesis
using SPN_r_SPN_s_leg0_eq Maps.comp_char [of rs.S\<^sub>0 "Maps.PRJ\<^sub>0 rs.R\<^sub>0 rs.S\<^sub>1"]
rs.S\<^sub>0_def Maps.PRJ\<^sub>0_def rs.composable CMP_def
by simp
qed
finally show ?thesis by simp
qed
qed
show "Cod.leg1 \<odot> Chn (CMP r s) = Dom.leg1"
proof (intro Maps.arr_eqI)
show 2: "Maps.seq Cod.leg1 (Chn (CMP r s))"
using 1 Maps.dom_char Maps.cod_char by blast
show 3: "Maps.arr Dom.leg1" by simp
show "Maps.Dom (Cod.leg1 \<odot> Chn (CMP r s)) = Maps.Dom Dom.leg1"
using 1 2 Maps.dom_char Maps.cod_char Maps.comp_char
Dom.leg_in_hom Maps.in_hom_char Maps.seq_char
by auto
show "Maps.Cod (Cod.leg1 \<odot> Chn (CMP r s)) = Maps.Cod Dom.leg1"
using 2 3 Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"]
Dom.apex_def Maps.dom_char SPN_r_SPN_s_leg1_eq
Maps.comp_char [of rs.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"] CMP_def
by simp
show "Maps.Map (Cod.leg1 \<odot> Chn (CMP r s)) = Maps.Map Dom.leg1"
proof -
have "Maps.Map (Cod.leg1 \<odot> Chn (CMP r s)) = Maps.Comp \<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk>"
using 1 2 Maps.dom_char Maps.cod_char
Maps.comp_char [of Cod.leg1 "Chn (CMP r s)"] CMP_def
by simp
also have "... = Maps.Comp \<lbrakk>tab\<^sub>1 r\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
proof -
have "Maps.Comp \<lbrakk>tab\<^sub>1 (r \<star> s)\<rbrakk> \<lbrakk>rs.cmp\<rbrakk> = \<lbrakk>tab\<^sub>1 (r \<star> s) \<star> rs.cmp\<rbrakk>"
using Maps.Comp_eq_iso_class_memb Maps.CLS_hcomp cmp.is_map rs.cmp_def
by auto
also have "... = Maps.Comp \<lbrakk>tab\<^sub>1 r\<rbrakk> \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using Maps.Comp_eq_iso_class_memb
Maps.CLS_hcomp [of "tab\<^sub>1 r" rs.\<rho>\<sigma>.p\<^sub>1] iso_class_eqI rs.cmp_props(5)
by auto
finally show ?thesis by simp
qed
also have "... = Maps.Map Dom.leg1"
proof -
have "Maps.seq rs.R\<^sub>1 (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1)"
by (intro Maps.seqI, simp_all add: rs.composable)
moreover have "\<lbrakk>prj\<^sub>1 (Maps.REP rs.S\<^sub>1) (Maps.REP rs.R\<^sub>0)\<rbrakk> = \<lbrakk>rs.\<rho>\<sigma>.p\<^sub>1\<rbrakk>"
using rs.prj_tab_agreement iso_class_eqI by auto
moreover have "Maps.Dom (Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1) = src rs.\<rho>\<sigma>.p\<^sub>1"
using rs.prj_tab_agreement Maps.PRJ\<^sub>1_def rs.composable
isomorphic_implies_hpar(3)
by auto
ultimately show ?thesis
using SPN_r_SPN_s_leg1_eq Maps.comp_char [of rs.R\<^sub>1 "Maps.PRJ\<^sub>1 rs.R\<^sub>0 rs.S\<^sub>1"]
rs.R\<^sub>1_def Maps.PRJ\<^sub>1_def rs.composable CMP_def
by simp
qed
finally show ?thesis by simp
(*
* Very simple, right? Yeah, once you sort through the notational morass and
* figure out what equals what.
*)
qed
qed
qed
show "Span.in_hom (CMP r s) (HoSPN_SPN.map (r, s)) (SPNoH.map (r, s))"
using Span.arr_char arrow_of_spans_axioms Span.dom_char Span.cod_char
CMP_def SPN.FF_def VV.arr_char rs.composable
by auto
thus "Span.in_hhom (CMP r s) (SPN.map\<^sub>0 (src s)) (SPN.map\<^sub>0 (trg r))"
using assms VV.ide_char VV.arr_char VV.in_hom_char SPN.FF_def
Span.src_dom [of "CMP r s"] Span.trg_dom [of "CMP r s"]
by fastforce
qed
lemma compositor_simps [simp]:
assumes "ide r" and "ide s" and "src r = trg s"
shows "Span.arr (CMP r s)"
and "Span.src (CMP r s) = SPN.map\<^sub>0 (src s)" and "Span.trg (CMP r s) = SPN.map\<^sub>0 (trg r)"
and "Span.dom (CMP r s) = HoSPN_SPN.map (r, s)"
and "Span.cod (CMP r s) = SPNoH.map (r, s)"
using assms compositor_in_hom [of r s] by auto
lemma compositor_is_iso:
assumes "ide r" and "ide s" and "src r = trg s"
shows "Span.iso (CMP r s)"
proof -
interpret rs: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg r s
using assms by unfold_locales auto
have "Span.arr (CMP r s)"
using assms compositor_in_hom by blast
moreover have "Maps.iso \<lbrakk>\<lbrakk>rs.cmp\<rbrakk>\<rbrakk>"
using assms Maps.iso_char'
by (metis (mono_tags, lifting) Maps.CLS_in_hom Maps.Map.simps(1) Maps.in_homE
equivalence_is_left_adjoint ide_in_iso_class rs.cmp_props(3) rs.cmp_simps(2))
ultimately show ?thesis
unfolding CMP_def
using assms Span.iso_char by simp
qed
interpretation \<Xi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
proof
fix rs
assume rs: "VV.ide rs"
let ?r = "fst rs"
let ?s = "snd rs"
show "Span.in_hom (CMP ?r ?s) (HoSPN_SPN.map rs) (SPNoH.map rs)"
using rs compositor_in_hom [of ?r ?s] VV.ide_char VV.arr_char by simp
next
fix \<mu>\<nu>
assume \<mu>\<nu>: "VV.arr \<mu>\<nu>"
let ?\<mu> = "fst \<mu>\<nu>"
let ?\<nu> = "snd \<mu>\<nu>"
show "CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)) \<bullet> HoSPN_SPN.map \<mu>\<nu> =
SPNoH.map \<mu>\<nu> \<bullet> CMP (fst (VV.dom \<mu>\<nu>)) (snd (VV.dom \<mu>\<nu>))"
proof -
let ?LHS = "CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)) \<bullet> HoSPN_SPN.map \<mu>\<nu>"
let ?RHS = "SPNoH.map \<mu>\<nu> \<bullet> CMP (fst (VV.dom \<mu>\<nu>)) (snd (VV.dom \<mu>\<nu>))"
have LHS:
"Span.in_hom ?LHS (HoSPN_SPN.map (VV.dom \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
proof
show "Span.in_hom (HoSPN_SPN.map \<mu>\<nu>) (HoSPN_SPN.map (VV.dom \<mu>\<nu>))
(HoSPN_SPN.map (VV.cod \<mu>\<nu>))"
using \<mu>\<nu> by blast
show "Span.in_hom (CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>)))
(HoSPN_SPN.map (VV.cod \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
using \<mu>\<nu> VV.cod_simp by (auto simp add: VV.arr_char)
qed
have RHS:
"Span.in_hom ?RHS (HoSPN_SPN.map (VV.dom \<mu>\<nu>)) (SPNoH.map (VV.cod \<mu>\<nu>))"
using \<mu>\<nu> VV.dom_simp VV.cod_simp by (auto simp add: VV.arr_char)
show "?LHS = ?RHS"
proof (intro Span.arr_eqI)
show "Span.par ?LHS ?RHS"
apply (intro conjI)
using LHS RHS apply auto[2]
proof -
show "Span.dom ?LHS = Span.dom ?RHS"
proof -
have "Span.dom ?LHS = HoSPN_SPN.map (VV.dom \<mu>\<nu>)"
using LHS by auto
also have "... = Span.dom ?RHS"
using RHS by auto
finally show ?thesis by simp
qed
show "Span.cod ?LHS = Span.cod ?RHS"
proof -
have "Span.cod ?LHS = SPNoH.map (VV.cod \<mu>\<nu>)"
using LHS by auto
also have "... = Span.cod ?RHS"
using RHS by auto
finally show ?thesis by simp
qed
qed
show "Chn ?LHS = Chn ?RHS"
proof -
interpret dom_\<mu>_\<nu>: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>dom ?\<mu>\<close> \<open>dom ?\<nu>\<close>
using \<mu>\<nu> VV.ide_char VV.arr_char by unfold_locales auto
interpret cod_\<mu>_\<nu>: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>cod ?\<mu>\<close> \<open>cod ?\<nu>\<close>
using \<mu>\<nu> VV.ide_char VV.arr_char by unfold_locales auto
interpret \<mu>_\<nu>: horizontal_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<close> \<open>tab_of_ide (dom ?\<mu>)\<close> \<open>tab\<^sub>0 (dom ?\<mu>)\<close> \<open>tab\<^sub>1 (dom ?\<mu>)\<close>
\<open>dom ?\<nu>\<close> \<open>tab_of_ide (dom ?\<nu>)\<close> \<open>tab\<^sub>0 (dom ?\<nu>)\<close> \<open>tab\<^sub>1 (dom ?\<nu>)\<close>
\<open>cod ?\<mu>\<close> \<open>tab_of_ide (cod ?\<mu>)\<close> \<open>tab\<^sub>0 (cod ?\<mu>)\<close> \<open>tab\<^sub>1 (cod ?\<mu>)\<close>
\<open>cod ?\<nu>\<close> \<open>tab_of_ide (cod ?\<nu>)\<close> \<open>tab\<^sub>0 (cod ?\<nu>)\<close> \<open>tab\<^sub>1 (cod ?\<nu>)\<close>
?\<mu> ?\<nu>
using \<mu>\<nu> VV.arr_char by unfold_locales auto
let ?\<mu>\<nu> = "?\<mu> \<star> ?\<nu>"
interpret dom_\<mu>\<nu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom ?\<mu>\<nu>\<close>
using \<mu>\<nu> by (unfold_locales, simp)
interpret cod_\<mu>\<nu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod ?\<mu>\<nu>\<close>
using \<mu>\<nu> by (unfold_locales, simp)
interpret \<mu>\<nu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close> \<open>tab_of_ide (dom ?\<mu>\<nu>)\<close> \<open>tab\<^sub>0 (dom ?\<mu>\<nu>)\<close> \<open>tab\<^sub>1 (dom ?\<mu>\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close> \<open>tab_of_ide (cod ?\<mu>\<nu>)\<close> \<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close> \<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
?\<mu>\<nu>
using \<mu>\<nu> by unfold_locales auto
have Chn_LHS_eq:
"Chn ?LHS = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot> Span.chine_hcomp (SPN (fst \<mu>\<nu>)) (SPN (snd \<mu>\<nu>))"
proof -
have "Chn ?LHS = Chn (CMP (fst (VV.cod \<mu>\<nu>)) (snd (VV.cod \<mu>\<nu>))) \<odot>
Chn (HoSPN_SPN.map \<mu>\<nu>)"
using \<mu>\<nu> LHS Span.Chn_vcomp by blast
also have "... = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot> Chn (HoSPN_SPN.map \<mu>\<nu>)"
using \<mu>\<nu> VV.arr_char VV.cod_char CMP_def by simp
also have "... = \<lbrakk>\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>\<rbrakk> \<odot>
Span.chine_hcomp (SPN (fst \<mu>\<nu>)) (SPN (snd \<mu>\<nu>))"
using \<mu>\<nu> VV.arr_char SPN.FF_def Span.hcomp_def by simp
finally show ?thesis by blast
qed
have Chn_RHS_eq:
"Chn ?RHS = Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof -
have "Chn ?RHS = Chn (SPN (?\<mu> \<star> ?\<nu>)) \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using \<mu>\<nu> RHS Span.vcomp_def VV.arr_char CMP_def Span.arr_char Span.not_arr_Null
VV.dom_simp
by auto
moreover have "Chn (SPN (?\<mu> \<star> ?\<nu>)) =
Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
proof -
have "Chn (SPN (?\<mu> \<star> ?\<nu>)) =
Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>spn ?\<mu>\<nu>\<rbrakk>"
using \<mu>\<nu> SPN_def by simp
also have "... = Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using spn_def by simp
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
let ?Chn_LHS =
"Maps.MkArr (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<odot>
Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)"
let ?Chn_RHS =
"Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>))) (src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<odot>
Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
have "?Chn_LHS = ?Chn_RHS"
proof (intro Maps.arr_eqI)
interpret LHS: arrow_of_spans Maps.comp ?LHS
using LHS Span.arr_char by auto
interpret RHS: arrow_of_spans Maps.comp ?RHS
using RHS Span.arr_char by auto
show 1: "Maps.arr ?Chn_LHS"
using LHS.chine_in_hom Chn_LHS_eq by auto
show 2: "Maps.arr ?Chn_RHS"
using RHS.chine_in_hom Chn_RHS_eq by auto
text \<open>
Here is where we use \<open>dom_\<mu>_\<nu>.chine_hcomp_SPN_SPN\<close>,
which depends on our having chosen the ``right'' pullbacks for \<open>Maps(B)\<close>.
The map \<open>Chn_LHS\<close> has as its domain the apex of the
horizontal composite of the components of @{term "VV.dom \<mu>\<nu>"},
whereas \<open>Chn_RHS\<close> has as its
domain the apex of the chosen tabulation of \<open>r\<^sub>0\<^sup>* \<star> s\<^sub>1\<close>.
We need these to be equal in order for \<open>Chn_LHS\<close> and \<open>Chn_RHS\<close> to be equal.
\<close>
show "Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS"
proof -
have "Maps.Dom ?Chn_LHS = Maps.Dom (Maps.dom ?Chn_LHS)"
using \<mu>\<nu> 1 Maps.Dom_dom by presburger
also have
"... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>)))"
proof -
have "... = Maps.Dom (Maps.dom (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)))"
using 1 Maps.seq_char Maps.Dom_comp by auto
also have "... = Maps.Dom (Maps.pbdom (Leg0 (Dom (SPN ?\<mu>)))
(Leg1 (Dom (SPN ?\<nu>))))"
using \<mu>\<nu> VV.arr_char Span.chine_hcomp_in_hom [of "SPN ?\<nu>" "SPN ?\<mu>"]
by auto
also have "... = Maps.Dom (Maps.dom (Maps.pbdom (Leg0 (Dom (SPN ?\<mu>)))
(Leg1 (Dom (SPN ?\<nu>)))))"
proof -
have "Maps.cospan (Leg0 (Dom (SPN (fst \<mu>\<nu>)))) (Leg1 (Dom (SPN (snd \<mu>\<nu>))))"
using \<mu>\<nu> VV.arr_char SPN_in_hom Span.arr_char Span.dom_char SPN_def
Maps.CLS_in_hom Maps.arr_char Maps.cod_char dom_\<mu>_\<nu>.composable
dom_\<mu>_\<nu>.RS_simps(16) dom_\<mu>_\<nu>.S\<^sub>1_def dom_\<mu>_\<nu>.RS_simps(1)
dom_\<mu>_\<nu>.R\<^sub>0_def Maps.pbdom_in_hom
by simp
thus ?thesis
using \<mu>\<nu> VV.arr_char Maps.pbdom_in_hom by simp
qed
also have "... = Maps.Dom
(Maps.dom (Maps.pbdom (Leg0 (Dom (SPN (dom ?\<mu>))))
(Leg1 (Dom (SPN (dom ?\<nu>))))))"
using \<mu>\<nu> SPN_def VV.arr_char by simp
also have "... = Maps.Dom
(Maps.dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>))))"
using \<mu>\<nu> VV.arr_char ide_dom
by (simp add: Span.chine_hcomp_ide_ide)
also have "... = Maps.Dom (Span.chine_hcomp (SPN (dom ?\<mu>)) (SPN (dom ?\<nu>)))"
using Maps.Dom_dom Maps.in_homE SPN.preserves_reflects_arr SPN.preserves_src
SPN.preserves_trg Span.chine_hcomp_in_hom dom_\<mu>_\<nu>.composable
dom_\<mu>_\<nu>.r.base_simps(2) dom_\<mu>_\<nu>.s.base_simps(2)
by (metis (no_types, lifting))
finally show ?thesis by simp
qed
also have "... = src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"
using "dom_\<mu>_\<nu>.chine_hcomp_SPN_SPN" by simp
also have "... = Maps.Dom ?Chn_RHS"
using 2 Maps.seq_char Maps.Dom_comp by auto
finally show ?thesis by simp
qed
show "Maps.Cod ?Chn_LHS = Maps.Cod ?Chn_RHS"
proof -
have "Maps.Cod ?Chn_LHS = src (tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>))"
using \<mu>\<nu> 1 VV.arr_char Maps.seq_char by auto
also have "... = src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>))"
using \<mu>\<nu> VV.arr_char cod_\<mu>\<nu>.tab_simps(2) by auto
also have "... = Maps.Cod ?Chn_RHS"
by (metis (no_types, lifting) "2" Maps.Cod.simps(1) Maps.Cod_comp Maps.seq_char)
finally show ?thesis by simp
qed
show "Maps.Map ?Chn_LHS = Maps.Map ?Chn_RHS"
proof -
have RHS: "Maps.Map ?Chn_RHS = iso_class (\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp)"
proof -
have "Maps.Map ?Chn_RHS = Maps.Comp \<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using \<mu>\<nu> 2 VV.arr_char Maps.Map_comp
Maps.comp_char
[of "Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
"Maps.MkArr (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)
(src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))
\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"]
by simp
also have "... = \<lbrakk>\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof -
have "\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))"
proof -
have "\<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.tab) (src (tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)))"
using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_props(1-3)
by (metis (mono_tags, lifting) equivalence_is_left_adjoint mem_Collect_eq)
thus ?thesis
using \<mu>\<nu> VV.arr_char dom_\<mu>\<nu>.tab_simps(2) by simp
qed
moreover have "\<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<in>
Maps.Hom (src (tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)))
(src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))"
using \<mu>\<nu> VV.arr_char \<mu>\<nu>.chine_in_hom \<mu>\<nu>.is_map by auto
moreover have
"\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<in> Maps.Comp \<lbrakk>\<mu>\<nu>.chine\<rbrakk> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
proof
show "is_iso_class \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using is_iso_classI by simp
show "is_iso_class \<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using is_iso_classI by simp
show "dom_\<mu>_\<nu>.cmp \<in> \<lbrakk>dom_\<mu>_\<nu>.cmp\<rbrakk>"
using ide_in_iso_class [of dom_\<mu>_\<nu>.cmp] by simp
show "\<mu>\<nu>.chine \<in> \<lbrakk>\<mu>\<nu>.chine\<rbrakk>"
using ide_in_iso_class by simp
show "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> \<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp"
using \<mu>\<nu> VV.arr_char \<mu>\<nu>.chine_simps dom_\<mu>_\<nu>.cmp_simps dom_\<mu>\<nu>.tab_simps(2)
isomorphic_reflexive
by auto
qed
ultimately show ?thesis
using \<mu>\<nu> dom_\<mu>_\<nu>.cmp_props \<mu>\<nu>.chine_in_hom \<mu>\<nu>.chine_is_induced_map
Maps.Comp_eq_iso_class_memb
by blast
qed
finally show ?thesis by simp
qed
have LHS: "Maps.Map ?Chn_LHS = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof -
have "Maps.Map ?Chn_LHS =
Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>
(Maps.Map
(Maps.tuple (Maps.CLS (spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1))
(Maps.CLS (tab\<^sub>0 (cod ?\<mu>)))
(Maps.CLS (tab\<^sub>1 (cod ?\<nu>)))
(Maps.CLS (spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0))))"
proof -
have "Maps.Map ?Chn_LHS =
Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>
(Maps.Map (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)))"
using \<mu>\<nu> 1 VV.arr_char Maps.Map_comp cod_\<mu>\<nu>.tab_simps(2)
Maps.comp_char
[of "Maps.MkArr (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)
(src (tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)))
\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
"Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)"]
by simp
moreover have "Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>) =
Maps.tuple
(Maps.CLS (spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1))
(Maps.CLS (tab\<^sub>0 (cod ?\<mu>)))
(Maps.CLS (tab\<^sub>1 (cod ?\<nu>)))
(Maps.CLS (spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0))"
proof -
have "Maps.PRJ\<^sub>0
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<and>
Maps.PRJ\<^sub>1
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
interpret X: identity_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>(tab\<^sub>0 (dom ?\<mu>))\<^sup>* \<star> tab\<^sub>1 (dom ?\<nu>)\<close>
using \<mu>\<nu> VV.arr_char
by (unfold_locales, simp)
have "Maps.PRJ\<^sub>0
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>tab\<^sub>0 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>0_def
using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
dom_\<mu>_\<nu>.RS_simps(18) dom_\<mu>_\<nu>.RS_simps(3) dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.S\<^sub>1_def
by auto
moreover
have "Maps.PRJ\<^sub>1
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>)
(Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>) \<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) =
\<lbrakk>\<lbrakk>tab\<^sub>1 ((Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>))\<rbrakk>\<rbrakk>"
unfolding Maps.PRJ\<^sub>1_def
using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.RS_simps(1) dom_\<mu>_\<nu>.RS_simps(16)
dom_\<mu>_\<nu>.RS_simps(18) dom_\<mu>_\<nu>.RS_simps(3) dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.S\<^sub>1_def
by auto
moreover
have "(Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<mu>))) (trg (snd \<mu>\<nu>))
\<lbrakk>tab\<^sub>0 (dom ?\<mu>)\<rbrakk>))\<^sup>* \<star>
Maps.REP (Maps.MkArr (src (tab\<^sub>0 (dom ?\<nu>))) (trg ?\<nu>)
\<lbrakk>tab\<^sub>1 (dom ?\<nu>)\<rbrakk>) \<cong>
(tab\<^sub>0 (dom ?\<mu>))\<^sup>* \<star> tab\<^sub>1 (dom ?\<nu>)"
using VV.arr_char \<mu>\<nu> dom_\<mu>_\<nu>.S\<^sub>1_def dom_\<mu>_\<nu>.s.leg1_simps(3)
dom_\<mu>_\<nu>.s.leg1_simps(4) trg_dom dom_\<mu>_\<nu>.R\<^sub>0_def
dom_\<mu>_\<nu>.prj_tab_agreement(1) isomorphic_symmetric
by simp
ultimately show ?thesis
using X.isomorphic_implies_same_tab isomorphic_symmetric by metis
qed
thus ?thesis
unfolding Span.chine_hcomp_def
using \<mu>\<nu> VV.arr_char SPN_def isomorphic_reflexive
Maps.comp_CLS [of "spn ?\<mu>" dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1 "spn ?\<mu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of "spn ?\<nu>" dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 "spn ?\<nu> \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"]
by simp
qed
moreover have "Maps.Dom (Span.chine_hcomp (SPN ?\<mu>) (SPN ?\<nu>)) =
src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0"
by (metis (no_types, lifting) "1" "2" Maps.Dom.simps(1) Maps.comp_char
\<open>Maps.Dom ?Chn_LHS = Maps.Dom ?Chn_RHS\<close>)
ultimately show ?thesis by simp
qed
also have "... = Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
proof -
let ?tuple = "Maps.tuple \<lbrakk>\<lbrakk>spn (fst \<mu>\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 (cod ?\<mu>)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (cod ?\<nu>)\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>spn (snd \<mu>\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
have "iso_class \<mu>_\<nu>.chine = Maps.Map ?tuple"
using \<mu>_\<nu>.CLS_chine spn_def Maps.Map.simps(1)
by (metis (no_types, lifting))
thus ?thesis by simp
qed
also have "... = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof -
have "\<lbrakk>\<mu>_\<nu>.chine\<rbrakk> \<in> Maps.Hom (src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0)"
proof -
have "\<guillemotleft>\<mu>_\<nu>.chine : src dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 \<rightarrow> src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<guillemotright>"
using \<mu>\<nu> VV.arr_char by simp
thus ?thesis
using \<mu>_\<nu>.is_map ide_in_iso_class left_adjoint_is_ide by blast
qed
moreover have "\<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<in>
Maps.Hom (src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0) (src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)))"
proof -
have "\<guillemotleft>cod_\<mu>_\<nu>.cmp : src cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0 \<rightarrow> src (tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>))\<guillemotright>"
using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_in_hom cod_\<mu>\<nu>.tab_simps(2)
by simp
thus ?thesis
using cod_\<mu>_\<nu>.cmp_props equivalence_is_left_adjoint left_adjoint_is_ide
ide_in_iso_class [of cod_\<mu>_\<nu>.cmp]
by (metis (mono_tags, lifting) mem_Collect_eq)
qed
moreover have
"cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine \<in> Maps.Comp \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
proof
show "is_iso_class \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
using \<mu>_\<nu>.w_simps(1) is_iso_classI by blast
show "is_iso_class \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
using cod_\<mu>_\<nu>.cmp_simps(2) is_iso_classI by blast
show "\<mu>_\<nu>.chine \<in> \<lbrakk>\<mu>_\<nu>.chine\<rbrakk>"
using ide_in_iso_class by simp
show "cod_\<mu>_\<nu>.cmp \<in> \<lbrakk>cod_\<mu>_\<nu>.cmp\<rbrakk>"
using ide_in_iso_class by simp
show "cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
by (simp add: isomorphic_reflexive)
qed
ultimately show ?thesis
using \<mu>\<nu> cod_\<mu>_\<nu>.cmp_props \<mu>_\<nu>.chine_in_hom \<mu>_\<nu>.chine_is_induced_map
Maps.Comp_eq_iso_class_memb
by simp
qed
finally show ?thesis by simp
qed
have EQ: "\<lbrakk>\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp\<rbrakk> = \<lbrakk>cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine\<rbrakk>"
proof (intro iso_class_eqI)
show "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
proof -
interpret dom_cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
dom_\<mu>_\<nu>.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>tab_of_ide (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>tab\<^sub>0 (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>tab\<^sub>1 (dom ?\<mu> \<star> dom ?\<nu>)\<close>
\<open>dom ?\<mu>\<nu>\<close>
using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_interpretation by simp
interpret cod_cmp: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>_\<nu>.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (cod ?\<nu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (cod ?\<mu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
\<open>tab_of_ide (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>tab\<^sub>0 (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu> \<star> cod ?\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close>
using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_interpretation by simp
interpret L: vertical_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
\<open>dom_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>tab_of_ide (dom ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>0 (dom ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (dom ?\<mu>\<nu>)\<close>
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>\<nu>.tab
\<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
\<open>dom ?\<mu>\<nu>\<close>
\<open>?\<mu> \<star> ?\<nu>\<close>
using \<mu>\<nu> VV.arr_char dom_\<mu>_\<nu>.cmp_in_hom
by unfold_locales auto
interpret R: vertical_composite_of_arrows_of_tabulations_in_maps
V H \<a> \<i> src trg
\<open>dom ?\<mu>\<nu>\<close>
\<open>dom_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (dom ?\<nu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (dom ?\<mu>) \<star> dom_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
\<open>cod_\<mu>_\<nu>.\<rho>\<sigma>.tab\<close>
\<open>tab\<^sub>0 (cod ?\<nu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>0\<close>
\<open>tab\<^sub>1 (cod ?\<mu>) \<star> cod_\<mu>_\<nu>.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>cod ?\<mu>\<nu>\<close>
cod_\<mu>\<nu>.tab
\<open>tab\<^sub>0 (cod ?\<mu>\<nu>)\<close>
\<open>tab\<^sub>1 (cod ?\<mu>\<nu>)\<close>
\<open>?\<mu> \<star> ?\<nu>\<close>
\<open>cod ?\<mu>\<nu>\<close>
using \<mu>\<nu> VV.arr_char cod_\<mu>_\<nu>.cmp_in_hom
by unfold_locales auto
have "\<mu>\<nu>.chine \<star> dom_\<mu>_\<nu>.cmp \<cong> L.chine"
using \<mu>\<nu> VV.arr_char L.chine_char dom_\<mu>_\<nu>.cmp_def isomorphic_symmetric
by simp
also have "... = R.chine"
using L.is_ide \<mu>\<nu> comp_arr_dom comp_cod_arr isomorphic_reflexive by force
also have "... \<cong> cod_\<mu>_\<nu>.cmp \<star> \<mu>_\<nu>.chine"
using \<mu>\<nu> VV.arr_char R.chine_char cod_\<mu>_\<nu>.cmp_def by simp
finally show ?thesis by simp
qed
qed
show ?thesis
using LHS RHS EQ by simp
qed
qed
thus ?thesis
using Chn_LHS_eq Chn_RHS_eq by simp
qed
qed
qed
qed
interpretation \<Xi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Xi>.map
using VV.ide_char VV.arr_char \<Xi>.map_simp_ide compositor_is_iso
by (unfold_locales, simp)
lemma compositor_is_natural_transformation:
shows "transformation_by_components VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map
(\<lambda>rs. CMP (fst rs) (snd rs))"
..
lemma compositor_is_natural_isomorphism:
shows "natural_isomorphism VV.comp Span.vcomp HoSPN_SPN.map SPNoH.map \<Xi>.map"
..
end
subsubsection "Associativity Coherence"
locale three_composable_identities_in_bicategory_of_spans =
bicategory_of_spans V H \<a> \<i> src trg +
f: identity_in_bicategory_of_spans V H \<a> \<i> src trg f +
g: identity_in_bicategory_of_spans V H \<a> \<i> src trg g +
h: identity_in_bicategory_of_spans V H \<a> \<i> src trg h
for V :: "'a comp" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and f :: 'a
and g :: 'a
and h :: 'a +
assumes fg: "src f = trg g"
and gh: "src g = trg h"
begin
interpretation f: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f
using f.is_arrow_of_tabulations_in_maps by simp
interpretation h: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
h h.tab \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close> h h.tab \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close> h
using h.is_arrow_of_tabulations_in_maps by simp
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 50)
interpretation SPN: "functor" V Span.vcomp SPN
using SPN_is_functor by simp
interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
using SPN_is_weak_arrow_of_homs by simp
interpretation SPN_SPN: "functor" VV.comp Span.VV.comp SPN.FF
using SPN.functor_FF by auto
interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
SPN.FF \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<circ> snd \<mu>\<nu>\<close>
..
interpretation SPNoH: composite_functor VV.comp V Span.vcomp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
text \<open>
Here come a lot of interpretations for ``composite things''.
We need these in order to have relatively short, systematic names for entities that will
appear in the lemmas to follow.
The names of the interpretations use a prefix notation, where \<open>H\<close> refers to horizontal
composition of 1-cells and \<open>T\<close> refers to composite of tabulations.
So, for example, \<open>THfgh\<close> refers to the composite of the tabulation associated with the
horizontal composition \<open>f \<star> g\<close> with the tabulation associated with \<open>h\<close>.
\<close>
interpretation HHfgh: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>(f \<star> g) \<star> h\<close>
using fg gh by unfold_locales auto
interpretation HfHgh: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation Tfg: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg f g
using fg gh by unfold_locales auto
interpretation Tgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg g h
using fg gh by unfold_locales auto
interpretation THfgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>f \<star> g\<close> h
using fg gh by unfold_locales auto
interpretation THfgh: tabulation V H \<a> \<i> src trg \<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
using THfgh.\<rho>\<sigma>.composite_is_tabulation by simp
interpretation TfHgh: two_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg
f \<open>g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TfHgh: tabulation V H \<a> \<i> src trg \<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab
\<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
using TfHgh.\<rho>\<sigma>.composite_is_tabulation by simp
interpretation Tfg_Hfg: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g\<close> \<open>tab_of_ide (f \<star> g)\<close> \<open>tab\<^sub>0 (f \<star> g)\<close> \<open>tab\<^sub>1 (f \<star> g)\<close>
\<open>f \<star> g\<close>
by unfold_locales auto
interpretation Tgh_Hgh: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>g \<star> h\<close> \<open>tab_of_ide (g \<star> h)\<close> \<open>tab\<^sub>0 (g \<star> h)\<close> \<open>tab\<^sub>1 (g \<star> h)\<close>
\<open>g \<star> h\<close>
by unfold_locales auto
interpretation THfgh_HHfgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>tab_of_ide ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>0 ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>1 ((f \<star> g) \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TfHgh_HfHgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> \<open>tab_of_ide (f \<star> g \<star> h)\<close> \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpretation TTfgh: composite_tabulation_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
using fg gh by unfold_locales auto
interpretation TTfgh_THfgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>f \<star> g\<close> Tfg.\<rho>\<sigma>.tab \<open>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
\<open>f \<star> g\<close> \<open>tab_of_ide (f \<star> g)\<close> \<open>tab\<^sub>0 (f \<star> g)\<close> \<open>tab\<^sub>1 (f \<star> g)\<close>
h \<open>tab_of_ide h\<close> \<open>tab\<^sub>0 h\<close> \<open>tab\<^sub>1 h\<close>
\<open>f \<star> g\<close> h
..
interpretation TfTgh: composite_tabulation_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
using fg gh by unfold_locales auto
interpretation TfTgh_TfHgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> \<open>tab_of_ide (g \<star> h)\<close> \<open>tab\<^sub>0 (g \<star> h)\<close> \<open>tab\<^sub>1 (g \<star> h)\<close>
f \<open>g \<star> h\<close>
..
interpretation TfTgh_TfTgh:
horizontal_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>tab_of_ide f\<close> \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close>
\<open>g \<star> h\<close> Tgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close>
f \<open>g \<star> h\<close>
..
text \<open>
The following interpretation defines the associativity between the peaks
of the two composite tabulations \<open>TTfgh\<close> (associated to the left) and \<open>TfTgh\<close>
(associated to the right).
\<close>
(* TODO: Try to get rid of the .\<rho>\<sigma> in, e.g., Tfg.\<rho>\<sigma>.p\<^sub>1. *)
interpretation TTfgh_TfTgh:
arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.\<rho>\<sigma>.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfTgh.tab \<open>(tab\<^sub>0 h \<star> Tgh.\<rho>\<sigma>.p\<^sub>0) \<star> TfTgh.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfTgh.p\<^sub>1\<close>
\<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
text \<open>
This interpretation defines the map, from the apex of the tabulation associated
with the horizontal composite \<open>(f \<star> g) \<star> h\<close> to the apex of the tabulation associated
with the horizontal composite \<open>f \<star> g \<star> h\<close>, induced by the associativity isomorphism
\<open>\<a>[f, g, h]\<close> from \<open>(f \<star> g) \<star> h\<close> to \<open>f \<star> g \<star> h\<close>.
\<close>
interpretation HHfgh_HfHgh: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom (\<alpha> (f, g, h))\<close> \<open>tab_of_ide (dom (\<alpha> (f, g, h)))\<close>
\<open>tab\<^sub>0 (dom (\<alpha> (f, g, h)))\<close> \<open>tab\<^sub>1 (dom (\<alpha> (f, g, h)))\<close>
\<open>cod (\<alpha> (f, g, h))\<close> \<open>tab_of_ide (cod (\<alpha> (f, g, h)))\<close>
\<open>tab\<^sub>0 (cod (\<alpha> (f, g, h)))\<close> \<open>tab\<^sub>1 (cod (\<alpha> (f, g, h)))\<close>
\<open>\<alpha> (f, g, h)\<close>
proof -
have "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
((f \<star> g) \<star> h) (tab_of_ide ((f \<star> g) \<star> h)) (tab\<^sub>0 ((f \<star> g) \<star> h)) (tab\<^sub>1 ((f \<star> g) \<star> h))
(f \<star> g \<star> h) (tab_of_ide (f \<star> g \<star> h)) (tab\<^sub>0 (f \<star> g \<star> h)) (tab\<^sub>1 (f \<star> g \<star> h))
\<a>[f, g, h]"
using fg gh by unfold_locales auto
thus "arrow_of_tabulations_in_maps V H \<a> \<i> src trg
(dom (\<alpha> (f, g, h))) (tab_of_ide (dom (\<alpha> (f, g, h))))
(tab\<^sub>0 (dom (\<alpha> (f, g, h)))) (tab\<^sub>1 (dom (\<alpha> (f, g, h))))
(cod (\<alpha> (f, g, h))) (tab_of_ide (cod (\<alpha> (f, g, h))))
(tab\<^sub>0 (cod (\<alpha> (f, g, h)))) (tab\<^sub>1 (cod (\<alpha> (f, g, h))))
(\<alpha> (f, g, h))"
using fg gh \<alpha>_def by auto
qed
interpretation SPN_f: arrow_of_spans Maps.comp \<open>SPN f\<close>
using SPN_in_hom Span.arr_char [of "SPN f"] by simp
interpretation SPN_g: arrow_of_spans Maps.comp \<open>SPN g\<close>
using SPN_in_hom Span.arr_char [of "SPN g"] by simp
interpretation SPN_h: arrow_of_spans Maps.comp \<open>SPN h\<close>
using SPN_in_hom Span.arr_char [of "SPN h"] by simp
interpretation SPN_fgh: three_composable_identity_arrows_of_spans Maps.comp
Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 \<open>SPN f\<close> \<open>SPN g\<close> \<open>SPN h\<close>
using fg gh Span.arr_char SPN_in_hom SPN.preserves_ide Span.ide_char
apply unfold_locales by auto
text \<open>
The following relates the projections associated with the composite span \<open>SPN_fgh\<close>
with tabulations in the underlying bicategory.
\<close>
lemma prj_char:
shows "SPN_fgh.Prj\<^sub>1\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0 = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>1 = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>1\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
and "SPN_fgh.Prj\<^sub>0\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
show "SPN_fgh.Prj\<^sub>1\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "ide (Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
by (metis TTfgh.composable TTfgh.leg1_simps(2) Tfg.\<rho>\<sigma>.T0.antipar(2)
Tfg.\<rho>\<sigma>.T0.ide_right Tfg_Hfg.u_simps(3) f.T0.antipar(2) f.T0.ide_right
f.u_simps(3) fg g.ide_leg1 g.leg1_simps(4) h.ide_leg1 h.leg1_simps(4)
ide_hcomp hseqE hcomp_simps(1) tab\<^sub>1_simps(1))
thus ?thesis
using fg gh Tfg.\<rho>\<sigma>.prj_char TTfgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
Maps.comp_CLS [of Tfg.\<rho>\<sigma>.p\<^sub>1 TTfgh.p\<^sub>1 "Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1"]
by (simp add: TTfgh.composable Tfg.\<rho>\<sigma>.T0.antipar(2))
qed
show "SPN_fgh.Prj\<^sub>0\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "ide (Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
by (metis TTfgh.leg1_simps(2) bicategory_of_spans.tab\<^sub>0_simps(1)
bicategory_of_spans.tab\<^sub>1_simps(1) bicategory_of_spans_axioms
Tfg.\<rho>\<sigma>.T0.antipar(2) Tfg.\<rho>\<sigma>.T0.ide_right Tfg.composable f.T0.antipar(2)
f.T0.ide_right f.u_simps(3) g.ide_leg1 g.leg1_simps(4)
Tfg.u_simps(3) THfgh.composable h.ide_leg1 h.leg1_simps(4)
ide_hcomp hseqE hcomp_simps(1) tab\<^sub>1_simps(3))
thus ?thesis
using fg gh Tfg.\<rho>\<sigma>.prj_char TTfgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
Maps.comp_CLS [of Tfg.\<rho>\<sigma>.p\<^sub>0 TTfgh.p\<^sub>1 "Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1"]
by (simp add: Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable)
qed
show "SPN_fgh.Prj\<^sub>0 = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using isomorphic_reflexive TTfgh.prj_char Tfg.\<rho>\<sigma>.prj_char
Maps.comp_CLS [of "tab\<^sub>0 g" Tfg.\<rho>\<sigma>.p\<^sub>0 "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
by (simp add: Tfg.composable)
show "SPN_fgh.Prj\<^sub>1 = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using Tgh.\<rho>\<sigma>.prj_char isomorphic_reflexive Tgh.composable
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
TfTgh.prj_char
by simp
show "SPN_fgh.Prj\<^sub>1\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.prj_char TfTgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of Tgh.\<rho>\<sigma>.p\<^sub>1 TfTgh.p\<^sub>0 "Tgh.\<rho>\<sigma>.p\<^sub>1 \<star> TfTgh.p\<^sub>0"]
by simp
show "SPN_fgh.Prj\<^sub>0\<^sub>0 = \<lbrakk>\<lbrakk>Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.prj_char TfTgh.prj_char isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" "Tgh.\<rho>\<sigma>.p\<^sub>1" "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Maps.comp_CLS [of Tgh.\<rho>\<sigma>.p\<^sub>0 TfTgh.p\<^sub>0 "Tgh.\<rho>\<sigma>.p\<^sub>0 \<star> TfTgh.p\<^sub>0"]
by simp
qed
interpretation \<Phi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
using compositor_is_natural_transformation by simp
interpretation \<Phi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Phi>.map
using compositor_is_natural_isomorphism by simp
(*
* TODO: Figure out how this subcategory gets introduced.
* The simps in the locale are used in the subsequent proofs.
*)
interpretation VVV': subcategory VxVxV.comp
\<open>\<lambda>\<tau>\<mu>\<nu>. arr (fst \<tau>\<mu>\<nu>) \<and> arr (fst (snd \<tau>\<mu>\<nu>)) \<and> arr (snd (snd \<tau>\<mu>\<nu>)) \<and>
src (fst (snd \<tau>\<mu>\<nu>)) = trg (snd (snd \<tau>\<mu>\<nu>)) \<and>
src (fst \<tau>\<mu>\<nu>) = trg (fst (snd \<tau>\<mu>\<nu>))\<close>
using fg gh VVV.arr_char VV.arr_char VVV.subcategory_axioms by simp
text \<open>
We define abbreviations for the left and right-hand sides of the equation for
associativity coherence.
\<close>
(*
* TODO: \<Phi> doesn't really belong in this locale. Replace it with CMP and rearrange
* material so that this locale comes first and the definition of \<Phi> comes later
* in bicategory_of_spans.
*)
abbreviation LHS
where "LHS \<equiv> SPN \<a>[f, g, h] \<bullet> \<Phi>.map (f \<star> g, h) \<bullet> (\<Phi>.map (f, g) \<circ> SPN h)"
abbreviation RHS
where "RHS \<equiv> \<Phi>.map (f, g \<star> h) \<bullet> (SPN f \<circ> \<Phi>.map (g, h)) \<bullet>
Span.assoc (SPN f) (SPN g) (SPN h)"
lemma arr_LHS:
shows "Span.arr LHS"
using fg gh VVV.arr_char VVV.ide_char VV.arr_char VV.ide_char Span.hseqI'
HoHV_def compositor_in_hom \<alpha>_def
apply (intro Span.seqI)
apply simp_all
using SPN.FF_def
apply simp
proof -
have "SPN ((f \<star> g) \<star> h) = Span.cod (CMP (f \<star> g) h)"
using fg gh compositor_in_hom by simp
also have "... = Span.cod (CMP (f \<star> g) h \<bullet> (CMP f g \<circ> SPN h))"
proof -
have "Span.seq (CMP (f \<star> g) h) (CMP f g \<circ> SPN h)"
proof (intro Span.seqI Span.hseqI)
show 1: "Span.in_hhom (SPN h) (SPN.map\<^sub>0 (src h)) (SPN.map\<^sub>0 (trg h))"
using SPN.preserves_src SPN.preserves_trg by simp
show 2: "Span.in_hhom (CMP f g) (SPN.map\<^sub>0 (trg h)) (SPN.map\<^sub>0 (trg f))"
using compositor_in_hom SPN_fgh.\<nu>\<pi>.composable fg by auto
show 3: "Span.arr (CMP (f \<star> g) h)"
using TTfgh.composable Tfg.\<rho>\<sigma>.ide_base compositor_simps(1) h.is_ide by auto
show "Span.dom (CMP (f \<star> g) h) = Span.cod (CMP f g \<circ> SPN h)"
using 1 2 3 fg gh compositor_in_hom SPN_fgh.\<nu>\<pi>.composable SPN_in_hom SPN.FF_def
by auto
qed
thus ?thesis by simp
qed
finally show "SPN ((f \<star> g) \<star> h) = Span.cod (CMP (f \<star> g) h \<bullet> (CMP f g \<circ> SPN h))"
by blast
qed
lemma arr_RHS:
shows "Span.arr RHS"
using fg gh VV.ide_char VV.arr_char \<Phi>.map_simp_ide SPN.FF_def Span.hseqI'
by (intro Span.seqI, simp_all)
lemma par_LHS_RHS:
shows "Span.par LHS RHS"
proof (intro conjI)
show "Span.arr LHS"
using arr_LHS by simp
show "Span.arr RHS"
using arr_RHS by simp
show "Span.dom LHS = Span.dom RHS"
proof -
have "Span.dom LHS = Span.dom (\<Phi>.map (f, g) \<circ> SPN h)"
using arr_LHS by auto
also have "... = Span.dom (\<Phi>.map (f, g)) \<circ> Span.dom (SPN h)"
using arr_LHS Span.dom_hcomp [of "SPN h" "\<Phi>.map (f, g)"] by blast
also have "... = (SPN f \<circ> SPN g) \<circ> SPN h"
using fg \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def by simp
also have "... = Span.dom (Span.assoc (SPN f) (SPN g) (SPN h))"
using fg gh VVV.arr_char VVV.ide_char VV.arr_char VV.ide_char by simp
also have "... = Span.dom RHS"
using \<open>Span.arr RHS\<close> by auto
finally show ?thesis by blast
qed
show "Span.cod LHS = Span.cod RHS"
proof -
have "Span.cod LHS = Span.cod (SPN \<a>[f, g, h])"
using arr_LHS by simp
also have "... = SPN (f \<star> g \<star> h)"
unfolding \<alpha>_def
using fg gh VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char HoVH_def
by simp
also have "... = Span.cod RHS"
using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def
compositor_in_hom
by simp
finally show ?thesis by blast
qed
qed
lemma Chn_LHS_eq:
shows "Chn LHS =
\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Chn LHS = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot>
Span.chine_hcomp (CMP f g) (SPN h)"
proof -
have "Chn LHS = Chn (SPN \<a>[f, g, h]) \<odot> Chn (CMP (f \<star> g) h) \<odot>
Chn (CMP f g \<circ> SPN h)"
using fg gh arr_LHS \<Phi>.map_simp_ide VV.ide_char VV.arr_char Span.Chn_vcomp
by auto
moreover have "Chn (SPN \<a>[f, g, h]) = Maps.CLS HHfgh_HfHgh.chine"
using fg gh SPN_def VVV.arr_char VV.arr_char spn_def \<alpha>_def by simp
moreover have "Chn (CMP (f \<star> g) h) = Maps.CLS THfgh_HHfgh.chine"
using fg gh CMP_def THfgh.cmp_def by simp
moreover have "Chn (CMP f g \<circ> SPN h) = Span.chine_hcomp (CMP f g) (SPN h)"
using fg gh Span.hcomp_def by simp
ultimately show ?thesis by simp
qed
moreover have "Span.chine_hcomp (CMP f g) (SPN h) = \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Span.chine_hcomp (CMP f g) (SPN h) =
Maps.tuple
(\<lbrakk>\<lbrakk>Tfg.cmp\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)"
proof -
have "\<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk>"
using fg isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>0 g" "Tfg.\<rho>\<sigma>.p\<^sub>0" "tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0"]
by simp
moreover have "span_in_category.apex Maps.comp
\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 h\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>\<rparr> =
\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk>"
proof -
interpret h: span_in_category Maps.comp \<open>\<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 h\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>\<rparr>\<close>
using h.determines_span by simp
interpret dom_h: identity_arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom h\<close> \<open>tab_of_ide (dom h)\<close> \<open>tab\<^sub>0 (dom h)\<close> \<open>tab\<^sub>1 (dom h)\<close>
\<open>cod h\<close> \<open>tab_of_ide (cod h)\<close> \<open>tab\<^sub>0 (cod h)\<close> \<open>tab\<^sub>1 (cod h)\<close>
h
by (simp add: h.is_arrow_of_tabulations_in_maps
identity_arrow_of_tabulations_in_maps.intro
identity_arrow_of_tabulations_in_maps_axioms.intro)
have "Maps.arr h.leg0"
using h.leg_simps(1) by simp
hence "Maps.dom h.leg0 = \<lbrakk>\<lbrakk>dom_h.chine\<rbrakk>\<rbrakk>"
using Maps.dom_char Maps.CLS_in_hom
apply simp
proof -
have "h.is_induced_map (src (tab\<^sub>0 h))"
using h.is_induced_map_iff dom_h.\<Delta>_eq_\<rho> h.apex_is_induced_by_cell by force
hence "src (tab\<^sub>0 h) \<cong> h.chine"
using h.chine_is_induced_map h.induced_map_unique by simp
thus "\<lbrakk>src (tab\<^sub>0 h)\<rbrakk> = \<lbrakk>h.chine\<rbrakk>"
using iso_class_eqI by simp
qed
thus ?thesis
using h.apex_def spn_def by simp
qed
ultimately show ?thesis
unfolding Span.chine_hcomp_def
using fg gh CMP_def Tfg.\<rho>\<sigma>.prj_char Span.hcomp_def by simp
qed
also have "... = \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "\<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
Maps.tuple \<lbrakk>\<lbrakk>Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>h.chine \<star> TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using TTfgh_THfgh.CLS_chine by simp
also have "... =
Maps.tuple (\<lbrakk>\<lbrakk>Tfg_Hfg.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>h.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>)"
proof -
have "\<lbrakk>\<lbrakk>Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg_Hfg.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint TTfgh.p\<^sub>1"
using Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable by simp
moreover have "Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1 \<cong> Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1"
using TTfgh_THfgh.prj_chine(2) isomorphic_reflexive isomorphic_implies_hpar(2)
by blast
ultimately show ?thesis
using Tfg_Hfg.is_map
Maps.comp_CLS [of Tfg_Hfg.chine TTfgh.p\<^sub>1 "Tfg_Hfg.chine \<star> TTfgh.p\<^sub>1"]
by simp
qed
moreover have "\<lbrakk>\<lbrakk>h.chine \<star> TTfgh.p\<^sub>0\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>h.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint TTfgh.p\<^sub>0"
by (simp add: Tfg.\<rho>\<sigma>.T0.antipar(2) THfgh.composable)
moreover have "h.chine \<star> TTfgh.p\<^sub>0 \<cong> h.chine \<star> TTfgh.p\<^sub>0"
using TTfgh_THfgh.prj_chine(1) isomorphic_reflexive isomorphic_implies_hpar(2)
by blast
ultimately show ?thesis
using h.is_map Maps.comp_CLS [of h.chine TTfgh.p\<^sub>0 "h.chine \<star> TTfgh.p\<^sub>0"]
by simp
qed
ultimately show ?thesis by argo
qed
also have "... =
Maps.tuple (\<lbrakk>\<lbrakk>Tfg.cmp\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>1 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 (f \<star> g)\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>spn h\<rbrakk>\<rbrakk> \<odot> Maps.PRJ\<^sub>0 \<lbrakk>\<lbrakk>tab\<^sub>0 g \<star> Tfg.\<rho>\<sigma>.p\<^sub>0\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>)"
using Tfg.cmp_def spn_def TTfgh.prj_char by simp
finally show ?thesis by simp
qed
finally show ?thesis by blast
qed
ultimately show ?thesis by simp
qed
abbreviation tuple_BC
where "tuple_BC \<equiv> Maps.tuple SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 SPN_fgh.Prj\<^sub>0"
abbreviation tuple_ABC
where "tuple_ABC \<equiv> Maps.tuple SPN_fgh.Prj\<^sub>1\<^sub>1
SPN_fgh.\<mu>.leg0
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
tuple_BC"
abbreviation tuple_BC'
where "tuple_BC' \<equiv> Maps.tuple \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
abbreviation tuple_ABC'
where "tuple_ABC' \<equiv> Maps.tuple \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>
tuple_BC'"
lemma csq:
shows "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
and "Maps.commutative_square SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof -
show 1: "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
proof
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.legs_form_cospan(1) by simp
show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
using SPN_fgh.prj_simps(2-3,5-6) by presburger
show "Maps.dom SPN_fgh.\<nu>.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1"
using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2)
by auto
show "SPN_fgh.\<nu>.leg0 \<odot> SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\<pi>.leg1 \<odot> SPN_fgh.Prj\<^sub>0"
by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
Maps.pullback_commutes' SPN_fgh.\<mu>\<nu>.dom.leg_simps(1)
SPN_fgh.\<mu>\<nu>.leg0_composite SPN_fgh.cospan_\<nu>\<pi>)
qed
show "Maps.commutative_square
SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using fg gh SPN_fgh.prj_simps(10) by blast
show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
using fg gh 1 Maps.tuple_simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(1)
SPN_fgh.prj_simps(4) SPN_fgh.prj_simps(5)
by presburger
show "Maps.dom SPN_fgh.\<mu>.leg0 = Maps.cod SPN_fgh.Prj\<^sub>1\<^sub>1"
using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto
show "SPN_fgh.\<mu>.leg0 \<odot> SPN_fgh.Prj\<^sub>1\<^sub>1 = (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_BC"
using 1 fg gh Maps.comp_assoc Maps.prj_tuple
by (metis (no_types, lifting) Maps.pullback_commutes' SPN_fgh.cospan_\<mu>\<nu>)
qed
qed
lemma tuple_ABC_eq_ABC':
shows "tuple_BC = tuple_BC'"
and "tuple_ABC = tuple_ABC'"
proof -
have "SPN_fgh.Prj\<^sub>1\<^sub>1 = \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using prj_char by simp
moreover have "SPN_fgh.\<mu>.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk>"
by simp
moreover have "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1 = \<lbrakk>\<lbrakk>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<rbrakk>\<rbrakk>"
using Tgh.\<rho>\<sigma>.prj_char isomorphic_reflexive Tgh.composable
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
by (simp add: g.T0.antipar(2))
moreover show "tuple_BC = tuple_BC'"
using prj_char Tfg.\<rho>\<sigma>.prj_char by simp
ultimately show "tuple_ABC = tuple_ABC'"
by argo
qed
lemma tuple_BC_in_hom:
shows "Maps.in_hom tuple_BC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src Tgh.\<rho>\<sigma>.p\<^sub>0))"
proof
show 1: "Maps.arr tuple_BC"
using csq(1) by simp
have 2: "Maps.commutative_square
\<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
by (metis Tfg.S\<^sub>0_def Tfg.span_legs_eq(3) Tgh.S\<^sub>1_def Tgh.span_legs_eq(4) csq(1)
prj_char(2) prj_char(3))
show "Maps.dom tuple_BC = Maps.MkIde (src TTfgh.p\<^sub>0)"
proof -
have "Maps.dom tuple_BC' = Maps.dom \<lbrakk>\<lbrakk>Tfg.\<rho>\<sigma>.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using 2 Maps.tuple_simps by simp
also have "... = Chn (Span.hcomp (Span.hcomp (SPN f) (SPN g)) (SPN h))"
using Maps.dom_char
by (metis SPN_fgh.prj_simps(5) prj_char(2))
also have "... = Maps.MkIde (src TTfgh.p\<^sub>0)"
using 1 fg gh Maps.dom_char csq(1) prj_char(3) tuple_ABC_eq_ABC'(1)
Maps.Dom.simps(1) Maps.tuple_simps(2) SPN_fgh.prj_simps(3,5-6)
by presburger
finally have "Maps.dom tuple_BC' = Maps.MkIde (src TTfgh.p\<^sub>0)"
by blast
thus ?thesis
using tuple_ABC_eq_ABC' by simp
qed
show "Maps.cod tuple_BC = Maps.MkIde (src Tgh.\<rho>\<sigma>.p\<^sub>0)"
proof -
have "Maps.cod tuple_BC' = Maps.pbdom \<lbrakk>\<lbrakk>tab\<^sub>0 g\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 h\<rbrakk>\<rbrakk>"
using 1 2 fg gh Maps.tuple_in_hom by blast
also have "... = Maps.MkIde (src Tgh.\<rho>\<sigma>.p\<^sub>0)"
using 1 2 fg gh Maps.pbdom_def
by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\<nu>\<pi>.are_identities(2)
SPN_fgh.\<nu>\<pi>.composable Span.chine_hcomp_ide_ide Tfg.S\<^sub>0_def Tfg.span_legs_eq(3)
Tgh.S\<^sub>1_def Tgh.chine_hcomp_SPN_SPN Tgh.span_legs_eq(4) g.is_ide)
finally show ?thesis
using tuple_ABC_eq_ABC' by simp
qed
qed
lemma tuple_ABC_in_hom:
shows "Maps.in_hom tuple_ABC (Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src TfTgh.p\<^sub>0))"
proof
show 1: "Maps.arr tuple_ABC"
using SPN_fgh.chine_assoc_def SPN_fgh.chine_assoc_in_hom by auto
show "Maps.dom tuple_ABC = Maps.MkIde (src TTfgh.p\<^sub>0)"
proof -
have "Maps.dom tuple_ABC = Maps.dom SPN_fgh.chine_assoc"
by (simp add: SPN_fgh.chine_assoc_def)
also have "... = Chn ((SPN f \<circ> SPN g) \<circ> SPN h)"
using SPN_fgh.chine_assoc_in_hom by blast
also have "... = Maps.MkIde (src TTfgh.p\<^sub>0)"
by (metis (lifting) Maps.Dom.simps(1) Maps.dom_char SPN_fgh.prj_simps(3)
SPN_fgh.prj_simps(6) prj_char(3))
finally show ?thesis by blast
qed
show "Maps.cod tuple_ABC = Maps.MkIde (src TfTgh.p\<^sub>0)"
proof -
have "Maps.cod tuple_ABC = Maps.cod SPN_fgh.chine_assoc"
by (simp add: SPN_fgh.chine_assoc_def)
also have 1: "... = Chn (SPN f \<circ> SPN g \<circ> SPN h)"
using SPN_fgh.chine_assoc_in_hom by blast
also have "... = Maps.MkIde (src TfTgh.p\<^sub>0)"
by (metis (lifting) Maps.Dom.simps(1) Maps.cod_char Maps.seq_char
SPN_fgh.prj_chine_assoc(1) SPN_fgh.prj_simps(1) TfTgh.leg1_in_hom(1)
TfTgh_TfTgh.u_in_hom 1 in_hhomE prj_char(4) src_hcomp)
finally show ?thesis by argo
qed
qed
lemma Chn_RHS_eq:
shows "Chn RHS = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> tuple_ABC'"
proof -
have "Chn RHS =
Chn (\<Phi>.map (f, g \<star> h)) \<odot> Chn (SPN f \<circ> \<Phi>.map (g, h)) \<odot>
Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
proof -
have "Chn RHS = Chn (\<Phi>.map (f, g \<star> h)) \<odot>
Chn ((SPN f \<circ> \<Phi>.map (g, h)) \<bullet> Span.assoc (SPN f) (SPN g) (SPN h))"
using arr_RHS Span.vcomp_eq Span.Chn_vcomp by blast
also have "... = Chn (\<Phi>.map (f, g \<star> h)) \<odot> Chn (SPN f \<circ> \<Phi>.map (g, h)) \<odot>
Chn (Span.assoc (SPN f) (SPN g) (SPN h))"
proof -
have "Span.seq (SPN f \<circ> \<Phi>.map (g, h)) (Span.assoc (SPN f) (SPN g) (SPN h))"
using arr_RHS by auto
thus ?thesis
using fg gh Span.vcomp_eq [of "SPN f \<circ> \<Phi>.map (g, h)"
"Span.assoc (SPN f) (SPN g) (SPN h)"]
by simp
qed
finally show ?thesis by blast
qed
moreover have "Chn (\<Phi>.map (f, g \<star> h)) = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk>"
using arr_RHS fg gh \<Phi>.map_simp_ide VV.ide_char VV.arr_char CMP_def TfHgh.cmp_def
by simp
moreover have "Chn (SPN f \<circ> \<Phi>.map (g, h)) = Span.chine_hcomp (SPN f) (CMP g h)"
using fg gh Span.hcomp_def \<Phi>.map_simp_ide VV.ide_char VV.arr_char SPN.FF_def
by simp
moreover have "Chn (Span.assoc (SPN f) (SPN g) (SPN h)) = tuple_ABC"
using fg gh Span.\<alpha>_ide VVV.ide_char VVV.arr_char VV.ide_char VV.arr_char
SPN_fgh.chine_assoc_def Span.\<alpha>_def
by simp
moreover have "Span.chine_hcomp (SPN f) (CMP g h) = \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Span.chine_hcomp (SPN f) (CMP g h) =
Maps.tuple
(\<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (g \<star> h)\<rbrakk>\<rbrakk>
(\<lbrakk>\<lbrakk>Tgh_Hgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>)"
proof -
interpret f: span_in_category Maps.comp
\<open>\<lparr>Leg0 = Maps.MkArr (src (tab\<^sub>0 f)) (trg g) \<lbrakk>tab\<^sub>0 f\<rbrakk>,
Leg1 = Maps.MkArr (src (tab\<^sub>0 f)) (trg f) \<lbrakk>tab\<^sub>1 f\<rbrakk>\<rparr>\<close>
using f.determines_span
by (simp add: Tfg.composable)
interpret f: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f f.tab \<open>tab\<^sub>0 f\<close> \<open>tab\<^sub>1 f\<close> f
using f.is_arrow_of_tabulations_in_maps by simp
have "f.apex = Maps.CLS f.chine"
proof (intro Maps.arr_eqI)
show "Maps.arr f.apex" by simp
show "Maps.arr \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using Maps.CLS_in_hom f.is_map by blast
show "Maps.Dom f.apex = Maps.Dom \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto
show "Maps.Cod f.apex = Maps.Cod \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
using f.apex_def Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable by auto
show "Maps.Map f.apex = Maps.Map \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
proof -
have "Maps.Map f.apex = \<lbrakk>src (tab\<^sub>0 f)\<rbrakk>"
using f.apex_def Maps.dom_char Tfg.RS_simps(2) Tfg.R\<^sub>0_def Tfg.composable
by auto
also have "... = \<lbrakk>f.chine\<rbrakk>"
proof (intro iso_class_eqI)
have "f.is_induced_map (src (tab\<^sub>0 f))"
using f.apex_is_induced_by_cell comp_cod_arr by auto
thus "src (tab\<^sub>0 f) \<cong> f.chine"
using f.induced_map_unique f.chine_is_induced_map by simp
qed
also have "... = Maps.Map \<lbrakk>\<lbrakk>f.chine\<rbrakk>\<rbrakk>"
by simp
finally show ?thesis by simp
qed
qed
thus ?thesis
unfolding Span.chine_hcomp_def
using fg gh CMP_def Tgh.\<rho>\<sigma>.prj_char Span.hcomp_def isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" Tgh.\<rho>\<sigma>.p\<^sub>1 "tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1"]
Tgh.cmp_def TfTgh.prj_char
by simp
qed
also have "... = Maps.tuple \<lbrakk>\<lbrakk>f.chine \<star> TfTgh.p\<^sub>1\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>tab\<^sub>0 f\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>tab\<^sub>1 (g \<star> h)\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>Tgh_Hgh.chine \<star> TfTgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using isomorphic_reflexive TfHgh.composable f.is_map TfHgh.composable Tgh_Hgh.is_map
Maps.comp_CLS [of f.chine TfTgh.p\<^sub>1 "f.chine \<star> TfTgh.p\<^sub>1"]
Maps.comp_CLS [of Tgh_Hgh.chine TfTgh.p\<^sub>0 "Tgh_Hgh.chine \<star> TfTgh.p\<^sub>0"]
by auto
also have "... = \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk>"
using TfTgh_TfHgh.CLS_chine by simp
finally show ?thesis by blast
qed
ultimately have "Chn RHS =\<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> tuple_ABC"
by simp
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> tuple_ABC'"
using tuple_ABC_eq_ABC' by simp
finally show ?thesis by simp
qed
interpretation g\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg \<open>tab\<^sub>1 h\<close> \<open>tab\<^sub>0 g\<close>
using gh by unfold_locales auto
interpretation f\<^sub>0g\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg \<open>tab\<^sub>1 g\<close> \<open>tab\<^sub>0 f\<close>
using fg by unfold_locales auto
interpretation f\<^sub>0gh\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>tab\<^sub>1 g \<star> Tgh.\<rho>\<sigma>.p\<^sub>1\<close> \<open>tab\<^sub>0 f\<close>
using fg gh Tgh.\<rho>\<sigma>.leg1_is_map
by unfold_locales auto
interpretation fg\<^sub>0h\<^sub>1: cospan_of_maps_in_bicategory_of_spans V H \<a> \<i> src trg
\<open>tab\<^sub>1 h\<close> \<open>tab\<^sub>0 g \<star> Tfg.p\<^sub>0\<close>
using TTfgh.r\<^sub>0s\<^sub>1_is_cospan by simp
lemma src_tab_eq:
shows "(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu> =
TTfgh.tab"
proof -
have "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu> =
(\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> TTfgh.tab"
unfolding TTfgh.tab_def
using TTfgh_TfTgh.chine_is_induced_map TTfgh.tab_def TTfgh_TfTgh.\<Delta>_simps(4)
by auto
moreover have "iso (\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0)"
by (simp add: fg gh)
moreover have "inv (\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) = \<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
using fg gh by simp
ultimately show ?thesis
using TTfgh_TfTgh.\<Delta>_simps(1)
invert_side_of_triangle(1)
[of "TfTgh.composite_cell TTfgh_TfTgh.chine TTfgh_TfTgh.the_\<theta> \<cdot> TTfgh_TfTgh.the_\<nu>"
"\<a>[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0" TTfgh.tab]
by argo
qed
text \<open>
We need to show that the associativity isomorphism (defined in terms of tupling) coincides
with \<open>TTfgh_TfTgh.chine\<close> (defined in terms of tabulations). In order to do this,
we need to know how the latter commutes with projections. That is the purpose of
the following lemma. Unfortunately, it requires some lengthy calculations,
which I haven't seen any way to avoid.
\<close>
lemma prj_chine:
shows "\<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
and "\<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have 1: "ide TfTgh.p\<^sub>1"
by (simp add: TfTgh.composable)
have 2: "ide TTfgh_TfTgh.chine"
by simp
have 3: "src TfTgh.p\<^sub>1 = trg TTfgh_TfTgh.chine"
using TTfgh_TfTgh.chine_in_hom(1) by simp
have 4: "src (tab\<^sub>1 f) = trg TfTgh.p\<^sub>1"
using TfTgh.leg1_simps(2) by blast
text \<open>
The required isomorphisms will each be established via \<open>T2\<close>, using the equation
\<open>src_tab_eq\<close> (associativities omitted from diagram):
$$
\begin{array}{l}
\xymatrix{
&& \xtwocell[dddd]{}\omit{^{\rm the\_}\nu}
& \scriptstyle{{\rm TTfgh}.{\rm apex}} \ar[dd]^{{\rm chine}} \ar[dddlll]_{{\rm TfTgh}.p_1} \ar[dddrrr]^{{\rm TfTgh}.p_0}
& \xtwocell[dddd]{}\omit{^{\rm the\_}\theta} \\
&&&&& \\
&&& \scriptstyle{{\rm TfTgh.apex}} \ar[ddll]_{{\rm TfTgh}.p_1} \ar[dr]^{{\rm TfTgh}.p_0} && \\
\scriptstyle{f.{\rm apex}} \ar[dd]_{f.{\rm tab}_1}
&& \dtwocell\omit{^<-7>{f_0gh_1.\phi}}
&& \scriptstyle{{\rm Tgh.apex}} \ar[dl]_{{\rm Tgh}.p_1} \ar[dr]^{{\rm Tgh}.p_0} \ddtwocell\omit{^{g_0h_1.\phi}}
&& \scriptstyle{h.{\rm apex}} \ar[dd]^{h.{\rm tab}_0} \\
& \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}}
&& \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}}
&& \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0} \dtwocell\omit{^h.{\rm tab}} \\
\scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f}
&& \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h}
}
\\
\\
\hspace{7cm}=
\\
\\
\xymatrix{
&&& \scriptstyle{{\rm TTfgh.apex}} \ar[dl]_{{\rm TTfgh}.p_1} \ar[ddrr]^{{\rm TTfgh}.p_0} && \\
&& \scriptstyle{{\rm Tfg.apex}} \ar[dl]_{{\rm Tfg}.p_1} \ar[dr]^{{\rm Tfg}.p_0} \ddtwocell\omit{^{f_0g_1.\phi}}
& \dtwocell\omit{^<-7>{fg_0h_1.\phi}} &&& \\
& \scriptstyle{f.{\rm apex}} \ar[dl]_{f.{\rm tab}_1} \ar[dr]^{f.{\rm tab}_0} \dtwocell\omit{^f.{\rm tab}}
&& \scriptstyle{g.{\rm apex}} \ar[dl]_{g.{\rm tab}_1} \ar[dr]^{g.{\rm tab}_0} \dtwocell\omit{^g.{\rm tab}}
&& \scriptstyle{h.{\rm apex}} \ar[dl]_{h.{\rm tab}_1} \ar[dr]^{h.{\rm tab}_0} \dtwocell\omit{^h.{\rm tab}} \\
\scriptstyle{{\rm trg}~f} && \scriptstyle{{\rm src}~f = {\rm trg}~g} \ar[ll]^{f}
&& \scriptstyle{{\rm src}~g = {\rm trg}~h} \ar[ll]^{g} && \scriptstyle{{\rm src}~h} \ar[ll]^{h}
}
\end{array}
$$
There is a sequential dependence between the proofs, such as we have already
seen for \<open>horizontal_composite_of_arrows_of_tabulations_in_maps.prj_chine\<close>.
\<close>
define u\<^sub>f where "u\<^sub>f = g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define w\<^sub>f where "w\<^sub>f = Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1"
define w\<^sub>f' where "w\<^sub>f' = TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>f
where "\<theta>\<^sub>f = (g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot> (g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot> (g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> (\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> (f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
define \<theta>\<^sub>f'
where "\<theta>\<^sub>f' = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]"
define \<beta>\<^sub>f
where "\<beta>\<^sub>f = \<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
have w\<^sub>f: "ide w\<^sub>f"
using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
have w\<^sub>f_is_map: "is_left_adjoint w\<^sub>f"
using w\<^sub>f_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (simp add: left_adjoints_compose)
have w\<^sub>f': "ide w\<^sub>f'"
unfolding w\<^sub>f'_def by simp
have w\<^sub>f'_is_map: "is_left_adjoint w\<^sub>f'"
unfolding w\<^sub>f'_def
using 3 TTfgh_TfTgh.is_map f\<^sub>0gh\<^sub>1.leg1_is_map
by (simp add: left_adjoints_compose)
have \<theta>\<^sub>f: "\<guillemotleft>\<theta>\<^sub>f : tab\<^sub>0 f \<star> w\<^sub>f \<Rightarrow> u\<^sub>f\<guillemotright>"
proof (unfold \<theta>\<^sub>f_def w\<^sub>f_def u\<^sub>f_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using f\<^sub>0g\<^sub>1.leg1_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto
show "\<guillemotleft>f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using f\<^sub>0g\<^sub>1.\<phi>_in_hom(2) Tfg.\<rho>\<sigma>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> ((g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using Tfg.\<rho>\<sigma>.T0.antipar(1)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1 :
((g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (g \<star> tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
(g \<star> tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> g \<star> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>g \<star> fg\<^sub>0h\<^sub>1.\<phi> : g \<star> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> g \<star> tab\<^sub>1 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>g \<star> h.tab \<star> TTfgh.p\<^sub>0 : g \<star> tab\<^sub>1 h \<star> TTfgh.p\<^sub>0 \<Rightarrow> g \<star> (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0\<guillemotright>"
using gh fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] :
g \<star> (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0 \<Rightarrow> g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using gh fg\<^sub>0h\<^sub>1.\<phi>_in_hom fg\<^sub>0h\<^sub>1.p\<^sub>1_simps
by (intro hcomp_in_vhom, auto)
qed
have \<theta>\<^sub>f': "\<guillemotleft>\<theta>\<^sub>f' : tab\<^sub>0 f \<star> w\<^sub>f' \<Rightarrow> u\<^sub>f\<guillemotright>"
proof (unfold \<theta>\<^sub>f'_def w\<^sub>f'_def u\<^sub>f_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine \<Rightarrow> (tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using "1" "2" "3" "4" assoc'_in_hom(2) f.ide_u f.leg1_simps(3) by auto
show "\<guillemotleft>f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine \<Rightarrow>
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.\<phi>_in_hom(2)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> (((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>(\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
(((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan g\<^sub>0h\<^sub>1.\<phi>_in_hom(2)
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine :
((g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan
by (intro hcomp_in_vhom, auto)
show "\<guillemotleft>can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) :
((g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> g \<star> h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan by auto
show "\<guillemotleft>g \<star> h \<star> TTfgh_TfTgh.the_\<theta> :
g \<star> h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> g \<star> h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<guillemotright>"
using f\<^sub>0gh\<^sub>1.cospan g\<^sub>0h\<^sub>1.cospan TTfgh_TfTgh.the_\<theta>_in_hom
by (intro hcomp_in_vhom, auto)
qed
have \<beta>\<^sub>f: "\<guillemotleft>\<beta>\<^sub>f : tab\<^sub>1 f \<star> w\<^sub>f \<Rightarrow> tab\<^sub>1 f \<star> w\<^sub>f'\<guillemotright>"
proof (unfold \<beta>\<^sub>f_def w\<^sub>f_def w\<^sub>f'_def, intro comp_in_homI)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
tab\<^sub>1 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using TTfgh.leg1_in_hom(2) assoc'_in_hom by auto
show "\<guillemotleft>TTfgh_TfTgh.the_\<nu> :
(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using TTfgh_TfTgh.the_\<nu>_in_hom TTfgh_TfTgh.the_\<nu>_props by simp
show "\<guillemotleft>\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
(tab\<^sub>1 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine \<Rightarrow> tab\<^sub>1 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using 1 2 3 4 by auto
qed
have iso_\<beta>\<^sub>f: "iso \<beta>\<^sub>f"
unfolding \<beta>\<^sub>f_def
using 1 2 3 4 \<beta>\<^sub>f \<beta>\<^sub>f_def isos_compose
apply (intro isos_compose)
apply (metis TTfgh.composable TTfgh.leg1_in_hom(2) Tfg.\<rho>\<sigma>.T0.antipar(2)
Tfg.\<rho>\<sigma>.T0.ide_right Tfg.\<rho>\<sigma>.leg1_in_hom(2) Tfg_Hfg.u_simps(3)
f.T0.antipar(2) f.T0.ide_right f.ide_leg1 f\<^sub>0g\<^sub>1.cospan g.ide_leg1
h.ide_leg1 h.leg1_simps(4) hcomp_in_vhomE ide_hcomp
iso_assoc' tab\<^sub>1_simps(1))
using TTfgh_TfTgh.the_\<nu>_props(2) f.ide_leg1 iso_assoc by blast+
have u\<^sub>f: "ide u\<^sub>f"
using \<theta>\<^sub>f ide_cod by blast
have w\<^sub>f_in_hhom: "in_hhom w\<^sub>f (src u\<^sub>f) (src (tab\<^sub>0 f))"
using u\<^sub>f w\<^sub>f u\<^sub>f_def w\<^sub>f_def by simp
have w\<^sub>f'_in_hhom: "in_hhom w\<^sub>f' (src u\<^sub>f) (src (tab\<^sub>0 f))"
using u\<^sub>f w\<^sub>f' w\<^sub>f'_def u\<^sub>f_def by simp
have 5: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>f \<Rightarrow> w\<^sub>f'\<guillemotright> \<and> \<beta>\<^sub>f = tab\<^sub>1 f \<star> \<gamma> \<and> \<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>)"
proof -
have eq\<^sub>f: "f.composite_cell w\<^sub>f \<theta>\<^sub>f = f.composite_cell w\<^sub>f' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f"
proof -
text \<open>
I don't see any alternative here to just grinding out the calculation.
The idea is to bring \<open>f.composite_cell w\<^sub>f \<theta>\<^sub>f\<close> into a form in which
\<open>src_tab_eq\<close> can be applied to eliminate \<open>\<theta>\<^sub>f\<close> in favor of \<open>\<theta>\<^sub>f'\<close>.
\<close>
have "f.composite_cell w\<^sub>f \<theta>\<^sub>f =
((f \<star> g \<star> \<a>[h, tab\<^sub>0 h, fg\<^sub>0h\<^sub>1.p\<^sub>0]) \<cdot>
(f \<star> g \<star> h.tab \<star> fg\<^sub>0h\<^sub>1.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> f\<^sub>0g\<^sub>1.p\<^sub>0, fg\<^sub>0h\<^sub>1.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, f\<^sub>0g\<^sub>1.p\<^sub>0] \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> f\<^sub>0g\<^sub>1.p\<^sub>0) \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1, fg\<^sub>0h\<^sub>1.p\<^sub>1])) \<cdot>
\<a>[f, tab\<^sub>0 f, f\<^sub>0g\<^sub>1.p\<^sub>1 \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1] \<cdot>
(f.tab \<star> f\<^sub>0g\<^sub>1.p\<^sub>1 \<star> fg\<^sub>0h\<^sub>1.p\<^sub>1)"
unfolding w\<^sub>f_def \<theta>\<^sub>f_def
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps Tgh.composable whisker_left by simp (* 12 sec, 30 sec cpu *)
also have "... =
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by simp
also have "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) =
f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, (h \<star> tab\<^sub>0 h) \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> h.tab \<star> TTfgh.p\<^sub>0)) \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc
assoc'_naturality [of f g "\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>1 h \<star> TTfgh.p\<^sub>0] \<cdot>
(f \<star> g \<star> fg\<^sub>0h\<^sub>1.\<phi>)) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc
assoc'_naturality [of f g "h.tab \<star> TTfgh.p\<^sub>0"]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1]) \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc
assoc'_naturality [of f g fg\<^sub>0h\<^sub>1.\<phi>]
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot> \<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] =
\<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot> \<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1] =
\<lbrace>(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
also have "... =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
by (intro E.eval_eqI, simp_all)
also have "... = \<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> \<a>[f, tab\<^sub>0 f \<star> Tfg.p\<^sub>1, TTfgh.p\<^sub>1] =
\<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> ((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\<phi>_in_hom
assoc_naturality [of f f\<^sub>0g\<^sub>1.\<phi> TTfgh.p\<^sub>1]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1)"
proof -
have "(f \<star> (g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> \<a>[f, tab\<^sub>1 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] =
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> ((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\<phi>_in_hom
assoc_naturality [of f "g.tab \<star> Tfg.p\<^sub>0" TTfgh.p\<^sub>1]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of f.tab Tfg.p\<^sub>1 TTfgh.p\<^sub>1] comp_assoc
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) =
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr whisker_right comp_assoc_assoc'
whisker_left [of f "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]" "\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(((\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) =
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1"
using fg fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc'
whisker_right
[of TTfgh.p\<^sub>1 "\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0]" "\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<star> TTfgh.p\<^sub>1) \<cdot>
((f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<star> TTfgh.p\<^sub>1) \<cdot>
((f.tab \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1)) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
whisker_right comp_assoc
by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>[f, (g \<star> tab\<^sub>0 g) \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, (\<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>)\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
also have "... = \<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tfg.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh.p\<^sub>1\<^bold>\<rangle>\<^bold>]\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
by (intro E.eval_eqI, auto)
also have "... = \<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) =
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0])"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
comp_cod_arr comp_assoc_assoc'
by simp
thus ?thesis by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
((f \<star> g) \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g) \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[f \<star> g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0] \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0]) \<cdot>
(f \<star> (g.tab \<star> Tfg.p\<^sub>0)) \<cdot>
(f \<star> f\<^sub>0g\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, Tfg.p\<^sub>1] \<cdot>
(f.tab \<star> Tfg.p\<^sub>1)
\<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
TTfgh.tab \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using TTfgh.tab_def Tfg.\<rho>\<sigma>.tab_def by simp
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<cdot>
(g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> Tgh.p\<^sub>0) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1)
\<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu>) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using src_tab_eq TfTgh.tab_def Tgh.\<rho>\<sigma>.tab_def comp_assoc by simp
text \<open>Now we have to make this look like \<open>f.composite_cell w\<^sub>f' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f\<close>.\<close>
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>)) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu>) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<cdot>
(g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> Tgh.p\<^sub>0) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1)
\<star> TfTgh.p\<^sub>0 =
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0)"
using fg gh whisker_right whisker_left by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1)
\<star> TTfgh_TfTgh.chine =
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine)"
(* using fg gh whisker_right [of TTfgh_TfTgh.chine] by auto (* 2 min *) *)
proof -
have "arr (\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using fg gh
by (intro seqI' comp_in_homI) auto
moreover
have "arr ((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr ((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
moreover
have "arr (\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1))"
using calculation by blast
ultimately show ?thesis
using whisker_right [of TTfgh_TfTgh.chine] TTfgh_TfTgh.is_ide by presburger
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
(f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
comp_arr_dom comp_assoc_assoc'
by simp
thus ?thesis by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f.tab \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality [of f.tab TfTgh.p\<^sub>1 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) =
f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
comp_cod_arr comp_assoc_assoc'
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
proof -
have "(\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f \<star> tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] =
\<lbrace>(\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... = \<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 f\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>])\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
by (intro E.eval_eqI, auto)
also have "... = \<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
using \<a>'_def \<alpha>_def by simp
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
(*
* This one can't be shortcut with a straight coherence-based proof,
* due to the presence of f\<^sub>0gh\<^sub>1.\<phi>, g\<^sub>0h\<^sub>1.\<phi>, h.tab, with associativities that
* do not respect their domain and codomain.
*
* I also tried to avoid distributing the "f \<star>" in advance, in order to
* reduce the number of associativity proof steps, but it then becomes
* less automatic to prove the necessary "arr" facts to do the proof.
* So unfortunately the mindless grind seems to be the path of least
* resistance.
*)
have "((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> f\<^sub>0gh\<^sub>1.\<phi>) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, tab\<^sub>0 f \<star> TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality [of f f\<^sub>0gh\<^sub>1.\<phi> TTfgh_TfTgh.chine] comp_assoc
by simp
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality [of f "(g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> \<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, ((g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality
[of f "\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality [of f "(g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
(((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
have "((f \<star> (g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
assoc'_naturality
[of f "(g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0" TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)"
proof -
(* OK, we can perhaps shortcut the last few steps... *)
have "((f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0]) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> \<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((f \<star> (g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<lbrace>((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>])
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
((\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
also have "... =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>])
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)\<rbrace>"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
apply (intro E.eval_eqI) by simp_all
also have "... =
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.p\<^sub>0_simps f\<^sub>0g\<^sub>1.p\<^sub>1_simps
\<a>'_def \<alpha>_def
by simp
finally show ?thesis
using comp_assoc by presburger
qed
finally show ?thesis
using comp_assoc by presburger
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "(f \<star> ((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) =
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
using fg gh whisker_left by simp (* 15 sec elapsed, 30 sec cpu *)
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
(f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot>
((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) =
\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(\<a>[f \<star> g, h, tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
(((f \<star> g) \<star> h) \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
proof -
have "(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> ((f \<star> g \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) =
\<a>\<^sup>-\<^sup>1[f, g, h] \<star> TTfgh_TfTgh.the_\<theta>"
using fg gh comp_arr_dom comp_cod_arr
interchange [of "\<a>\<^sup>-\<^sup>1[f, g, h]" "f \<star> g \<star> h"
"tab\<^sub>0 h \<star> TTfgh.p\<^sub>0" TTfgh_TfTgh.the_\<theta>]
by simp
also have "... = (((f \<star> g) \<star> h) \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh comp_arr_dom comp_cod_arr
interchange [of "(f \<star> g) \<star> h" "\<a>\<^sup>-\<^sup>1[f, g, h]" TTfgh_TfTgh.the_\<theta>
"((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine"]
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
(\<a>[f, g, h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0] \<cdot>
((f \<star> g) \<star> h \<star> TTfgh_TfTgh.the_\<theta>)) \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh assoc_naturality [of "f \<star> g" h TTfgh_TfTgh.the_\<theta>] comp_assoc
by simp
also have "... =
(f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using fg gh assoc_naturality [of f g "h \<star> TTfgh_TfTgh.the_\<theta>"] comp_assoc
by simp
finally show ?thesis
using comp_assoc by presburger
qed
also have "... =
((f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
f \<star> can
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>[f, g, h \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[f \<star> g, h, ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, h] \<star> ((tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[f \<star> g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> h, (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[f, (g \<star> h) \<star> (tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f \<star> \<a>[g \<star> h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0] \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> (\<a>\<^sup>-\<^sup>1[g, h, tab\<^sub>0 h \<star> Tgh.p\<^sub>0] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f \<star> ((g \<star> \<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0]) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>,
\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>,
((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>\<^bold>]
\<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>,
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>f\<^bold>\<rangle>, (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>) \<^bold>\<star> (\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>,
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star>
\<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>\<^bold>]) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... =
can (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))
(\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))"
using fg gh
apply (unfold can_def)
apply (intro E.eval_eqI)
by simp_all
also have "... =
f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
using fg gh whisker_can_left_0 by simp
finally show ?thesis by blast
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... =
(f \<star> (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]) \<cdot>
\<a>[f, tab\<^sub>0 f, TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine] \<cdot>
(f.tab \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
TTfgh_TfTgh.the_\<nu> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]"
proof -
have "((f \<star> g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(f \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(f \<star> (((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])) =
f \<star> (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine]"
proof -
have 1: "arr ((g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
using fg gh
apply (intro seqI' comp_in_homI) by auto
moreover
have 2: "arr (can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine])"
using calculation by blast
ultimately show ?thesis
using whisker_left f.ide_base by presburger
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = f.composite_cell w\<^sub>f' \<theta>\<^sub>f' \<cdot> \<beta>\<^sub>f"
unfolding w\<^sub>f'_def \<theta>\<^sub>f'_def \<beta>\<^sub>f_def
using comp_assoc by presburger
finally show ?thesis by blast
qed
show ?thesis
using w\<^sub>f w\<^sub>f' \<theta>\<^sub>f \<theta>\<^sub>f' \<beta>\<^sub>f f.T2 [of w\<^sub>f w\<^sub>f' \<theta>\<^sub>f u\<^sub>f \<theta>\<^sub>f' \<beta>\<^sub>f] eq\<^sub>f by fast
qed
obtain \<gamma>\<^sub>f where \<gamma>\<^sub>f: "\<guillemotleft>\<gamma>\<^sub>f : w\<^sub>f \<Rightarrow> w\<^sub>f'\<guillemotright> \<and> \<beta>\<^sub>f = tab\<^sub>1 f \<star> \<gamma>\<^sub>f \<and> \<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f)"
using 5 by auto
show "\<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>f"
using \<gamma>\<^sub>f BS3 w\<^sub>f_is_map w\<^sub>f'_is_map by blast
hence "isomorphic w\<^sub>f w\<^sub>f'"
using \<gamma>\<^sub>f isomorphic_def isomorphic_symmetric by auto
thus ?thesis
using w\<^sub>f w\<^sub>f_def w\<^sub>f'_def Maps.CLS_eqI isomorphic_symmetric by auto
qed
text \<open>
On to the next equation:
\[
\<open>\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>\<close>.
\]
We have to make use of the equation \<open>\<theta>\<^sub>f = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f)\<close> in this part,
similarly to how the equation \<open>src_tab_eq\<close> was used to replace
\<open>TTfgh.tab\<close> in the first part.
\<close>
define u\<^sub>g where "u\<^sub>g = h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define w\<^sub>g where "w\<^sub>g = Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1"
define w\<^sub>g' where "w\<^sub>g' = Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>g
where "\<theta>\<^sub>g = \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
define \<theta>\<^sub>g'
where "\<theta>\<^sub>g' = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
define \<beta>\<^sub>g
where "\<beta>\<^sub>g = \<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot> \<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
have u\<^sub>g: "ide u\<^sub>g"
unfolding u\<^sub>g_def by simp
have w\<^sub>g: "ide w\<^sub>g"
unfolding w\<^sub>g_def using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
have w\<^sub>g_is_map: "is_left_adjoint w\<^sub>g"
unfolding w\<^sub>g_def
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps left_adjoints_compose by simp
have w\<^sub>g': "ide w\<^sub>g'"
unfolding w\<^sub>g'_def by simp
have w\<^sub>g'_is_map: "is_left_adjoint w\<^sub>g'"
unfolding w\<^sub>g'_def
using TTfgh_TfTgh.is_map left_adjoints_compose by simp
have \<theta>\<^sub>g: "\<guillemotleft>\<theta>\<^sub>g : tab\<^sub>0 g \<star> w\<^sub>g \<Rightarrow> u\<^sub>g\<guillemotright>"
using w\<^sub>g_def u\<^sub>g_def \<theta>\<^sub>g_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\<phi>_in_hom by auto
have \<theta>\<^sub>g': "\<guillemotleft>\<theta>\<^sub>g' : tab\<^sub>0 g \<star> w\<^sub>g' \<Rightarrow> u\<^sub>g\<guillemotright>"
unfolding w\<^sub>g'_def u\<^sub>g_def \<theta>\<^sub>g'_def
by (intro comp_in_homI) auto
have w\<^sub>g_in_hhom: "in_hhom w\<^sub>g (src u\<^sub>g) (src (tab\<^sub>0 g))"
unfolding w\<^sub>g_def u\<^sub>g_def by auto
have w\<^sub>g'_in_hhom: "in_hhom w\<^sub>g' (src u\<^sub>g) (src (tab\<^sub>0 g))"
unfolding w\<^sub>g'_def u\<^sub>g_def by auto
have \<beta>\<^sub>g: "\<guillemotleft>\<beta>\<^sub>g : tab\<^sub>1 g \<star> w\<^sub>g \<Rightarrow> tab\<^sub>1 g \<star> w\<^sub>g'\<guillemotright>"
proof (unfold \<beta>\<^sub>g_def w\<^sub>g_def, intro comp_in_homI)
(* auto can solve this, but it's too slow *)
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
tab\<^sub>1 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1 :
(tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> (tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\<phi>_in_hom f\<^sub>0g\<^sub>1.\<phi>_uniqueness(2)
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] :
(tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \<gamma>\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto
show "\<guillemotleft>tab\<^sub>0 f \<star> \<gamma>\<^sub>f : tab\<^sub>0 f \<star> Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps \<gamma>\<^sub>f w\<^sub>f_def w\<^sub>f'_def by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] :
tab\<^sub>0 f \<star> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine \<Rightarrow> (tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine\<guillemotright>"
by auto
show "\<guillemotleft>f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 f \<star> TfTgh.p\<^sub>1) \<star> TTfgh_TfTgh.chine
\<Rightarrow> ((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine\<guillemotright>"
using f\<^sub>0gh\<^sub>1.\<phi>_in_hom
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] :
((tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
by auto
show "\<guillemotleft>\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
(tab\<^sub>1 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine \<Rightarrow> tab\<^sub>1 g \<star> w\<^sub>g'\<guillemotright>"
using w\<^sub>g'_def by auto
qed
have eq\<^sub>g: "g.composite_cell w\<^sub>g \<theta>\<^sub>g = g.composite_cell w\<^sub>g' \<theta>\<^sub>g' \<cdot> \<beta>\<^sub>g"
proof -
have "g.composite_cell w\<^sub>g \<theta>\<^sub>g =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot>
(h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
fg\<^sub>0h\<^sub>1.\<phi> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
unfolding w\<^sub>g_def \<theta>\<^sub>g_def by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]) \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_left
comp_assoc
by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
(\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]))) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
proof -
have "(\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) =
g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
by (simp add: comp_assoc)
qed
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]) \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1)"
using comp_assoc by presburger
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(g.tab \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1))"
using fg gh f\<^sub>0g\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>0_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc pentagon'
invert_opposite_sides_of_square
[of "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1"
"(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1])"
"\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]" "\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1]"]
by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps assoc'_naturality [of g.tab Tfg.p\<^sub>0 TTfgh.p\<^sub>1] by simp
also have "... =
(g \<star> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]) \<cdot>
(g \<star> h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
(g \<star> fg\<^sub>0h\<^sub>1.\<phi>) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
(\<a>[g, tab\<^sub>0 g, Tfg.p\<^sub>0] \<star> TTfgh.p\<^sub>1) \<cdot>
((g.tab \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot>
(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] =
((f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1]) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using comp_assoc by presburger
also have "... = ((f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
((tab\<^sub>0 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps whisker_right comp_assoc_assoc' by simp
also have "... = ((f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0g\<^sub>1.\<phi>_uniqueness comp_cod_arr by simp
also have "... = ((tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) =
f\<^sub>0g\<^sub>1.\<phi> \<cdot> inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1"
using f\<^sub>0g\<^sub>1.\<phi>_uniqueness whisker_right by simp
also have "... = (tab\<^sub>1 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1"
using f\<^sub>0g\<^sub>1.\<phi>_uniqueness comp_arr_inv' by simp
finally show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_cod_arr by simp
finally have "(f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot> (inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] = \<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<theta>\<^sub>f \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
unfolding \<theta>\<^sub>f_def using comp_assoc by presburger
also have "... = \<theta>\<^sub>f' \<cdot> (tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
using \<gamma>\<^sub>f comp_assoc by simp
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
unfolding \<theta>\<^sub>f'_def using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 f, TfTgh.p\<^sub>1, TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 f \<star> \<gamma>\<^sub>f) \<cdot>
\<a>[tab\<^sub>0 f, Tfg.p\<^sub>1, TTfgh.p\<^sub>1] \<cdot>
(inv f\<^sub>0g\<^sub>1.\<phi> \<star> TTfgh.p\<^sub>1) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]"
proof -
have "(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
(f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine) =
f\<^sub>0gh\<^sub>1.\<phi> \<star> TTfgh_TfTgh.chine"
using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_arr_dom comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by fastforce
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<beta>\<^sub>g"
unfolding \<beta>\<^sub>g_def using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<beta>\<^sub>g"
proof -
have "(((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "(g.tab \<star> Tgh.p\<^sub>1)" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "((g.tab \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>1 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of g.tab Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
proof -
have "(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc g\<^sub>0h\<^sub>1.\<phi>_in_hom by simp
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
proof -
have "(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) =
(\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1)) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_assoc comp_assoc_assoc' by simp
also have "... = (\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr comp_assoc_assoc' by simp
also have "... = (((g \<star> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1)) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine)"
using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps whisker_right comp_assoc_assoc' by simp
also have "... = (\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine"
using g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps comp_cod_arr by simp
finally show ?thesis by presburger
qed
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<beta>\<^sub>g"
using comp_assoc by presburger
also have "... = (g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
((((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using comp_assoc by presburger
also have "... = ((((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine]) \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "(((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "g \<star> g\<^sub>0h\<^sub>1.\<phi>" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of "g \<star> h.tab \<star> Tgh.p\<^sub>0" TfTgh.p\<^sub>0 TTfgh_TfTgh.chine]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
(((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "((g \<star> g\<^sub>0h\<^sub>1.\<phi>) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of g g\<^sub>0h\<^sub>1.\<phi> "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
proof -
have "((g \<star> h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>1 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
assoc'_naturality [of g "h.tab \<star> Tgh.p\<^sub>0" "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
thus ?thesis
using comp_assoc by presburger
qed
finally show ?thesis by simp
qed
thus ?thesis
using comp_assoc by presburger
qed
also have "... = ((g \<star> h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(g \<star> can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
(g \<star> (h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine])) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
proof -
have "can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
g \<star> can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
can (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g \<star> (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \<a>'_def \<alpha>_def by simp
also have "... = can (((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
unfolding can_def
using gh
apply (intro E.eval_eqI) by simp_all
finally show ?thesis by blast
qed
moreover
have "\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
can ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<a>\<^sup>-\<^sup>1[g, (h \<star> tab\<^sub>0 h) \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh f\<^sub>0gh\<^sub>1.p\<^sub>0_simps canI_associator_0 \<a>'_def \<alpha>_def by simp
also have "... = can ((\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
unfolding can_def
using gh
apply (intro E.eval_eqI) by simp_all
finally show ?thesis by blast
qed
ultimately show ?thesis
using gh whisker_can_left_0 by simp
qed
moreover have "\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
proof -
have "\<a>[g, tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[g \<star> tab\<^sub>0 g \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
((\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1] \<star> TfTgh.p\<^sub>0) \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[(g \<star> tab\<^sub>0 g) \<star> Tgh.p\<^sub>1, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g \<star> tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] =
\<lbrace>\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
((\<^bold>\<a>\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>\<^bold>] \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[(\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>g\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \<a>'_def \<alpha>_def by simp
also have "... = \<lbrace>\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 g\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>1\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
apply (intro E.eval_eqI) by simp_all
also have "... = g \<star> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using gh g\<^sub>0h\<^sub>1.p\<^sub>1_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps \<a>'_def \<alpha>_def by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc by presburger
qed
also have "... = (g \<star>
(h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>)
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>
\<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[g, tab\<^sub>0 g, Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g.tab \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<beta>\<^sub>g"
using gh whisker_left by auto (* 11 sec *)
also have "... = g.composite_cell w\<^sub>g' \<theta>\<^sub>g' \<cdot> \<beta>\<^sub>g"
unfolding w\<^sub>g'_def \<theta>\<^sub>g'_def
using comp_assoc by presburger
finally show ?thesis by blast
qed
have 6: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>g \<Rightarrow> w\<^sub>g'\<guillemotright> \<and> \<beta>\<^sub>g = tab\<^sub>1 g \<star> \<gamma> \<and> \<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>)"
using w\<^sub>g w\<^sub>g' \<theta>\<^sub>g \<theta>\<^sub>g' \<beta>\<^sub>g eq\<^sub>g g.T2 [of w\<^sub>g w\<^sub>g' \<theta>\<^sub>g u\<^sub>g \<theta>\<^sub>g' \<beta>\<^sub>g] by blast
obtain \<gamma>\<^sub>g where \<gamma>\<^sub>g: "\<guillemotleft>\<gamma>\<^sub>g : w\<^sub>g \<Rightarrow> w\<^sub>g'\<guillemotright> \<and> \<beta>\<^sub>g = tab\<^sub>1 g \<star> \<gamma>\<^sub>g \<and> \<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)"
using 6 by auto
show "\<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>g"
using \<gamma>\<^sub>g BS3 w\<^sub>g_is_map w\<^sub>g'_is_map by blast
hence "isomorphic w\<^sub>g w\<^sub>g'"
using \<gamma>\<^sub>g isomorphic_def isomorphic_symmetric by auto
thus ?thesis
using w\<^sub>g w\<^sub>g' w\<^sub>g_def w\<^sub>g'_def Maps.CLS_eqI by auto
qed
text \<open>Now the last equation: similar, but somewhat simpler.\<close>
define u\<^sub>h where "u\<^sub>h = tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define w\<^sub>h where "w\<^sub>h = TTfgh.p\<^sub>0"
define w\<^sub>h' where "w\<^sub>h' = Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
define \<theta>\<^sub>h
where "\<theta>\<^sub>h = tab\<^sub>0 h \<star> TTfgh.p\<^sub>0"
define \<theta>\<^sub>h'
where "\<theta>\<^sub>h' = TTfgh_TfTgh.the_\<theta> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
define \<beta>\<^sub>h
where "\<beta>\<^sub>h = \<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
have u\<^sub>h: "ide u\<^sub>h"
unfolding u\<^sub>h_def by simp
have w\<^sub>h: "ide w\<^sub>h"
unfolding w\<^sub>h_def by simp
have w\<^sub>h_is_map: "is_left_adjoint w\<^sub>h"
unfolding w\<^sub>h_def by simp
have w\<^sub>h': "ide w\<^sub>h'"
unfolding w\<^sub>h'_def by simp
have w\<^sub>h'_is_map: "is_left_adjoint w\<^sub>h'"
unfolding w\<^sub>h'_def
using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps TTfgh_TfTgh.is_map left_adjoints_compose by simp
have \<theta>\<^sub>h: "\<guillemotleft>\<theta>\<^sub>h : tab\<^sub>0 h \<star> w\<^sub>h \<Rightarrow> u\<^sub>h\<guillemotright>"
unfolding \<theta>\<^sub>h_def w\<^sub>h_def u\<^sub>h_def by auto
have \<theta>\<^sub>h': "\<guillemotleft>\<theta>\<^sub>h' : tab\<^sub>0 h \<star> w\<^sub>h' \<Rightarrow> u\<^sub>h\<guillemotright>"
unfolding \<theta>\<^sub>h'_def w\<^sub>h'_def u\<^sub>h_def
using g\<^sub>0h\<^sub>1.p\<^sub>0_simps f\<^sub>0gh\<^sub>1.p\<^sub>0_simps
by (intro comp_in_homI) auto
have \<beta>\<^sub>h: "\<guillemotleft>\<beta>\<^sub>h : tab\<^sub>1 h \<star> w\<^sub>h \<Rightarrow> tab\<^sub>1 h \<star> w\<^sub>h'\<guillemotright>"
proof (unfold \<beta>\<^sub>h_def w\<^sub>h_def w\<^sub>h'_def, intro comp_in_homI)
(* auto can solve this, but it's too slow *)
show "\<guillemotleft>inv fg\<^sub>0h\<^sub>1.\<phi> : tab\<^sub>1 h \<star> TTfgh.p\<^sub>0 \<Rightarrow> (tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.\<phi>_uniqueness by blast
show "\<guillemotleft>\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] :
(tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>tab\<^sub>0 g \<star> \<gamma>\<^sub>g :
tab\<^sub>0 g \<star> Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1 \<Rightarrow> tab\<^sub>0 g \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using \<gamma>\<^sub>g w\<^sub>g_def w\<^sub>g'_def fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
tab\<^sub>0 g \<star> Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
show "\<guillemotleft>g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine :
(tab\<^sub>0 g \<star> Tgh.p\<^sub>1) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> (tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by force
show "\<guillemotleft>\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] :
(tab\<^sub>1 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine
\<Rightarrow> tab\<^sub>1 h \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<guillemotright>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps by auto
qed
have eq\<^sub>h: "h.composite_cell w\<^sub>h \<theta>\<^sub>h = h.composite_cell w\<^sub>h' \<theta>\<^sub>h' \<cdot> \<beta>\<^sub>h"
proof -
text \<open>
Once again, the strategy is to form the subexpression
\[
\<open>\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]\<close>
\]
which is equal to \<open>\<theta>\<^sub>g\<close>, so that we can make use of the equation \<open>\<theta>\<^sub>g = \<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)\<close>.
\<close>
have "h.composite_cell w\<^sub>h \<theta>\<^sub>h =
(h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0)"
unfolding w\<^sub>h_def \<theta>\<^sub>h_def by simp
also have "... = \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0)"
proof -
have "(h \<star> tab\<^sub>0 h \<star> TTfgh.p\<^sub>0) \<cdot> \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] = \<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0]"
using comp_cod_arr by simp
thus ?thesis
using comp_assoc by metis
qed
also have "... = (\<a>[h, tab\<^sub>0 h, TTfgh.p\<^sub>0] \<cdot> (h.tab \<star> TTfgh.p\<^sub>0) \<cdot>
fg\<^sub>0h\<^sub>1.\<phi> \<cdot> \<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> (\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi> =
(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> ((tab\<^sub>0 g \<star> Tfg.p\<^sub>0) \<star> TTfgh.p\<^sub>1) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps comp_assoc_assoc' by simp
also have "... = (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using fg\<^sub>0h\<^sub>1.p\<^sub>1_simps fg\<^sub>0h\<^sub>1.\<phi>_uniqueness comp_cod_arr by simp
also have "... = (h.tab \<star> TTfgh.p\<^sub>0) \<cdot> (tab\<^sub>1 h \<star> TTfgh.p\<^sub>0)"
using comp_arr_inv' fg\<^sub>0h\<^sub>1.\<phi>_uniqueness by simp
also have "... = h.tab \<star> TTfgh.p\<^sub>0"
using comp_arr_dom fg\<^sub>0h\<^sub>1.p\<^sub>0_simps by simp
finally have "(h.tab \<star> TTfgh.p\<^sub>0) \<cdot> fg\<^sub>0h\<^sub>1.\<phi> \<cdot> (\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1]) \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi> =
h.tab \<star> TTfgh.p\<^sub>0"
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<theta>\<^sub>g \<cdot> \<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
unfolding \<theta>\<^sub>g_def by simp
also have "... = (\<theta>\<^sub>g' \<cdot> (tab\<^sub>0 g \<star> \<gamma>\<^sub>g)) \<cdot> \<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot> inv fg\<^sub>0h\<^sub>1.\<phi>"
using \<gamma>\<^sub>g by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
unfolding \<theta>\<^sub>g'_def
using comp_assoc by presburger
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
((\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
(h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(\<a>[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((h.tab \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using comp_assoc by presburger
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using assoc_naturality [of h.tab Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"] comp_assoc
by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
((\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine)) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "(\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) =
h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
using comp_cod_arr comp_assoc_assoc' by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine])) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using comp_assoc by presburger
also have "... = ((h \<star> TTfgh_TfTgh.the_\<theta>) \<cdot>
(h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>))) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) =
can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
using \<a>'_def \<alpha>_def by simp
also have "... = can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
can (((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
proof -
have "\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>h\<^bold>\<rangle>, \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace> =
can (((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
unfolding can_def
apply (intro E.eval_eqI) by simp_all
thus ?thesis by simp
qed
also have "... = can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
by simp
also have "... = h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
using whisker_can_left_0 by simp
finally have "can (\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> ((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(((\<^bold>\<langle>h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) \<cdot>
(\<a>\<^sup>-\<^sup>1[h \<star> tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) =
h \<star> can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)"
by simp
thus ?thesis
using comp_assoc by presburger
qed
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta> \<cdot>
can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
using whisker_left [of h] comp_assoc by simp
also have "... = (h \<star> TTfgh_TfTgh.the_\<theta> \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]) \<cdot>
\<a>[h, tab\<^sub>0 h, Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(h.tab \<star> Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>[tab\<^sub>1 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(g\<^sub>0h\<^sub>1.\<phi> \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine) \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 g, Tgh.p\<^sub>1, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine] \<cdot>
(tab\<^sub>0 g \<star> \<gamma>\<^sub>g) \<cdot>
\<a>[tab\<^sub>0 g, Tfg.p\<^sub>0, TTfgh.p\<^sub>1] \<cdot>
inv fg\<^sub>0h\<^sub>1.\<phi>"
proof -
have "can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) =
\<lbrace>\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>] \<^bold>\<cdot>
\<^bold>\<a>\<^sup>-\<^sup>1\<^bold>[\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle>, \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>, \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>\<^bold>]\<rbrace>"
unfolding can_def
apply (intro E.eval_eqI) by auto
also have "... = \<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
using \<a>'_def \<alpha>_def by simp
finally have "can (((\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle>) \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>)
(\<^bold>\<langle>tab\<^sub>0 h\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>Tgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TfTgh.p\<^sub>0\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>TTfgh_TfTgh.chine\<^bold>\<rangle>) =
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h \<star> Tgh.p\<^sub>0, TfTgh.p\<^sub>0, TTfgh_TfTgh.chine] \<cdot>
\<a>\<^sup>-\<^sup>1[tab\<^sub>0 h, Tgh.p\<^sub>0, TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine]"
by simp
thus ?thesis by simp
qed
also have "... = h.composite_cell w\<^sub>h' \<theta>\<^sub>h' \<cdot> \<beta>\<^sub>h"
unfolding w\<^sub>h'_def \<theta>\<^sub>h'_def \<beta>\<^sub>h_def
using comp_assoc by presburger
finally show ?thesis by simp
qed
have 7: "\<exists>!\<gamma>. \<guillemotleft>\<gamma> : w\<^sub>h \<Rightarrow> w\<^sub>h'\<guillemotright> \<and> \<beta>\<^sub>h = tab\<^sub>1 h \<star> \<gamma> \<and> \<theta>\<^sub>h = \<theta>\<^sub>h' \<cdot> (tab\<^sub>0 h \<star> \<gamma>)"
using w\<^sub>h w\<^sub>h' \<theta>\<^sub>h \<theta>\<^sub>h' \<beta>\<^sub>h eq\<^sub>h h.T2 [of w\<^sub>h w\<^sub>h' \<theta>\<^sub>h u\<^sub>h \<theta>\<^sub>h' \<beta>\<^sub>h] by blast
obtain \<gamma>\<^sub>h where \<gamma>\<^sub>h: "\<guillemotleft>\<gamma>\<^sub>h : w\<^sub>h \<Rightarrow> w\<^sub>h'\<guillemotright> \<and> \<beta>\<^sub>h = tab\<^sub>1 h \<star> \<gamma>\<^sub>h \<and> \<theta>\<^sub>h = \<theta>\<^sub>h' \<cdot> (tab\<^sub>0 h \<star> \<gamma>\<^sub>h)"
using 7 by auto
show "\<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
proof -
have "iso \<gamma>\<^sub>h"
using \<gamma>\<^sub>h BS3 w\<^sub>h_is_map w\<^sub>h'_is_map by blast
hence "isomorphic w\<^sub>h w\<^sub>h'"
using \<gamma>\<^sub>h isomorphic_def isomorphic_symmetric by auto
thus ?thesis
using w\<^sub>h w\<^sub>h' w\<^sub>h_def w\<^sub>h'_def Maps.CLS_eqI [of w\<^sub>h w\<^sub>h'] by simp
qed
qed
text \<open>
Finally, we can show that @{term TTfgh_TfTgh.chine} is given by tupling.
\<close>
lemma CLS_chine:
shows "\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = tuple_ABC"
proof -
have "tuple_ABC = SPN_fgh.chine_assoc"
using SPN_fgh.chine_assoc_def by simp
also have "... = \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.arr_eqI)
show "Maps.arr SPN_fgh.chine_assoc"
using SPN_fgh.chine_assoc_in_hom by auto
show "Maps.arr \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using Maps.CLS_in_hom TTfgh_TfTgh.is_map by blast
show "Maps.Dom SPN_fgh.chine_assoc = Maps.Dom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using SPN_fgh.chine_assoc_def Maps.dom_char tuple_ABC_in_hom TTfgh_TfTgh.chine_in_hom
by fastforce
show "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "Maps.Cod SPN_fgh.chine_assoc = Maps.Cod tuple_ABC"
using SPN_fgh.chine_assoc_def by simp
also have "... = src (prj\<^sub>0 (tab\<^sub>1 g \<star> prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)) (tab\<^sub>0 f))"
by (metis (lifting) Maps.Dom.simps(1) Maps.seq_char SPN_fgh.prj_chine_assoc(1)
SPN_fgh.prj_simps(1) calculation f\<^sub>0gh\<^sub>1.leg1_simps(3) prj_char(4))
also have "... = Maps.Cod \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using Maps.cod_char TTfgh_TfTgh.chine_in_hom by simp
finally show ?thesis by blast
qed
show "Maps.Map SPN_fgh.chine_assoc = Maps.Map \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have 0: "Chn (Span.hcomp (SPN f) (Span.hcomp (SPN g) (SPN h))) =
Maps.MkIde (src TfTgh.p\<^sub>0)"
using fg gh
by (metis (mono_tags, lifting) Maps.in_homE Maps.seqE SPN_fgh.prj_chine_assoc(1)
SPN_fgh.prj_simps(1) SPN_fgh.prj_simps(13) calculation tuple_ABC_in_hom)
have "tuple_ABC = \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.prj_joint_monic
[of SPN_fgh.\<mu>.leg0 "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1"
tuple_ABC "\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"])
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.\<nu>\<pi>.dom.is_span SPN_fgh.\<nu>\<pi>.leg1_composite SPN_fgh.cospan_\<mu>\<nu>
by auto
show "Maps.seq SPN_fgh.Prj\<^sub>1 tuple_ABC"
using 0 tuple_ABC_in_hom SPN_fgh.prj_in_hom(4) by auto
show "Maps.seq SPN_fgh.Prj\<^sub>1 \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof
show "Maps.in_hom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map
by blast
show "Maps.in_hom SPN_fgh.Prj\<^sub>1 \<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> SPN_fgh.\<mu>.apex"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.prj_in_hom(4) by blast
show "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
proof -
have "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = Maps.MkIde (src TfTgh.p\<^sub>0)"
by simp
also have "... = Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger
finally show ?thesis by blast
qed
show "SPN_fgh.\<mu>.apex = Maps.dom SPN_fgh.\<mu>.leg0"
using SPN_f.dom.apex_def by blast
qed
qed
have 2: "Maps.commutative_square SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1
SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
proof
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.legs_form_cospan(1) by simp
show "Maps.span SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0"
using SPN_fgh.prj_simps(2-3,5-6) by presburger
show "Maps.dom SPN_fgh.\<nu>.leg0 = Maps.cod SPN_fgh.Prj\<^sub>0\<^sub>1"
using SPN_fgh.prj_simps(8) SPN_g.dom.is_span SPN_g.dom.leg_simps(2)
by auto
show "SPN_fgh.\<nu>.leg0 \<odot> SPN_fgh.Prj\<^sub>0\<^sub>1 = SPN_fgh.\<pi>.leg1 \<odot> SPN_fgh.Prj\<^sub>0"
by (metis (no_types, lifting) Maps.cod_comp Maps.comp_assoc
Maps.pullback_commutes' SPN_fgh.\<mu>\<nu>.dom.leg_simps(1)
SPN_fgh.\<mu>\<nu>.leg0_composite SPN_fgh.cospan_\<nu>\<pi>)
qed
have 1: "Maps.commutative_square
SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1)
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using fg gh SPN_fgh.prj_simps(10) by blast
show "Maps.span SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC"
using fg gh csq(2) by blast
show "Maps.dom SPN_fgh.\<mu>.leg0 = Maps.cod SPN_fgh.Prj\<^sub>1\<^sub>1"
using fg gh SPN_f.dom.leg_simps(2) SPN_fgh.prj_simps(7) by auto
show "SPN_fgh.\<mu>.leg0 \<odot> SPN_fgh.Prj\<^sub>1\<^sub>1 =
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_BC"
using 2 fg gh Maps.comp_assoc csq(2)
Maps.prj_tuple [of SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 SPN_fgh.Prj\<^sub>0\<^sub>1 SPN_fgh.Prj\<^sub>0]
by blast
qed
show "SPN_fgh.Prj\<^sub>1 \<odot> tuple_ABC = SPN_fgh.Prj\<^sub>1 \<odot> Maps.CLS TTfgh_TfTgh.chine"
proof -
have "SPN_fgh.Prj\<^sub>1 \<odot> tuple_ABC = SPN_fgh.Prj\<^sub>1\<^sub>1"
using csq(2) by simp
also have "... = \<lbrakk>\<lbrakk>Tfg.p\<^sub>1 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using prj_char by simp
also have "... = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using prj_chine(1) by simp
also have "... = \<lbrakk>\<lbrakk>TfTgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint TfTgh.p\<^sub>1"
by (simp add: fg)
moreover have "is_left_adjoint TTfgh_TfTgh.chine"
using TTfgh_TfTgh.is_map by simp
moreover have "TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine \<cong> TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine"
using fg gh isomorphic_reflexive by simp
ultimately show ?thesis
using Maps.comp_CLS
[of TfTgh.p\<^sub>1 TTfgh_TfTgh.chine "TfTgh.p\<^sub>1 \<star> TTfgh_TfTgh.chine"]
by simp
qed
also have "... = SPN_fgh.Prj\<^sub>1 \<odot> Maps.CLS TTfgh_TfTgh.chine"
using prj_char by simp
finally show ?thesis by blast
qed
show
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_ABC =
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot> tuple_ABC =
tuple_BC"
using csq(2)
Maps.prj_tuple [of SPN_fgh.\<mu>.leg0 "SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1"
SPN_fgh.Prj\<^sub>1\<^sub>1 tuple_BC]
by simp
also have "... =
Maps.comp
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (Maps.comp SPN_fgh.\<nu>.leg1 SPN_fgh.\<nu>\<pi>.prj\<^sub>1))
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof (intro Maps.prj_joint_monic
[of SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 tuple_BC
"Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"])
show "Maps.cospan SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1"
using SPN_fgh.\<nu>\<pi>.legs_form_cospan(1) by simp
show "Maps.seq SPN_fgh.\<nu>\<pi>.prj\<^sub>1 tuple_BC"
proof
show "Maps.in_hom tuple_BC
(Maps.MkIde (src TTfgh.p\<^sub>0)) (Maps.MkIde (src Tgh.p\<^sub>0))"
using tuple_BC_in_hom by simp
show "Maps.in_hom SPN_fgh.\<nu>\<pi>.prj\<^sub>1 (Maps.MkIde (src Tgh.p\<^sub>0)) SPN_fgh.\<nu>.apex"
proof -
have "Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 = Maps.MkIde (src Tgh.p\<^sub>0)"
using fg gh Maps.pbdom_def
by (metis (no_types, lifting) SPN.preserves_ide SPN_fgh.\<nu>\<pi>.are_identities(2)
SPN_fgh.\<nu>\<pi>.composable Span.chine_hcomp_ide_ide Tgh.chine_hcomp_SPN_SPN
g.is_ide)
thus ?thesis
using SPN_fgh.\<nu>\<pi>.prj_in_hom(1) by simp
qed
qed
show "Maps.seq SPN_fgh.\<nu>\<pi>.prj\<^sub>1
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>)"
proof
show "Maps.in_hom SPN_fgh.\<nu>\<pi>.prj\<^sub>1
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1) SPN_fgh.\<nu>.apex"
using SPN_fgh.\<nu>\<pi>.prj_in_hom(1) by simp
show "Maps.in_hom
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>)
\<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1)"
proof
show "Maps.in_hom \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> \<lbrakk>\<lbrakk>src TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.chine_in_hom Maps.CLS_in_hom TTfgh_TfTgh.is_map
by blast
show "Maps.in_hom
(Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1))
\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>
(Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1)"
proof
show "Maps.cospan SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using SPN_fgh.prj_in_hom(4) by blast
show "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
Maps.pbdom SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
proof -
have "\<lbrakk>\<lbrakk>trg TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> = Maps.MkIde (src TfTgh.p\<^sub>0)"
by simp
also have "... = Maps.pbdom SPN_fgh.\<mu>.leg0
(SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using 0 Maps.pbdom_def SPN_fgh.chine_composite(2) by presburger
finally show ?thesis by blast
qed
show "Maps.pbdom SPN_fgh.\<nu>.leg0 SPN_fgh.\<pi>.leg1 =
Maps.dom (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1)"
using fg gh Maps.pbdom_def SPN_fgh.\<nu>\<pi>.apex_composite
SPN_fgh.\<nu>\<pi>.dom.apex_def SPN_fgh.\<nu>\<pi>.dom.is_span
SPN_fgh.\<nu>\<pi>.leg1_composite
by presburger
qed
qed
qed
show "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot> tuple_BC =
SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot> tuple_BC = SPN_fgh.Prj\<^sub>0"
using csq(1) by simp
also have "... = SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>0 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>Tgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.prj_char TfTgh.prj_char(1) isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \<star> Tgh.p\<^sub>1"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
Maps.comp_CLS
[of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
Maps.comp_CLS [of Tgh.p\<^sub>0 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
"Tgh.p\<^sub>0 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>TTfgh.p\<^sub>0\<rbrakk>\<rbrakk>"
using prj_chine(3) by simp
also have "... = SPN_fgh.Prj\<^sub>0"
using prj_char by simp
finally show ?thesis by argo
qed
finally show ?thesis by blast
qed
show "SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot> tuple_BC =
SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot> tuple_BC = SPN_fgh.Prj\<^sub>0\<^sub>1"
using csq(1) by simp
also have "... = SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
have "SPN_fgh.\<nu>\<pi>.prj\<^sub>1 \<odot>
Maps.PRJ\<^sub>0 SPN_fgh.\<mu>.leg0 (SPN_fgh.\<nu>.leg1 \<odot> SPN_fgh.\<nu>\<pi>.prj\<^sub>1) \<odot>
\<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>Tgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh Tgh.\<rho>\<sigma>.prj_char TfTgh.prj_char(1) isomorphic_reflexive
Maps.comp_CLS [of "tab\<^sub>1 g" "prj\<^sub>1 (tab\<^sub>1 h) (tab\<^sub>0 g)" "tab\<^sub>1 g \<star> Tgh.p\<^sub>1"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map isomorphic_reflexive
Maps.comp_CLS
[of TfTgh.p\<^sub>0 TTfgh_TfTgh.chine "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh TTfgh_TfTgh.is_map left_adjoints_compose isomorphic_reflexive
Maps.comp_CLS [of Tgh.p\<^sub>1 "TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"
"Tgh.p\<^sub>1 \<star> TfTgh.p\<^sub>0 \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = \<lbrakk>\<lbrakk>Tfg.p\<^sub>0 \<star> TTfgh.p\<^sub>1\<rbrakk>\<rbrakk>"
using prj_chine(2) by simp
also have "... = SPN_fgh.Prj\<^sub>0\<^sub>1"
using prj_char by simp
finally show ?thesis by argo
qed
finally show ?thesis by blast
qed
qed
finally show ?thesis by simp
qed
qed
thus ?thesis
using SPN_fgh.chine_assoc_def by simp
qed
qed
finally show ?thesis by simp
qed
text \<open>
At long last, we can show associativity coherence for \<open>SPN\<close>.
\<close>
lemma assoc_coherence:
shows "LHS = RHS"
proof (intro Span.arr_eqI)
show "Span.par LHS RHS"
using par_LHS_RHS by blast
show "Chn LHS = Chn RHS"
proof -
have "Chn LHS = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
using Chn_LHS_eq by simp
also have "... = \<lbrakk>\<lbrakk>HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "\<lbrakk>\<lbrakk>THfgh_HHfgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map
TTfgh_THfgh.is_map left_adjoints_compose
Maps.comp_CLS
[of THfgh_HHfgh.chine TTfgh_THfgh.chine "THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"]
by simp
moreover
have "\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk> =
\<lbrakk>\<lbrakk>HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine\<rbrakk>\<rbrakk>"
proof -
have "ide (HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine)"
proof -
have "ide (THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine)"
using fg gh HHfgh_HfHgh.is_map THfgh_HHfgh.is_map TTfgh_THfgh.is_map
left_adjoint_is_ide left_adjoints_compose
by auto
moreover have "src HHfgh_HfHgh.chine = trg (THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine)"
using fg gh HHfgh_HfHgh.chine_in_hom \<alpha>_def by auto
ultimately show ?thesis by simp
qed
thus ?thesis
using fg gh isomorphic_reflexive HHfgh_HfHgh.is_map THfgh_HHfgh.is_map
TTfgh_THfgh.is_map left_adjoints_compose
Maps.comp_CLS
[of HHfgh_HfHgh.chine "THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"
"HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"]
by auto
qed
ultimately show ?thesis by argo
qed
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
proof -
interpret A: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfTgh.tab \<open>(tab\<^sub>0 h \<star> Tgh.p\<^sub>0) \<star> TfTgh.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfTgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>\<a>[f, g, h]\<close> \<open>f \<star> g \<star> h\<close>
..
interpret B: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> TfHgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 (g \<star> h) \<star> TfHgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 f \<star> TfHgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>\<a>[f, g, h]\<close> \<open>f \<star> g \<star> h\<close>
using fg gh by unfold_locales auto
interpret C: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> HHfgh.tab \<open>tab\<^sub>0 ((f \<star> g) \<star> h)\<close> \<open>tab\<^sub>1 ((f \<star> g) \<star> h)\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
interpret D: vertical_composite_of_arrows_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>(f \<star> g) \<star> h\<close> TTfgh.tab \<open>tab\<^sub>0 h \<star> TTfgh.p\<^sub>0\<close> \<open>(tab\<^sub>1 f \<star> Tfg.p\<^sub>1) \<star> TTfgh.p\<^sub>1\<close>
\<open>(f \<star> g) \<star> h\<close> THfgh.\<rho>\<sigma>.tab \<open>tab\<^sub>0 h \<star> THfgh.\<rho>\<sigma>.p\<^sub>0\<close> \<open>tab\<^sub>1 (f \<star> g) \<star> THfgh.\<rho>\<sigma>.p\<^sub>1\<close>
\<open>f \<star> g \<star> h\<close> HfHgh.tab \<open>tab\<^sub>0 (f \<star> g \<star> h)\<close> \<open>tab\<^sub>1 (f \<star> g \<star> h)\<close>
\<open>(f \<star> g) \<star> h\<close> \<open>\<a>[f, g, h]\<close>
using fg gh by unfold_locales auto
have "HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine \<cong> D.chine"
proof -
have "D.chine \<cong> D.\<pi>.chine \<star> TTfgh_THfgh.chine"
using D.chine_char by simp
also have "... \<cong> C.chine \<star> TTfgh_THfgh.chine"
using fg gh comp_arr_dom isomorphic_reflexive by simp
also have "... \<cong> (C.\<pi>.chine \<star> THfgh_HHfgh.chine) \<star> TTfgh_THfgh.chine"
using C.chine_char hcomp_isomorphic_ide by simp
also have "... \<cong> (HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine) \<star> TTfgh_THfgh.chine"
proof -
have "C.\<pi>.chine = HHfgh_HfHgh.chine"
using fg gh comp_arr_dom comp_cod_arr \<alpha>_def by simp
hence "isomorphic C.\<pi>.chine HHfgh_HfHgh.chine"
using isomorphic_reflexive by simp
thus ?thesis
using hcomp_isomorphic_ide by simp
qed
also have "... \<cong> HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine"
proof -
have "ide HHfgh_HfHgh.chine \<and> ide THfgh_HHfgh.chine \<and> ide TTfgh_THfgh.chine"
by simp
moreover have "src HHfgh_HfHgh.chine = trg THfgh_HHfgh.chine \<and>
src THfgh_HHfgh.chine = trg TTfgh_THfgh.chine"
using fg gh HHfgh_HfHgh.chine_in_hom THfgh_HHfgh.chine_in_hom
TTfgh_THfgh.chine_in_hom \<alpha>_def
by auto
ultimately show ?thesis
using fg gh iso_assoc isomorphic_def
assoc_in_hom [of HHfgh_HfHgh.chine THfgh_HHfgh.chine TTfgh_THfgh.chine]
by auto
qed
finally show ?thesis
using isomorphic_symmetric by blast
qed
also have "... \<cong> B.chine"
proof -
have "D.chine = B.chine"
using fg gh comp_arr_dom comp_cod_arr by simp
thus ?thesis
using isomorphic_reflexive by simp
qed
also have "... \<cong> TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
proof -
have "B.chine \<cong> TfHgh_HfHgh.chine \<star> B.\<mu>.chine"
using B.chine_char by simp
also have "... \<cong> TfHgh_HfHgh.chine \<star> A.chine"
using fg gh comp_cod_arr isomorphic_reflexive by simp
also have "... \<cong> TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
using A.chine_char hcomp_ide_isomorphic by simp
finally show ?thesis by blast
qed
finally have "HHfgh_HfHgh.chine \<star> THfgh_HHfgh.chine \<star> TTfgh_THfgh.chine \<cong>
TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
by blast
thus ?thesis
using fg gh Maps.CLS_eqI isomorphic_implies_hpar(1) by blast
qed
also have "... = \<lbrakk>\<lbrakk>TfHgh_HfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TfTgh_TfHgh.chine\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>TTfgh_TfTgh.chine\<rbrakk>\<rbrakk>"
using fg gh isomorphic_reflexive TfTgh_TfHgh.is_map TfHgh_HfHgh.is_map TTfgh_TfTgh.is_map
left_adjoints_compose
Maps.comp_CLS
[of TfTgh_TfHgh.chine TTfgh_TfTgh.chine "TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"]
Maps.comp_CLS
[of TfHgh_HfHgh.chine "TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"
"TfHgh_HfHgh.chine \<star> TfTgh_TfHgh.chine \<star> TTfgh_TfTgh.chine"]
by simp
also have "... = Chn RHS"
using Chn_RHS_eq CLS_chine tuple_ABC_eq_ABC'(2) by simp
finally show ?thesis
by blast
qed
qed
end
subsubsection "SPN is an Equivalence Pseudofunctor"
context bicategory_of_spans
begin
interpretation Maps: maps_category V H \<a> \<i> src trg ..
interpretation Span: span_bicategory Maps.comp Maps.PRJ\<^sub>0 Maps.PRJ\<^sub>1 ..
no_notation Fun.comp (infixl "\<circ>" 55)
notation Span.vcomp (infixr "\<bullet>" 55)
notation Span.hcomp (infixr "\<circ>" 53)
notation Maps.comp (infixr "\<odot>" 55)
notation isomorphic (infix "\<cong>" 50)
interpretation SPN: "functor" V Span.vcomp SPN
using SPN_is_functor by simp
interpretation SPN: weak_arrow_of_homs V src trg Span.vcomp Span.src Span.trg SPN
using SPN_is_weak_arrow_of_homs by simp
interpretation HoSPN_SPN: composite_functor VV.comp Span.VV.comp Span.vcomp
SPN.FF \<open>\<lambda>\<mu>\<nu>. Span.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
..
interpretation SPNoH: composite_functor VV.comp V Span.vcomp
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> SPN
..
interpretation \<Phi>: transformation_by_components VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<open>\<lambda>rs. CMP (fst rs) (snd rs)\<close>
using compositor_is_natural_transformation by simp
interpretation \<Phi>: natural_isomorphism VV.comp Span.vcomp
HoSPN_SPN.map SPNoH.map \<Phi>.map
using compositor_is_natural_isomorphism by simp
abbreviation \<Phi>
where "\<Phi> \<equiv> \<Phi>.map"
interpretation SPN: pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>
proof
show "\<And>f g h. \<lbrakk> ide f; ide g; ide h; src f = trg g; src g = trg h \<rbrakk> \<Longrightarrow>
SPN \<a>[f, g, h] \<bullet> \<Phi> (f \<star> g, h) \<bullet> (\<Phi> (f, g) \<circ> SPN h) =
\<Phi> (f, g \<star> h) \<bullet> (SPN f \<circ> \<Phi> (g, h)) \<bullet> Span.assoc (SPN f) (SPN g) (SPN h)"
proof -
fix f g h
assume f: "ide f" and g: "ide g" and h: "ide h"
assume fg: "src f = trg g" and gh: "src g = trg h"
interpret fgh: three_composable_identities_in_bicategory_of_spans V H \<a> \<i> src trg f g h
using f g h fg gh
by (unfold_locales, simp)
show "fgh.LHS = fgh.RHS"
using fgh.assoc_coherence by simp
qed
qed
lemma SPN_is_pseudofunctor:
shows "pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>"
..
interpretation SPN: equivalence_pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>
proof
show "\<And>\<mu> \<mu>'. \<lbrakk>par \<mu> \<mu>'; SPN \<mu> = SPN \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
proof -
fix \<mu> \<mu>'
assume par: "par \<mu> \<mu>'"
assume eq: "SPN \<mu> = SPN \<mu>'"
interpret dom_\<mu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close>
using par apply unfold_locales by auto
interpret cod_\<mu>: identity_in_bicategory_of_spans V H \<a> \<i> src trg \<open>cod \<mu>\<close>
using par apply unfold_locales by auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<open>tab_of_ide (dom \<mu>)\<close> \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<open>tab_of_ide (cod \<mu>)\<close> \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>
using par apply unfold_locales by auto
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>
using par apply unfold_locales by auto
interpret \<mu>': arrow_of_tabulations_in_maps V H \<a> \<i> src trg
\<open>dom \<mu>\<close> \<open>tab_of_ide (dom \<mu>)\<close> \<open>tab\<^sub>0 (dom \<mu>)\<close> \<open>tab\<^sub>1 (dom \<mu>)\<close>
\<open>cod \<mu>\<close> \<open>tab_of_ide (cod \<mu>)\<close> \<open>tab\<^sub>0 (cod \<mu>)\<close> \<open>tab\<^sub>1 (cod \<mu>)\<close>
\<mu>'
using par apply unfold_locales by auto
interpret \<mu>': arrow_in_bicategory_of_spans V H \<a> \<i> src trg \<open>dom \<mu>\<close> \<open>cod \<mu>\<close> \<mu>'
using par apply unfold_locales by auto
have "\<mu>.chine \<cong> \<mu>'.chine"
using par eq SPN_def spn_def Maps.CLS_eqI \<mu>.is_ide by auto
hence "\<mu>.\<Delta> = \<mu>'.\<Delta>"
using \<mu>.\<Delta>_naturality \<mu>'.\<Delta>_naturality
by (metis \<mu>.\<Delta>_simps(4) \<mu>'.\<Delta>_simps(4) \<mu>.chine_is_induced_map \<mu>'.chine_is_induced_map
\<mu>.induced_map_preserved_by_iso)
thus "\<mu> = \<mu>'"
using par \<mu>.\<mu>_in_terms_of_\<Delta> \<mu>'.\<mu>_in_terms_of_\<Delta> by metis
qed
show "\<And>a'. Span.obj a' \<Longrightarrow> \<exists>a. obj a \<and> Span.equivalent_objects (SPN.map\<^sub>0 a) a'"
proof -
fix a'
assume a': "Span.obj a'"
let ?a = "Maps.Dom (Chn a')"
have a: "obj ?a"
using a' Span.obj_char Span.ide_char Maps.ide_char by blast
moreover have "Span.equivalent_objects (SPN.map\<^sub>0 ?a) a'"
proof -
have "SPN.map\<^sub>0 ?a = a'"
proof (intro Span.arr_eqI)
have "Chn (SPN.map\<^sub>0 ?a) = Chn (Span.src (SPN ?a))"
using a a' by auto
also have "... = Maps.MkIde (Maps.Dom (Chn a'))"
proof -
have "Maps.arr \<lbrakk>\<lbrakk>tab\<^sub>0 (dom (Maps.Dom (Chn a')))\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (tab\<^sub>0 (dom (Maps.Dom (Chn a'))))"
using a by auto
thus ?thesis
using Maps.CLS_in_hom by auto
qed
moreover have "arr (Maps.Dom (Chn a'))"
using a by auto
moreover have "Span.arr (SPN (Maps.Dom (Chn a')))"
using a a' SPN_in_hom by auto
ultimately show ?thesis
using a a' SPN_def Span.src_def Maps.cod_char obj_simps(2) by simp
qed
also have "... = Chn a'"
using a' Maps.MkIde_Dom Span.obj_char Span.ide_char by simp
finally show "Chn (SPN.map\<^sub>0 ?a) = Chn a'" by simp
show "Span.par (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) a'"
using a a' Span.obj_char
apply (intro conjI)
using SPN.map\<^sub>0_simps(1) Span.obj_def
apply blast
apply simp
apply (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src
\<open>Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\<close> obj_def)
by (metis (no_types, lifting) SPN.map\<^sub>0_def SPN.preserves_arr Span.obj_src
\<open>Chn (SPN.map\<^sub>0 (Maps.Dom (Chn a'))) = Chn a'\<close> obj_def)
qed
thus ?thesis
using Span.equivalent_objects_reflexive
by (simp add: a')
qed
ultimately show "\<exists>a. obj a \<and> Span.equivalent_objects (SPN.map\<^sub>0 a) a'"
by auto
qed
show "\<And>a b g. \<lbrakk>obj a; obj b; Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 b); Span.ide g\<rbrakk>
\<Longrightarrow> \<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> Span.isomorphic (SPN f) g"
proof -
fix a b g
assume a: "obj a" and b: "obj b"
and g_in_hhom: "Span.in_hhom g (SPN.map\<^sub>0 a) (SPN.map\<^sub>0 b)"
and ide_g: "Span.ide g"
have arr_a: "arr a"
using a by auto
have arr_b: "arr b"
using b by auto
show "\<exists>f. \<guillemotleft>f : a \<rightarrow> b\<guillemotright> \<and> ide f \<and> Span.isomorphic (SPN f) g"
proof -
interpret g: arrow_of_spans Maps.comp g
using ide_g Span.ide_char by blast
interpret g: identity_arrow_of_spans Maps.comp g
using ide_g Span.ide_char
by unfold_locales auto
interpret REP_leg0: map_in_bicategory V H \<a> \<i> src trg \<open>Maps.REP g.leg0\<close>
using Maps.REP_in_Map [of g.leg0]
by unfold_locales auto
have 0: "\<guillemotleft>Maps.REP g.leg0 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg0\<guillemotright>"
using g.dom.leg_in_hom Maps.REP_in_hhom
by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod
g.dom.leg_simps(1))
have 1: "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
using g.dom.leg_in_hom Maps.REP_in_hhom
by (metis (no_types, lifting) Maps.Dom_cod Maps.REP_simps(2) Maps.arr_cod
g.dom.leg_simps(3))
let ?f = "Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*"
have f_in_hhom: "\<guillemotleft>?f : a \<rightarrow> b\<guillemotright>"
proof
show "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> b\<guillemotright>"
proof -
have "\<guillemotleft>Maps.REP g.leg1 : src (Maps.REP g.apex) \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
using 1 by simp
moreover have "Maps.Cod g.leg1 = b"
proof -
have "src (Maps.REP g.dtrg) = src (Maps.REP (Leg0 (Dom (SPN.map\<^sub>0 b))))"
using g_in_hhom Span.trg_def [of g] by auto
also have "... = src (Maps.REP (Maps.cod \<lbrakk>\<lbrakk>tab\<^sub>0 b\<rbrakk>\<rbrakk>))"
using b arr_b SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>trg (tab\<^sub>0 b)\<rbrakk>\<rbrakk>)"
using b Maps.CLS_in_hom [of "tab\<^sub>0 b"] by force
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>b\<rbrakk>\<rbrakk>)"
using b by fastforce
also have "... = b"
using b obj_simps by auto
finally show ?thesis by simp
qed
ultimately show ?thesis by argo
qed
show "\<guillemotleft>(Maps.REP g.leg0)\<^sup>* : a \<rightarrow> src (Maps.REP g.apex)\<guillemotright>"
proof -
have "\<guillemotleft>Maps.REP g.leg0 : src (Maps.REP g.apex) \<rightarrow> a\<guillemotright>"
proof -
have "src (Maps.REP g.dsrc) = src (Maps.REP (Leg0 (Dom (SPN.map\<^sub>0 a))))"
using g_in_hhom Span.src_def [of g] by auto
also have "... = src (Maps.REP (Maps.cod \<lbrakk>\<lbrakk>tab\<^sub>0 a\<rbrakk>\<rbrakk>))"
using a arr_a SPN.map\<^sub>0_def Span.src_def SPN_in_hom by auto
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>trg (tab\<^sub>0 a)\<rbrakk>\<rbrakk>)"
using a Maps.CLS_in_hom [of "tab\<^sub>0 a"] by force
also have "... = src (Maps.REP \<lbrakk>\<lbrakk>a\<rbrakk>\<rbrakk>)"
using a by fastforce
also have "... = a"
using a obj_simps by auto
finally show ?thesis by fast
qed
thus ?thesis
using REP_leg0.antipar REP_leg0.ide_right
apply (intro in_hhomI) by auto
qed
qed
moreover have ide_f: "ide ?f"
using REP_leg0.antipar f_in_hhom by fastforce
moreover have "Span.isomorphic (SPN ?f) g"
proof -
have SPN_f_eq: "SPN ?f = \<lparr>Chn = \<lbrakk>\<lbrakk>spn ?f\<rbrakk>\<rbrakk>,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>, Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>\<rparr>\<rparr>"
using calculation(1) SPN_def [of ?f] REP_leg0.antipar by auto
text \<open>
We need an invertible arrow of spans from \<open>SPN f\<close> to \<open>g\<close>.
There exists a tabulation \<open>(REP g.leg0, \<rho>, REP g.leg1)\<close> of \<open>f\<close>.
There is also a tabulation \<open>(tab\<^sub>0 f, \<rho>', tab\<^sub>1 f) of f\<close>.
As these are tabulations of the same arrow, they are equivalent.
This yields an equivalence map which is an arrow of spans from
\<open>(tab\<^sub>0 f, tab\<^sub>1 f)\<close> to \<open>(REP g.leg0, \<rho>, REP g.leg1)\<close>.
Its isomorphism class is an invertible arrow of spans in maps
from \<open>(CLS (tab\<^sub>0 f), CLS (tab\<^sub>1 f))\<close> to \<open>(g.leg0, g.leg1)\<close>.
\<close>
interpret f: identity_in_bicategory_of_spans V H \<a> \<i> src trg ?f
using ide_f apply unfold_locales by auto
interpret f: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
?f f.tab \<open>tab\<^sub>0 ?f\<close> \<open>tab\<^sub>1 ?f\<close> ?f f.tab \<open>tab\<^sub>0 ?f\<close> \<open>tab\<^sub>1 ?f\<close> ?f
using f.is_arrow_of_tabulations_in_maps by simp
interpret g: span_of_maps V H \<a> \<i> src trg \<open>Maps.REP g.leg0\<close> \<open>Maps.REP g.leg1\<close>
using Span.arr_char
by (unfold_locales, blast+)
have 2: "src (Maps.REP g.leg0) = src (Maps.REP g.leg1)"
using 0 1 by fastforce
hence "\<exists>\<rho>. tabulation (\<cdot>) (\<star>) \<a> \<i> src trg ?f \<rho> (Maps.REP g.leg0) (Maps.REP g.leg1)"
using BS2' [of "Maps.REP g.leg0" "Maps.REP g.leg1" ?f] isomorphic_reflexive
Span.arr_char
by auto
hence "tabulation V H \<a> \<i> src trg ?f
(REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f) (Maps.REP g.leg0) (Maps.REP g.leg1)"
using 2 REP_leg0.canonical_tabulation [of "Maps.REP g.leg1"] by auto
hence 3: "\<exists>w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'.
equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>w' : src (Maps.REP g.leg0) \<rightarrow> src (tab\<^sub>0 ?f)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
f.tab = (?f \<star> \<theta>) \<cdot> \<a>[?f, Maps.REP g.leg0, w] \<cdot>
(REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : tab\<^sub>0 ?f \<star> w' \<Rightarrow> Maps.REP g.leg0\<guillemotright> \<and>
\<guillemotleft>\<nu>' : Maps.REP g.leg1 \<Rightarrow> tab\<^sub>1 ?f \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f =
(?f \<star> \<theta>') \<cdot> \<a>[?f, tab\<^sub>0 ?f, w'] \<cdot> (f.tab \<star> w') \<cdot> \<nu>'"
using f.apex_unique_up_to_equivalence
[of "REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f"
"Maps.REP g.leg0" "Maps.REP g.leg1"]
by simp
obtain w w' \<phi> \<psi> \<theta> \<nu> \<theta>' \<nu>'
where 4: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg w' w \<psi> \<phi> \<and>
\<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>w' : src (Maps.REP g.leg0) \<rightarrow> src (tab\<^sub>0 ?f)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu> \<and>
f.tab = (?f \<star> \<theta>) \<cdot> \<a>[?f, Maps.REP g.leg0, w] \<cdot>
(REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f \<star> w) \<cdot> \<nu> \<and>
\<guillemotleft>\<theta>' : tab\<^sub>0 ?f \<star> w' \<Rightarrow> Maps.REP g.leg0\<guillemotright> \<and>
\<guillemotleft>\<nu>' : Maps.REP g.leg1 \<Rightarrow> tab\<^sub>1 ?f \<star> w'\<guillemotright> \<and> iso \<nu>' \<and>
REP_leg0.trnr\<^sub>\<eta> (Maps.REP g.leg1) ?f =
(?f \<star> \<theta>') \<cdot> \<a>[?f, tab\<^sub>0 ?f, w'] \<cdot> (f.tab \<star> w') \<cdot> \<nu>'"
using 3 by meson
hence w\<theta>\<nu>: "equivalence_map w \<and> \<guillemotleft>w : src (tab\<^sub>0 ?f) \<rightarrow> src (Maps.REP g.leg0)\<guillemotright> \<and>
\<guillemotleft>\<theta> : Maps.REP g.leg0 \<star> w \<Rightarrow> tab\<^sub>0 ?f\<guillemotright> \<and>
\<guillemotleft>\<nu> : tab\<^sub>1 ?f \<Rightarrow> Maps.REP g.leg1 \<star> w\<guillemotright> \<and> iso \<nu>"
using equivalence_map_def quasi_inverses_def quasi_inverses_symmetric
by meson
let ?W = "\<lparr>Chn = \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>, Dom = Dom (SPN ?f), Cod = Dom g\<rparr>"
have W_in_hom: "Span.in_hom ?W (SPN ?f) g"
proof
have "arrow_of_spans Maps.comp ?W"
proof
interpret Dom_W: span_in_category Maps.comp \<open>Dom ?W\<close>
proof (unfold_locales, intro conjI)
show "Maps.arr (Leg0 (Dom ?W))"
apply (intro Maps.arrI)
apply auto
by (metis f.base_simps(2) f.satisfies_T0 f.u_in_hom src_hcomp)
show "Maps.arr (Leg1 (Dom ?W))"
using 1
apply (intro Maps.arrI)
apply auto
proof -
let ?f = "tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)"
assume 1: "\<guillemotleft>Maps.REP g.leg1 : Maps.Dom g.apex \<rightarrow> Maps.Cod g.leg1\<guillemotright>"
have "\<guillemotleft>?f : src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))
\<rightarrow> Maps.Cod g.leg1\<guillemotright> \<and>
is_left_adjoint ?f \<and> \<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk> = \<lbrakk>?f\<rbrakk>"
using 1 by simp
thus "\<exists>f. \<guillemotleft>f : src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))
\<rightarrow> Maps.Cod g.leg1\<guillemotright> \<and>
is_left_adjoint f \<and>
\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk> = \<lbrakk>f\<rbrakk>"
by blast
qed
show "Maps.dom (Leg0 (Dom ?W)) = Maps.dom (Leg1 (Dom ?W))"
proof -
have "Maps.dom (Leg0 (Dom ?W)) =
Maps.MkIde (src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))"
using Maps.dom_char
apply simp
by (metis (no_types, lifting) Maps.CLS_in_hom Maps.in_homE f.base_simps(2)
f.satisfies_T0 f.u_simps(3) hcomp_simps(1))
also have "... = Maps.dom (Leg1 (Dom ?W))"
using Maps.dom_char Maps.CLS_in_hom f.leg1_is_map f_in_hhom
apply simp
by (metis (no_types, lifting) Maps.in_homE Maps.REP_simps(3) f.base_simps(2)
f.leg1_is_map f.leg1_simps(3) f.leg1_simps(4) g.dom.leg_simps(3)
trg_hcomp)
finally show ?thesis by blast
qed
qed
show "Maps.span (Leg0 (Dom ?W)) (Leg1 (Dom ?W))"
using Dom_W.span_in_category_axioms Dom_W.is_span by blast
interpret Cod_W: span_in_category Maps.comp \<open>Cod ?W\<close>
apply unfold_locales by fastforce
show "Maps.span (Leg0 (Cod ?W)) (Leg1 (Cod ?W))"
by fastforce
show "Maps.in_hom (Chn ?W) Dom_W.apex Cod_W.apex"
proof
show 1: "Maps.arr (Chn ?W)"
using w\<theta>\<nu> Maps.CLS_in_hom [of w] equivalence_is_adjoint by auto
show "Maps.dom (Chn ?W) = Dom_W.apex"
proof -
have "Maps.dom (Chn ?W) = Maps.MkIde (src w)"
using 1 w\<theta>\<nu> Maps.dom_char by simp
also have "... = Dom_W.apex"
proof -
have "src w = src (tab\<^sub>0 ?f)"
using w\<theta>\<nu> by blast
thus ?thesis
using Dom_W.apex_def Maps.arr_char Maps.dom_char
apply simp
by (metis (no_types, lifting) f.base_simps(2) f.satisfies_T0
f.u_in_hom hcomp_simps(1))
qed
finally show ?thesis by fastforce
qed
show "Maps.cod (Chn ?W) = Cod_W.apex"
proof -
have "Maps.cod (Chn ?W) = Maps.MkIde (trg w)"
using 1 w\<theta>\<nu> Maps.cod_char by simp
also have "... = Cod_W.apex"
proof -
have "trg w = src (Maps.REP g.leg0)"
using w\<theta>\<nu> by blast
thus ?thesis
using Cod_W.apex_def Maps.arr_char Maps.cod_char
apply simp
using g.dom.apex_def Maps.dom_char Maps.REP_simps(2) g.dom.is_span
by presburger
qed
finally show ?thesis by fastforce
qed
qed
show "Cod_W.leg0 \<odot> Chn ?W = Dom_W.leg0"
proof -
have "Cod_W.leg0 \<odot> Chn ?W = g.leg0 \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using g.dom.leg_simps(1) Maps.CLS_REP [of g.leg0]
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg0 \<star> w\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (Maps.REP g.leg0)"
by fast
moreover have "is_left_adjoint w"
using w\<theta>\<nu> equivalence_is_adjoint by simp
moreover have "Maps.REP g.leg0 \<star> w \<cong> Maps.REP g.leg0 \<star> w"
using w\<theta>\<nu> isomorphic_reflexive Maps.REP_in_hhom
by (metis (no_types, lifting) REP_leg0.ide_left adjoint_pair_antipar(1)
calculation(2) ide_hcomp in_hhomE)
ultimately show ?thesis
using w\<theta>\<nu> Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
by blast
qed
also have "... = \<lbrakk>\<lbrakk>tab\<^sub>0 ?f\<rbrakk>\<rbrakk>"
proof -
have "iso \<theta>"
proof -
have "is_left_adjoint (Maps.REP g.leg0 \<star> w)"
using w\<theta>\<nu> equivalence_is_adjoint Maps.REP_in_hhom
by (simp add: g.leg0_is_map in_hhom_def left_adjoints_compose)
moreover have "is_left_adjoint (tab\<^sub>0 ?f)"
by simp
ultimately show ?thesis
using w\<theta>\<nu> BS3 by blast
qed
thus ?thesis
using w\<theta>\<nu> Maps.CLS_eqI equivalence_is_adjoint
by (meson isomorphic_def isomorphic_implies_hpar(1))
qed
finally show ?thesis by fastforce
qed
show "Cod_W.leg1 \<odot> Chn ?W = Dom_W.leg1"
proof -
have "Cod_W.leg1 \<odot> Chn ?W = g.leg1 \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
by simp
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using g.dom.leg_simps(3) Maps.CLS_REP by presburger
also have "... = \<lbrakk>\<lbrakk>Maps.REP g.leg1 \<star> w\<rbrakk>\<rbrakk>"
proof -
have "is_left_adjoint (Maps.REP g.leg1)"
by fast
moreover have "is_left_adjoint w"
using w\<theta>\<nu> equivalence_is_adjoint by simp
moreover have "Maps.REP g.leg1 \<star> w \<cong> Maps.REP g.leg1 \<star> w"
using w\<theta>\<nu> isomorphic_reflexive Maps.REP_in_hhom
by (metis (no_types, lifting) "2" calculation(2) g.dom.is_span
hcomp_ide_isomorphic Maps.ide_REP in_hhomE
right_adjoint_determines_left_up_to_iso)
ultimately show ?thesis
using w\<theta>\<nu> Maps.comp_CLS isomorphic_reflexive equivalence_is_adjoint
by blast
qed
also have "... = \<lbrakk>\<lbrakk>tab\<^sub>1 ?f\<rbrakk>\<rbrakk>"
proof -
have "ide (Maps.REP g.leg1 \<star> w)"
using 2 w\<theta>\<nu> equivalence_map_is_ide by auto
moreover have "Maps.REP g.leg1 \<star> w \<cong>
tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)"
using w\<theta>\<nu> equivalence_is_adjoint f.leg1_is_map
right_adjoint_determines_left_up_to_iso adjoint_pair_preserved_by_iso
by (meson adjoint_pair_antipar(2) ide_in_hom(2) ide_is_iso)
ultimately show ?thesis
using Maps.CLS_eqI by blast
qed
finally show ?thesis by fastforce
qed
qed
thus W: "Span.arr ?W"
using Span.arr_char by blast
interpret Dom_W:
span_in_category Maps.comp
\<open>\<lparr>Leg0 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))
(src (Maps.REP g.leg0)\<^sup>*)
(iso_class (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*))),
Leg1 = Maps.MkArr (src (tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))
(Maps.Cod g.leg1)
(iso_class (tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)))\<rparr>\<close>
using W Span.arr_char
by (simp add: arrow_of_spans_def)
interpret Cod_W: span_in_category Maps.comp \<open>Cod ?W\<close>
using W Span.arr_char
by (simp add: arrow_of_spans_def)
show "Span.dom ?W = SPN ?f"
proof -
have "Span.dom ?W =
\<lparr>Chn = Dom_W.apex,
Dom = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>,
Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>\<rparr>,
Cod = \<lparr>Leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>,
Leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 (Maps.REP g.leg1 \<star> (Maps.REP g.leg0)\<^sup>*)\<rbrakk>\<rbrakk>\<rparr>\<rparr>"
using 0 W Span.dom_char by simp
also have "... = SPN ?f"
using SPN_def Dom_W.apex_def Maps.dom_char Dom_W.is_span iso_class_eqI
spn_ide
apply simp
using ide_f by blast
finally show ?thesis by blast
qed
show "Span.cod ?W = g"
using 0 W Span.cod_char Cod_W.apex_def by simp
qed
moreover have "Span.iso ?W"
proof -
have "Maps.iso \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
proof -
have "Maps.arr \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk> \<and> w \<in> Maps.Map \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk> \<and> equivalence_map w"
proof (intro conjI)
show 1: "Maps.arr \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using w\<theta>\<nu> Maps.CLS_in_hom equivalence_is_adjoint by blast
show "equivalence_map w"
using w\<theta>\<nu> by blast
show "w \<in> Maps.Map \<lbrakk>\<lbrakk>w\<rbrakk>\<rbrakk>"
using 1 w\<theta>\<nu> equivalence_is_adjoint Maps.arr_char
by (simp add: equivalence_map_is_ide ide_in_iso_class)
qed
thus ?thesis
using Maps.iso_char' by blast
qed
thus ?thesis
using w\<theta>\<nu> W_in_hom Span.iso_char by auto
qed
ultimately show ?thesis
using Span.isomorphic_def by blast
qed
ultimately show ?thesis by blast
qed
qed
show "\<And>r s \<tau>. \<lbrakk>ide r; ide s; src r = src s; trg r = trg s; Span.in_hom \<tau> (SPN r) (SPN s)\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright> \<and> SPN \<mu> = \<tau>"
proof -
fix r s \<tau>
assume r: "ide r" and s: "ide s"
assume src_eq: "src r = src s" and trg_eq: "trg r = trg s"
assume \<tau>: "Span.in_hom \<tau> (SPN r) (SPN s)"
interpret \<tau>: arrow_of_spans Maps.comp \<tau>
using \<tau> Span.arr_char by auto
interpret r: identity_in_bicategory_of_spans V H \<a> \<i> src trg r
using r by unfold_locales auto
interpret s: identity_in_bicategory_of_spans V H \<a> \<i> src trg s
using s by unfold_locales auto
interpret s: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> s
using s.is_arrow_of_tabulations_in_maps by simp
have \<tau>_dom_leg0_eq: "\<tau>.dom.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 r\<rbrakk>\<rbrakk>"
using \<tau> Span.dom_char SPN_def [of r] by auto
have \<tau>_dom_leg1_eq: "\<tau>.dom.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 r\<rbrakk>\<rbrakk>"
using \<tau> Span.dom_char SPN_def [of r] by auto
have \<tau>_cod_leg0_eq: "\<tau>.cod.leg0 = \<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>"
using \<tau> Span.cod_char SPN_def [of s] by auto
have \<tau>_cod_leg1_eq: "\<tau>.cod.leg1 = \<lbrakk>\<lbrakk>tab\<^sub>1 s\<rbrakk>\<rbrakk>"
using \<tau> Span.cod_char SPN_def [of s] by auto
have 1: "tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<cong> tab\<^sub>0 r"
proof -
have "tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<cong> Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine"
proof -
have "Maps.REP \<tau>.cod.leg0 \<cong> tab\<^sub>0 s"
using \<tau>_cod_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.satisfies_T0
by presburger
moreover have "src (tab\<^sub>0 s) = trg (Maps.REP \<tau>.chine)"
using Maps.seq_char [of "\<lbrakk>\<lbrakk>tab\<^sub>0 s\<rbrakk>\<rbrakk>" \<tau>.chine] \<tau>.cod.leg_simps(1)
\<tau>.leg0_commutes \<tau>_cod_leg0_eq
by auto
ultimately show ?thesis
using hcomp_isomorphic_ide [of "tab\<^sub>0 s" "Maps.REP \<tau>.cod.leg0" "Maps.REP \<tau>.chine"]
isomorphic_symmetric
by fastforce
qed
also have "... \<cong> Maps.REP \<tau>.dom.leg0"
proof -
have "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg0\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg0\<rbrakk>\<rbrakk>"
using \<tau>.leg0_commutes Maps.CLS_REP \<tau>.chine_simps(1)
\<tau>.cod.leg_simps(1) \<tau>.dom.leg_simps(1)
by presburger
hence "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg0\<rbrakk>\<rbrakk>"
using Maps.comp_CLS [of "Maps.REP \<tau>.cod.leg0" "Maps.REP \<tau>.chine"
"Maps.REP \<tau>.cod.leg0 \<star> Maps.REP \<tau>.chine"]
isomorphic_reflexive
by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2) Maps.REP_simps(2-3)
\<tau>.chine_in_hom \<tau>.cod.leg_in_hom(1) \<tau>.dom.leg_simps(1) \<tau>.leg0_commutes
ide_hcomp Maps.ide_REP)
thus ?thesis
using Maps.CLS_eqI Maps.seq_char Maps.ide_REP
by (meson calculation isomorphic_implies_ide(2))
qed
also have "... \<cong> tab\<^sub>0 r"
using \<tau>_dom_leg0_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.satisfies_T0
by presburger
finally show ?thesis by blast
qed
obtain \<theta> where \<theta>: "\<guillemotleft>\<theta> : tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<Rightarrow> tab\<^sub>0 r\<guillemotright> \<and> iso \<theta>"
using 1 by blast
have 2: "tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<cong> tab\<^sub>1 r"
proof -
have "tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<cong> Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine"
proof -
have "Maps.REP \<tau>.cod.leg1 \<cong> tab\<^sub>1 s"
using \<tau>_cod_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS s.leg1_is_map
by presburger
moreover have "src (Maps.REP \<tau>.cod.leg1) = trg (Maps.REP \<tau>.chine)"
using Maps.seq_char by auto
ultimately show ?thesis
using hcomp_isomorphic_ide [of "Maps.REP \<tau>.cod.leg1" "tab\<^sub>1 s" "Maps.REP \<tau>.chine"]
isomorphic_symmetric
by fastforce
qed
also have "... \<cong> Maps.REP \<tau>.dom.leg1"
proof -
have "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg1\<rbrakk>\<rbrakk> \<odot> \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg1\<rbrakk>\<rbrakk>"
using \<tau>.leg1_commutes Maps.CLS_REP \<tau>.chine_simps(1)
\<tau>.cod.leg_simps(3) \<tau>.dom.leg_simps(3)
by presburger
hence "\<lbrakk>\<lbrakk>Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine\<rbrakk>\<rbrakk> = \<lbrakk>\<lbrakk>Maps.REP \<tau>.dom.leg1\<rbrakk>\<rbrakk>"
using Maps.comp_CLS [of "Maps.REP \<tau>.cod.leg1" "Maps.REP \<tau>.chine"
"Maps.REP \<tau>.cod.leg1 \<star> Maps.REP \<tau>.chine"]
isomorphic_reflexive
by (metis (no_types, lifting) Maps.seq_char Maps.REP_in_hhom(2)
Maps.REP_simps(2) Maps.REP_simps(3) \<tau>.chine_in_hom \<tau>.cod.leg_in_hom(2)
\<tau>.dom.leg_simps(3) \<tau>.leg1_commutes ide_hcomp Maps.ide_REP)
thus ?thesis
using Maps.CLS_eqI Maps.seq_char Maps.ide_REP
by (meson calculation isomorphic_implies_ide(2))
qed
also have "... \<cong> tab\<^sub>1 r"
using \<tau>_dom_leg1_eq Maps.CLS_REP Maps.CLS_eqI Maps.REP_CLS r.leg1_is_map
by presburger
finally show ?thesis by blast
qed
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 s \<star> Maps.REP \<tau>.chine\<guillemotright> \<and> iso \<nu>"
using 2 isomorphic_symmetric by blast
define \<Delta>
where "\<Delta> \<equiv> (s \<star> \<theta>) \<cdot> \<a>[s, tab\<^sub>0 s, Maps.REP \<tau>.chine] \<cdot> (s.tab \<star> Maps.REP \<tau>.chine) \<cdot> \<nu>"
have \<Delta>: "\<guillemotleft>\<Delta> : tab\<^sub>1 r \<Rightarrow> s \<star> tab\<^sub>0 r\<guillemotright>"
proof (unfold \<Delta>_def, intro comp_in_homI)
show "\<guillemotleft>\<nu> : tab\<^sub>1 r \<Rightarrow> tab\<^sub>1 s \<star> Maps.REP \<tau>.chine\<guillemotright>"
using \<nu> by simp
show 3: "\<guillemotleft>s.tab \<star> Maps.REP \<tau>.chine :
tab\<^sub>1 s \<star> Maps.REP \<tau>.chine \<Rightarrow> (s \<star> tab\<^sub>0 s) \<star> Maps.REP \<tau>.chine\<guillemotright>"
apply (intro hcomp_in_vhom)
apply auto
using "1" by fastforce
show "\<guillemotleft>\<a>[s, tab\<^sub>0 s, Maps.REP \<tau>.chine] :
(s \<star> tab\<^sub>0 s) \<star> Maps.REP \<tau>.chine \<Rightarrow> s \<star> tab\<^sub>0 s \<star> Maps.REP \<tau>.chine\<guillemotright>"
using s assoc_in_hom [of s "tab\<^sub>0 s" "Maps.REP \<tau>.chine"]
by (metis (no_types, lifting) Maps.ide_REP 3 \<tau>.chine_simps(1) hcomp_in_vhomE
ideD(2) ideD(3) s.ide_u s.tab_simps(2) s.u_simps(3))
show "\<guillemotleft>s \<star> \<theta> : s \<star> tab\<^sub>0 s \<star> Maps.REP \<tau>.chine \<Rightarrow> s \<star> tab\<^sub>0 r\<guillemotright>"
using 1 s \<theta> isomorphic_implies_hpar(4) src_eq by auto
qed
define \<mu> where "\<mu> \<equiv> r.T0.trnr\<^sub>\<epsilon> s \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
have \<mu>: "\<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright>"
proof (unfold \<mu>_def, intro comp_in_homI)
show "\<guillemotleft>inv (r.T0.trnr\<^sub>\<epsilon> r r.tab) : r \<Rightarrow> tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*\<guillemotright>"
using r.yields_isomorphic_representation by fastforce
show "\<guillemotleft>r.T0.trnr\<^sub>\<epsilon> s \<Delta> : tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>* \<Rightarrow> s\<guillemotright>"
using s \<Delta> src_eq r.T0.adjoint_transpose_right(2) [of s "tab\<^sub>1 r"] by auto
qed
interpret \<mu>: arrow_in_bicategory_of_spans V H \<a> \<i> src trg r s \<mu>
using \<mu> by unfold_locales auto
interpret \<mu>: arrow_of_tabulations_in_maps V H \<a> \<i> src trg
r r.tab \<open>tab\<^sub>0 r\<close> \<open>tab\<^sub>1 r\<close> s s.tab \<open>tab\<^sub>0 s\<close> \<open>tab\<^sub>1 s\<close> \<mu>
using \<mu>.is_arrow_of_tabulations_in_maps by simp
have \<Delta>_eq: "\<Delta> = \<mu>.\<Delta>"
proof -
have "r.T0.trnr\<^sub>\<epsilon> s \<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab) =
r.T0.trnr\<^sub>\<epsilon> s \<mu>.\<Delta> \<cdot> inv (r.T0.trnr\<^sub>\<epsilon> r r.tab)"
using \<mu> \<mu>_def \<mu>.\<mu>_in_terms_of_\<Delta> by auto
hence "r.T0.trnr\<^sub>\<epsilon> s \<Delta> = r.T0.trnr\<^sub>\<epsilon> s \<mu>.\<Delta>"
using r s \<Delta> r.T0.adjoint_transpose_right(2) r.yields_isomorphic_representation
iso_inv_iso iso_is_retraction retraction_is_epi epiE
by (metis \<mu>.in_hom \<mu>_def arrI)
thus ?thesis
using \<Delta> \<mu>.\<Delta>_in_hom(2) src_eq r.T0.adjoint_transpose_right(6)
bij_betw_imp_inj_on
[of "r.T0.trnr\<^sub>\<epsilon> s" "hom (tab\<^sub>1 r) (s \<star> tab\<^sub>0 r)" "hom (tab\<^sub>1 r \<star> (tab\<^sub>0 r)\<^sup>*) s"]
inj_on_def [of "r.T0.trnr\<^sub>\<epsilon> s" "hom (tab\<^sub>1 r) (s \<star> tab\<^sub>0 r)"]
by simp
qed
have "\<mu>.is_induced_map (Maps.REP \<tau>.chine)"
using \<theta> \<nu> \<Delta>_eq \<Delta>_def \<mu>.is_induced_map_iff \<tau>.chine_simps(1) Maps.ide_REP by blast
hence 3: "Maps.REP \<tau>.chine \<cong> \<mu>.chine"
using \<mu>.chine_is_induced_map \<mu>.induced_map_unique by simp
have "SPN \<mu> = \<tau>"
proof (intro Span.arr_eqI)
show "Span.par (SPN \<mu>) \<tau>"
using \<mu> \<tau> SPN_in_hom
by (metis (no_types, lifting) SPN.preserves_cod SPN.preserves_dom Span.in_homE
in_homE)
show "Chn (SPN \<mu>) = \<tau>.chine"
proof -
have "Chn (SPN \<mu>) = \<lbrakk>\<lbrakk>spn \<mu>\<rbrakk>\<rbrakk>"
using \<mu> SPN_def spn_def by auto
also have "... = \<lbrakk>\<lbrakk>\<mu>.chine\<rbrakk>\<rbrakk>"
using \<mu> spn_def by fastforce
also have "... = \<lbrakk>\<lbrakk>Maps.REP \<tau>.chine\<rbrakk>\<rbrakk>"
using 3 isomorphic_symmetric Maps.CLS_eqI iso_class_eqI isomorphic_implies_hpar(3)
isomorphic_implies_hpar(4)
by auto
also have "... = \<tau>.chine"
using Maps.CLS_REP \<tau>.chine_simps(1) by blast
finally show ?thesis by blast
qed
qed
thus "\<exists>\<mu>. \<guillemotleft>\<mu> : r \<Rightarrow> s\<guillemotright> \<and> SPN \<mu> = \<tau>"
using \<mu> by auto
qed
qed
theorem SPN_is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg SPN \<Phi>"
..
text \<open>
We have completed the proof of the second half of the main result (CKS Theorem 4):
\<open>B\<close> is biequivalent (via \<open>SPN\<close>) to \<open>Span(Maps(B))\<close>.
\<close>
corollary
shows "equivalent_bicategories V H \<a> \<i> src trg
Span.vcomp Span.hcomp Span.assoc Span.unit Span.src Span.trg"
using SPN_is_equivalence_pseudofunctor equivalent_bicategories_def by blast
end
end
diff --git a/thys/Bicategory/InternalAdjunction.thy b/thys/Bicategory/InternalAdjunction.thy
--- a/thys/Bicategory/InternalAdjunction.thy
+++ b/thys/Bicategory/InternalAdjunction.thy
@@ -1,3344 +1,3344 @@
(* Title: InternalAdjunction
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Adjunctions in a Bicategory"
theory InternalAdjunction
imports CanonicalIsos Strictness
begin
text \<open>
An \emph{internal adjunction} in a bicategory is a four-tuple \<open>(f, g, \<eta>, \<epsilon>)\<close>,
where \<open>f\<close> and \<open>g\<close> are antiparallel 1-cells and \<open>\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>\<close> and
\<open>\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> src g\<guillemotright>\<close> are 2-cells, such that the familiar ``triangle''
(or ``zig-zag'') identities are satisfied. We state the triangle identities
in two equivalent forms, each of which is convenient in certain situations.
\<close>
locale adjunction_in_bicategory =
adjunction_data_in_bicategory +
assumes triangle_left: "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
and triangle_right: "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
begin
lemma triangle_left':
shows "\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] = f"
using triangle_left triangle_equiv_form by simp
lemma triangle_right':
shows "\<r>[g] \<cdot> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g] = g"
using triangle_right triangle_equiv_form by simp
end
text \<open>
Internal adjunctions have a number of properties, which we now develop,
that generalize those of ordinary adjunctions involving functors and
natural transformations.
\<close>
context bicategory
begin
lemma adjunction_unit_determines_counit:
assumes "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>"
and "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>'"
shows "\<epsilon> = \<epsilon>'"
proof -
interpret E: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms(1) by auto
interpret E': adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using assms(2) by auto
text \<open>
Note that since we want to prove the the result for an arbitrary bicategory,
not just in for a strict bicategory, the calculation is a little more involved
than one might expect from a treatment that suppresses canonical isomorphisms.
\<close>
have "\<epsilon> = \<epsilon> \<cdot> (f \<star> \<r>[g] \<cdot> (g \<star> \<epsilon>') \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> \<l>\<^sup>-\<^sup>1[g])"
using E'.triangle_right' comp_arr_dom by simp
also have "... = \<epsilon> \<cdot> (f \<star> \<r>[g]) \<cdot> (f \<star> g \<star> \<epsilon>') \<cdot> (f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar whisker_left by simp
also have "... = \<epsilon> \<cdot> ((f \<star> \<r>[g]) \<cdot> (f \<star> g \<star> \<epsilon>')) \<cdot> (f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using comp_assoc by simp
also have "... = \<epsilon> \<cdot> \<r>[f \<star> g] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, src g] \<cdot> (f \<star> g \<star> \<epsilon>')) \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "f \<star> \<r>[g] = \<r>[f \<star> g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, src g]"
using E.antipar(1) runit_hcomp(3) by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<epsilon> \<cdot> \<r>[f \<star> g]) \<cdot> ((f \<star> g) \<star> \<epsilon>') \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar E'.counit_in_hom assoc'_naturality [of f g \<epsilon>'] comp_assoc by simp
also have "... = \<r>[trg f] \<cdot> ((\<epsilon> \<star> trg f) \<cdot> ((f \<star> g) \<star> \<epsilon>')) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar E.counit_in_hom runit_naturality [of \<epsilon>] comp_assoc by simp
also have "... = (\<l>[src g] \<cdot> (src g \<star> \<epsilon>')) \<cdot> (\<epsilon> \<star> f \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "(\<epsilon> \<star> trg f) \<cdot> ((f \<star> g) \<star> \<epsilon>') = (src g \<star> \<epsilon>') \<cdot> (\<epsilon> \<star> f \<star> g)"
using E.antipar interchange E.counit_in_hom comp_arr_dom comp_cod_arr
by (metis E'.counit_simps(1-3) E.counit_simps(1-3))
thus ?thesis
using E.antipar comp_assoc unitor_coincidence by simp
qed
also have "... = \<epsilon>' \<cdot> \<l>[f \<star> g] \<cdot> (\<epsilon> \<star> f \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "\<l>[src g] \<cdot> (src g \<star> \<epsilon>') = \<epsilon>' \<cdot> \<l>[f \<star> g]"
using E.antipar lunit_naturality [of \<epsilon>'] by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[trg f, f, g] \<cdot> (\<epsilon> \<star> f \<star> g)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g]) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar lunit_hcomp comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot>
(f \<star> \<a>[g, f, g])) \<cdot> (f \<star> \<eta> \<star> g) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar assoc'_naturality [of \<epsilon> f g] comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g \<star> f, g] \<cdot> (f \<star> \<eta> \<star> g)) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
proof -
have "\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g] \<cdot> (f \<star> \<a>[g, f, g]) =
(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, g]"
using E.antipar iso_assoc' pentagon' comp_assoc
invert_side_of_triangle(2)
[of "\<a>\<^sup>-\<^sup>1[f \<star> g, f, g] \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> g]"
"(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, g]" "f \<star> \<a>\<^sup>-\<^sup>1[g, f, g]"]
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
((f \<star> \<eta>) \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar assoc'_naturality [of f \<eta> g] comp_assoc by simp
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<star> g) \<cdot> ((\<epsilon> \<star> f) \<star> g) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> g) \<cdot>
((f \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[f] \<star> g)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g]) = \<r>\<^sup>-\<^sup>1[f] \<star> g"
proof -
have "\<r>\<^sup>-\<^sup>1[f] \<star> g = inv (\<r>[f] \<star> g)"
using E.antipar by simp
also have "... = inv ((f \<star> \<l>[g]) \<cdot> \<a>[f, trg g, g])"
using E.antipar by (simp add: triangle)
also have "... = \<a>\<^sup>-\<^sup>1[f, trg g, g] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[g])"
using E.antipar inv_comp by simp
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = \<epsilon>' \<cdot> (\<l>[f] \<cdot> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[f] \<star> g)"
using E.antipar whisker_right by simp
also have "... = \<epsilon>'"
using E.triangle_left' comp_arr_dom by simp
finally show ?thesis by simp
qed
end
subsection "Adjoint Transpose"
context adjunction_in_bicategory
begin
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
text \<open>
Just as for an ordinary adjunction between categories, an adjunction in a bicategory
determines bijections between hom-sets. There are two versions of this relationship:
depending on whether the transposition is occurring on the left (\emph{i.e.}~``output'')
side or the right (\emph{i.e.}~``input'') side.
\<close>
definition trnl\<^sub>\<eta>
where "trnl\<^sub>\<eta> v \<mu> \<equiv> (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
definition trnl\<^sub>\<epsilon>
where "trnl\<^sub>\<epsilon> u \<nu> \<equiv> \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)"
lemma adjoint_transpose_left:
assumes "ide u" and "ide v" and "src f = trg v" and "src g = trg u"
shows "trnl\<^sub>\<eta> v \<in> hom (f \<star> v) u \<rightarrow> hom v (g \<star> u)"
and "trnl\<^sub>\<epsilon> u \<in> hom v (g \<star> u) \<rightarrow> hom (f \<star> v) u"
and "\<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>"
and "\<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>"
and "bij_betw (trnl\<^sub>\<eta> v) (hom (f \<star> v) u) (hom v (g \<star> u))"
and "bij_betw (trnl\<^sub>\<epsilon> u) (hom v (g \<star> u)) (hom (f \<star> v) u)"
proof -
show A: "trnl\<^sub>\<eta> v \<in> hom (f \<star> v) u \<rightarrow> hom v (g \<star> u)"
using assms antipar trnl\<^sub>\<eta>_def by fastforce
show B: "trnl\<^sub>\<epsilon> u \<in> hom v (g \<star> u) \<rightarrow> hom (f \<star> v) u"
using assms antipar trnl\<^sub>\<epsilon>_def by fastforce
show C: "\<And>\<mu>. \<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : f \<star> v \<Rightarrow> u\<guillemotright>"
have "trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) =
\<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v])"
using trnl\<^sub>\<eta>_def trnl\<^sub>\<epsilon>_def by simp
also have "... = \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> g \<star> \<mu>)) \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar whisker_left comp_assoc by auto
also have "... = \<l>[u] \<cdot> ((\<epsilon> \<star> u) \<cdot> ((f \<star> g) \<star> \<mu>)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar assoc'_naturality [of f g \<mu>] comp_assoc by fastforce
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot>
(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot>
(f \<star> \<eta> \<star> v) \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
proof -
have "(\<epsilon> \<star> u) \<cdot> ((f \<star> g) \<star> \<mu>) = (trg u \<star> \<mu>) \<cdot> (\<epsilon> \<star> f \<star> v)"
using assms \<mu> antipar comp_cod_arr comp_arr_dom
interchange [of "trg u" \<epsilon> \<mu> "f \<star> v"] interchange [of \<epsilon> "f \<star> g" u \<mu>]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> \<a>[trg f, f, v] \<cdot>
((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
proof -
have 1: "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
\<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v]"
proof -
have "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
(\<epsilon> \<star> f \<star> v) \<cdot>
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v] \<cdot>
(f \<star> \<eta> \<star> v)"
proof -
have "(\<epsilon> \<star> f \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) \<cdot> (f \<star> \<eta> \<star> v) =
(\<epsilon> \<star> f \<star> v) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v])) \<cdot> (f \<star> \<eta> \<star> v)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f \<star> v) \<cdot>
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v] \<cdot>
(f \<star> \<eta> \<star> v)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, f \<star> v] \<cdot> (f \<star> \<a>[g, f, v]) =
\<a>[f \<star> g, f, v] \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, g \<star> f, v]"
using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0
canI_associator_hcomp(1-3)
by simp
thus ?thesis
using comp_assoc by simp
qed
finally show ?thesis by blast
qed
also have "... = ((\<epsilon> \<star> f \<star> v) \<cdot> \<a>[f \<star> g, f, v]) \<cdot>
(\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar assoc'_naturality [of f \<eta> v] comp_assoc by simp
also have "... = (\<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<star> v)) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar assoc_naturality [of \<epsilon> f v] by simp
also have "... = \<a>[trg f, f, v] \<cdot>
(((\<epsilon> \<star> f) \<star> v) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<star> v) \<cdot> ((f \<star> \<eta>) \<star> v)) \<cdot>
\<a>\<^sup>-\<^sup>1[f, trg v, v]"
using comp_assoc by simp
also have "... = \<a>[trg f, f, v] \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v]"
using assms \<mu> antipar whisker_right by simp
finally show ?thesis by simp
qed
show ?thesis
using 1 comp_assoc by metis
qed
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot>
\<a>[trg f, f, v] \<cdot> (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] \<star> v) \<cdot> \<a>\<^sup>-\<^sup>1[f, trg v, v] \<cdot> (f \<star> \<l>\<^sup>-\<^sup>1[v])"
using assms \<mu> antipar triangle_left by simp
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> can (\<^bold>\<langle>trg u\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>v\<^bold>\<rangle>) (\<^bold>\<langle>f\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>v\<^bold>\<rangle>)"
using assms \<mu> antipar canI_unitor_0 canI_associator_1
canI_associator_1(1-2) [of f v] whisker_can_right_0 whisker_can_left_0
by simp
also have "... = \<l>[u] \<cdot> (trg u \<star> \<mu>) \<cdot> \<l>\<^sup>-\<^sup>1[f \<star> v]"
unfolding can_def using assms antipar comp_arr_dom comp_cod_arr \<ll>_ide_simp
by simp
also have "... = (\<l>[u] \<cdot> \<l>\<^sup>-\<^sup>1[u]) \<cdot> \<mu>"
using assms \<mu> lunit'_naturality [of \<mu>] comp_assoc by auto
also have "... = \<mu>"
using assms \<mu> comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto
finally show "trnl\<^sub>\<epsilon> u (trnl\<^sub>\<eta> v \<mu>) = \<mu>" by simp
qed
show D: "\<And>\<nu>. \<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright> \<Longrightarrow> trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>"
proof -
fix \<nu>
assume \<nu>: "\<guillemotleft>\<nu> : v \<Rightarrow> g \<star> u\<guillemotright>"
have "trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) =
(g \<star> \<l>[u] \<cdot> (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
using trnl\<^sub>\<eta>_def trnl\<^sub>\<epsilon>_def by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> (g \<star> f \<star> \<nu>) \<cdot>
\<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar interchange [of g "g \<cdot> g \<cdot> g"] comp_assoc by auto
also have "... = ((g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot>
\<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u)) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v] =
\<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v] =
\<a>[g, f, g \<star> u] \<cdot> ((g \<star> f) \<star> \<nu>) \<cdot> (\<eta> \<star> v) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> f \<star> \<nu>) \<cdot> \<a>[g, f, v] = \<a>[g, f, g \<star> u] \<cdot> ((g \<star> f) \<star> \<nu>)"
using assms \<nu> antipar assoc_naturality [of g f \<nu>] by auto
thus ?thesis
using assms comp_assoc by metis
qed
also have "... = \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "((g \<star> f) \<star> \<nu>) \<cdot> (\<eta> \<star> v) = (\<eta> \<star> g \<star> u) \<cdot> (trg v \<star> \<nu>)"
using assms \<nu> antipar comp_arr_dom comp_cod_arr
interchange [of "g \<star> f" \<eta> \<nu> v] interchange [of \<eta> "trg v" "g \<star> u" \<nu>]
by auto
thus ?thesis
using comp_assoc by metis
qed
finally show ?thesis by blast
qed
thus ?thesis using comp_assoc by simp
qed
also have "... = \<l>[g \<star> u] \<cdot> (trg v \<star> \<nu>) \<cdot> \<l>\<^sup>-\<^sup>1[v]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
\<l>[g \<star> u]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
(g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot>
((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u) \<cdot>
\<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> (\<eta> \<star> g \<star> u) =
(g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot>
((\<eta> \<star> g \<star> u) \<cdot> \<a>[trg v, g, u]) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar comp_arr_dom comp_assoc comp_assoc_assoc'(1) by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot>
(\<a>[g \<star> f, g, u] \<cdot> ((\<eta> \<star> g) \<star> u)) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar assoc_naturality [of \<eta> g u] by simp
also have "... = (g \<star> \<l>[u]) \<cdot> (g \<star> \<epsilon> \<star> u) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> ((\<a>[g, trg u, u] \<cdot> \<a>\<^sup>-\<^sup>1[g, trg u, u]) \<cdot> (g \<star> \<epsilon> \<star> u)) \<cdot>
((g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "(\<a>[g, trg u, u] \<cdot> \<a>\<^sup>-\<^sup>1[g, trg u, u]) \<cdot> (g \<star> \<epsilon> \<star> u) = g \<star> \<epsilon> \<star> u"
using assms antipar comp_cod_arr comp_assoc_assoc'(1) by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<a>\<^sup>-\<^sup>1[g, trg u, u] \<cdot> (g \<star> \<epsilon> \<star> u)) \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u] \<cdot>
((\<eta> \<star> g) \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (((g \<star> \<epsilon>) \<star> u) \<cdot> (\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot>
(g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u)) \<cdot> \<a>\<^sup>-\<^sup>1[trg v, g, u]"
using assms antipar assoc'_naturality [of g \<epsilon> u] comp_assoc by simp
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot>
((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u) \<cdot>
\<a>\<^sup>-\<^sup>1[trg v, g, u]"
proof -
have "\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u] =
\<a>[g, f, g] \<star> u"
using assms antipar canI_associator_0 whisker_can_left_0 whisker_can_right_0
canI_associator_hcomp
by simp
hence "((g \<star> \<epsilon>) \<star> u) \<cdot>
(\<a>\<^sup>-\<^sup>1[g, f \<star> g, u] \<cdot> (g \<star> \<a>\<^sup>-\<^sup>1[f, g, u]) \<cdot> \<a>[g, f, g \<star> u] \<cdot> \<a>[g \<star> f, g, u]) \<cdot>
((\<eta> \<star> g) \<star> u) =
(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<star> u"
using assms antipar whisker_right by simp
thus ?thesis by simp
qed
finally show ?thesis by blast
qed
also have "... = (g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u]"
using assms antipar triangle_right by simp
also have "... = can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>) (\<^bold>\<langle>trg g\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>)"
proof -
have "(g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u] =
((g \<star> \<l>[u]) \<cdot> \<a>[g, trg u, u] \<cdot> (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g] \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[trg g, g, u])"
using comp_assoc by simp
moreover have "... = can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>) (\<^bold>\<langle>trg g\<^bold>\<rangle>\<^sub>0 \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>u\<^bold>\<rangle>)"
using assms antipar canI_unitor_0 canI_associator_1 [of g u] inv_can
whisker_can_left_0 whisker_can_right_0
by simp
ultimately show ?thesis by simp
qed
also have "... = \<l>[g \<star> u]"
unfolding can_def using assms comp_arr_dom comp_cod_arr \<ll>_ide_simp by simp
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = (\<l>[g \<star> u] \<cdot> \<l>\<^sup>-\<^sup>1[g \<star> u]) \<cdot> \<nu>"
using assms \<nu> lunit'_naturality comp_assoc by auto
also have "... = \<nu>"
using assms \<nu> comp_cod_arr iso_lunit comp_arr_inv inv_is_inverse by auto
finally show "trnl\<^sub>\<eta> v (trnl\<^sub>\<epsilon> u \<nu>) = \<nu>" by simp
qed
show "bij_betw (trnl\<^sub>\<eta> v) (hom (f \<star> v) u) (hom v (g \<star> u))"
using A B C D by (intro bij_betwI) auto
show "bij_betw (trnl\<^sub>\<epsilon> u) (hom v (g \<star> u)) (hom (f \<star> v) u)"
using A B C D by (intro bij_betwI) auto
qed
lemma trnl\<^sub>\<epsilon>_comp:
assumes "ide u" and "seq \<mu> \<nu>" and "src f = trg \<mu>"
shows "trnl\<^sub>\<epsilon> u (\<mu> \<cdot> \<nu>) = trnl\<^sub>\<epsilon> u \<mu> \<cdot> (f \<star> \<nu>)"
using assms trnl\<^sub>\<epsilon>_def whisker_left [of f \<mu> \<nu>] comp_assoc by auto
definition trnr\<^sub>\<eta>
where "trnr\<^sub>\<eta> v \<mu> \<equiv> (\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
definition trnr\<^sub>\<epsilon>
where "trnr\<^sub>\<epsilon> u \<nu> \<equiv> \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g)"
lemma adjoint_transpose_right:
assumes "ide u" and "ide v" and "src v = trg g" and "src u = trg f"
shows "trnr\<^sub>\<eta> v \<in> hom (v \<star> g) u \<rightarrow> hom v (u \<star> f)"
and "trnr\<^sub>\<epsilon> u \<in> hom v (u \<star> f) \<rightarrow> hom (v \<star> g) u"
and "\<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>"
and "\<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright> \<Longrightarrow> trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>"
and "bij_betw (trnr\<^sub>\<eta> v) (hom (v \<star> g) u) (hom v (u \<star> f))"
and "bij_betw (trnr\<^sub>\<epsilon> u) (hom v (u \<star> f)) (hom (v \<star> g) u)"
proof -
show A: "trnr\<^sub>\<eta> v \<in> hom (v \<star> g) u \<rightarrow> hom v (u \<star> f)"
using assms antipar trnr\<^sub>\<eta>_def by fastforce
show B: "trnr\<^sub>\<epsilon> u \<in> hom v (u \<star> f) \<rightarrow> hom (v \<star> g) u"
using assms antipar trnr\<^sub>\<epsilon>_def by fastforce
show C: "\<And>\<mu>. \<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright> \<Longrightarrow> trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>"
proof -
fix \<mu>
assume \<mu>: "\<guillemotleft>\<mu> : v \<star> g \<Rightarrow> u\<guillemotright>"
have "trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) =
\<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> ((\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v] \<star> g)"
unfolding trnr\<^sub>\<epsilon>_def trnr\<^sub>\<eta>_def by simp
also have "... = \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> (\<a>[u, f, g] \<cdot> ((\<mu> \<star> f) \<star> g)) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using assms \<mu> antipar whisker_right comp_assoc by auto
also have "... = \<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> ((\<mu> \<star> f \<star> g) \<cdot> \<a>[v \<star> g, f, g]) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using assms \<mu> antipar assoc_naturality [of \<mu> f g] by auto
also have "... = \<r>[u] \<cdot> ((u \<star> \<epsilon>) \<cdot> (\<mu> \<star> f \<star> g)) \<cdot> \<a>[v \<star> g, f, g] \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using comp_assoc by auto
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot> ((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
proof -
have "(u \<star> \<epsilon>) \<cdot> (\<mu> \<star> f \<star> g) = (\<mu> \<star> src u) \<cdot> ((v \<star> g) \<star> \<epsilon>)"
using assms \<mu> antipar comp_arr_dom comp_cod_arr
interchange [of \<mu> "v \<star> g" "src u" \<epsilon>] interchange [of u \<mu> \<epsilon> "f \<star> g"]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
(((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)) \<cdot>
(\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using comp_assoc by simp
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]) \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
proof -
have "((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) =
\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> \<a>[v, src v, g]"
proof -
have "((v \<star> g) \<star> \<epsilon>) \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g) =
((\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> \<a>[v, g, src u]) \<cdot> ((v \<star> g) \<star> \<epsilon>)) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)"
proof -
have "arr v \<and> dom v = v \<and> cod v = v"
using assms(2) ide_char by blast
moreover have "arr g \<and> dom g = g \<and> cod g = g"
using ide_right ide_char by blast
ultimately show ?thesis
by (metis (no_types) antipar(2) assms(3-4) assoc_naturality
counit_simps(1,3,5) hcomp_reassoc(1) comp_assoc)
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (\<a>[v, g, src u] \<cdot> ((v \<star> g) \<star> \<epsilon>)) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> ((v \<star> \<eta>) \<star> g)"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> ((v \<star> g \<star> \<epsilon>) \<cdot> \<a>[v, g, f \<star> g]) \<cdot>
\<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
(\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g]) \<cdot> ((v \<star> \<eta>) \<star> g)"
proof -
have "\<a>[v, g, src u] \<cdot> ((v \<star> g) \<star> \<epsilon>) = (v \<star> g \<star> \<epsilon>) \<cdot> \<a>[v, g, f \<star> g]"
using assms antipar assoc_naturality [of v g \<epsilon>] by simp
moreover have "(\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g]) \<cdot> ((v \<star> \<eta>) \<star> g) = (v \<star> \<eta>) \<star> g"
using assms antipar comp_cod_arr comp_assoc_assoc'(2) by simp
ultimately show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> g \<star> \<epsilon>) \<cdot>
\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g \<star> f, g] \<cdot> \<a>[v, g \<star> f, g] \<cdot> ((v \<star> \<eta>) \<star> g)"
using comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> ((v \<star> g \<star> \<epsilon>) \<cdot>
(\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g \<star> f, g]) \<cdot> (v \<star> \<eta> \<star> g)) \<cdot> \<a>[v, src v, g]"
using assms antipar assoc_naturality [of v \<eta> g] comp_assoc by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot>
((v \<star> g \<star> \<epsilon>) \<cdot> (v \<star> \<a>[g, f, g]) \<cdot> (v \<star> \<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]"
proof -
have "\<a>[v, g, f \<star> g] \<cdot> \<a>[v \<star> g, f, g] \<cdot> (\<a>\<^sup>-\<^sup>1[v, g, f] \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[v, g \<star> f, g] =
v \<star> \<a>[g, f, g]"
using assms antipar canI_associator_0 canI_associator_hcomp
whisker_can_left_0 whisker_can_right_0
by simp
thus ?thesis
using assms antipar whisker_left by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot>
(v \<star> (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot>
\<a>[v, src v, g]"
using assms antipar whisker_left by simp
finally show ?thesis by simp
qed
thus ?thesis by auto
qed
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot>
\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot>
\<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g)"
using triangle_right comp_assoc by simp
also have "... = \<r>[u] \<cdot> (\<mu> \<star> src u) \<cdot> \<r>\<^sup>-\<^sup>1[v \<star> g]"
proof -
have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) = \<r>\<^sup>-\<^sup>1[v \<star> g]"
proof -
have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) =
\<a>\<^sup>-\<^sup>1[v, g, trg f] \<cdot> can (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)"
using assms canI_unitor_0 canI_associator_1(2-3) whisker_can_left_0(1)
whisker_can_right_0
by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> can (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) (\<^bold>\<langle>v\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>g\<^bold>\<rangle>)"
using antipar by simp
(* TODO: There should be an alternate version of whisker_can_left for this. *)
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> (v \<star> can (\<^bold>\<langle>g\<^bold>\<rangle> \<^bold>\<star> \<^bold>\<langle>src g\<^bold>\<rangle>\<^sub>0) \<^bold>\<langle>g\<^bold>\<rangle>)"
using assms canI_unitor_0(2) whisker_can_left_0 by simp
also have "... = \<a>\<^sup>-\<^sup>1[v, g, src g] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g])"
using assms canI_unitor_0(2) by simp
also have "... = \<r>\<^sup>-\<^sup>1[v \<star> g]"
using assms runit_hcomp(2) by simp
finally have "\<a>\<^sup>-\<^sup>1[v, g, src u] \<cdot> (v \<star> \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]) \<cdot> \<a>[v, src v, g] \<cdot> (\<r>\<^sup>-\<^sup>1[v] \<star> g) =
\<r>\<^sup>-\<^sup>1[v \<star> g]"
by simp
thus ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = (\<r>[u] \<cdot> \<r>\<^sup>-\<^sup>1[u]) \<cdot> \<mu>"
using assms \<mu> runit'_naturality [of \<mu>] comp_assoc by auto
also have "... = \<mu>"
using \<mu> comp_cod_arr iso_runit inv_is_inverse comp_arr_inv by auto
finally show "trnr\<^sub>\<epsilon> u (trnr\<^sub>\<eta> v \<mu>) = \<mu>" by simp
qed
show D: "\<And>\<nu>. \<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright> \<Longrightarrow> trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>"
proof -
fix \<nu>
assume \<nu>: "\<guillemotleft>\<nu> : v \<Rightarrow> u \<star> f\<guillemotright>"
have "trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) =
(\<r>[u] \<cdot> (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
unfolding trnr\<^sub>\<eta>_def trnr\<^sub>\<epsilon>_def by simp
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
(((\<nu> \<star> g) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f]) \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar whisker_right comp_assoc by auto
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
(\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> (\<nu> \<star> g \<star> f)) \<cdot> (v \<star> \<eta>) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms \<nu> antipar assoc'_naturality [of \<nu> g f] by auto
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> ((\<nu> \<star> g \<star> f) \<cdot> (v \<star> \<eta>)) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = (\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot> (\<a>[u, f, g] \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u \<star> f, g, f] \<cdot> (((u \<star> f) \<star> \<eta>) \<cdot> (\<nu> \<star> src v)) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
proof -
have "(\<nu> \<star> g \<star> f) \<cdot> (v \<star> \<eta>) = ((u \<star> f) \<star> \<eta>) \<cdot> (\<nu> \<star> src v)"
using assms \<nu> antipar interchange [of "u \<star> f" \<nu> \<eta> "src v"]
interchange [of \<nu> v "g \<star> f" \<eta>] comp_arr_dom comp_cod_arr
by auto
thus ?thesis by simp
qed
also have "... = ((\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot>
((\<a>[u, f, g] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u \<star> f, g, f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> ((u \<star> \<epsilon>) \<star> f) \<cdot>
(\<a>\<^sup>-\<^sup>1[u, f \<star> g, f] \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> \<a>[u, f, g \<star> f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar canI_associator_hcomp canI_associator_0 whisker_can_left_0
whisker_can_right_0
by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> (((u \<star> \<epsilon>) \<star> f) \<cdot>
\<a>\<^sup>-\<^sup>1[u, f \<star> g, f]) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> (\<a>[u, f, g \<star> f]) \<cdot>
((u \<star> f) \<star> \<eta>)) \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = ((\<r>[u] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<epsilon> \<star> f)) \<cdot>
(u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> ((u \<star> f \<star> \<eta>) \<cdot> \<a>[u, f, src f])) \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar assoc'_naturality [of u \<epsilon> f] assoc_naturality [of u f \<eta>]
by auto
also have "... = (\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot>
((u \<star> \<epsilon> \<star> f) \<cdot> (u \<star> \<a>\<^sup>-\<^sup>1[f, g, f]) \<cdot> (u \<star> f \<star> \<eta>)) \<cdot> \<a>[u, f, src f] \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using comp_assoc by simp
also have "... = (\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot>
(u \<star> (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> \<a>[u, f, src f] \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar whisker_left by auto
also have "... = ((\<r>[u] \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f])) \<cdot>
(\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
using assms antipar triangle_left comp_assoc by simp
also have "... = \<r>[u \<star> f] \<cdot> (\<nu> \<star> src v) \<cdot> \<r>\<^sup>-\<^sup>1[v]"
proof -
have "(\<r>[u] \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f] \<cdot> (u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f] =
((u \<star> \<l>[f]) \<cdot> (\<a>[u, src u, f] \<cdot> \<a>\<^sup>-\<^sup>1[u, src u, f])) \<cdot>
(u \<star> \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]) \<cdot> \<a>[u, f, src f]"
using assms ide_left ide_right antipar triangle comp_assoc by metis
also have "... = (u \<star> \<r>[f]) \<cdot> \<a>[u, f, src f]"
using assms antipar canI_associator_1 canI_unitor_0 whisker_can_left_0
whisker_can_right_0 canI_associator_1
by simp
also have "... = \<r>[u \<star> f]"
using assms antipar runit_hcomp by simp
finally show ?thesis by simp
qed
also have "... = (\<r>[u \<star> f] \<cdot> \<r>\<^sup>-\<^sup>1[u \<star> f]) \<cdot> \<nu>"
using assms \<nu> runit'_naturality [of \<nu>] comp_assoc by auto
also have "... = \<nu>"
using assms \<nu> comp_cod_arr comp_arr_inv inv_is_inverse iso_runit by auto
finally show "trnr\<^sub>\<eta> v (trnr\<^sub>\<epsilon> u \<nu>) = \<nu>" by auto
qed
show "bij_betw (trnr\<^sub>\<eta> v) (hom (v \<star> g) u) (hom v (u \<star> f))"
using A B C D by (intro bij_betwI, auto)
show "bij_betw (trnr\<^sub>\<epsilon> u) (hom v (u \<star> f)) (hom (v \<star> g) u)"
using A B C D by (intro bij_betwI, auto)
qed
lemma trnr\<^sub>\<eta>_comp:
assumes "ide v" and "seq \<mu> \<nu>" and "src \<mu> = trg f"
shows "trnr\<^sub>\<eta> v (\<mu> \<cdot> \<nu>) = (\<mu> \<star> f) \<cdot> trnr\<^sub>\<eta> v \<nu>"
using assms trnr\<^sub>\<eta>_def whisker_right comp_assoc by auto
end
text \<open>
It is useful to have at hand the simpler versions of the preceding results that
hold in a normal bicategory and in a strict bicategory.
\<close>
locale adjunction_in_normal_bicategory =
normal_bicategory +
adjunction_in_bicategory
begin
lemma triangle_left:
shows "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = f"
using triangle_left strict_lunit strict_runit by simp
lemma triangle_right:
shows "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = g"
using triangle_right strict_lunit strict_runit by simp
lemma trnr\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<mu> \<in> hom (v \<star> g) u"
shows "trnr\<^sub>\<eta> v \<mu> = (\<mu> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[v, g, f] \<cdot> (v \<star> \<eta>)"
unfolding trnr\<^sub>\<eta>_def
using assms antipar strict_runit' comp_arr_ide [of "\<r>\<^sup>-\<^sup>1[v]" "v \<star> \<eta>"] hcomp_arr_obj
by auto
lemma trnr\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<nu> \<in> hom v (u \<star> f)"
shows "trnr\<^sub>\<epsilon> u \<nu> = (u \<star> \<epsilon>) \<cdot> \<a>[u, f, g] \<cdot> (\<nu> \<star> g)"
unfolding trnr\<^sub>\<epsilon>_def
using assms antipar strict_runit comp_ide_arr hcomp_arr_obj by auto
lemma trnl\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<mu> \<in> hom (f \<star> v) u"
shows "trnl\<^sub>\<eta> v \<mu> = (g \<star> \<mu>) \<cdot> \<a>[g, f, v] \<cdot> (\<eta> \<star> v)"
using assms trnl\<^sub>\<eta>_def antipar strict_lunit comp_arr_dom hcomp_obj_arr by auto
lemma trnl\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<nu> \<in> hom v (g \<star> u)"
shows "trnl\<^sub>\<epsilon> u \<nu> = (\<epsilon> \<star> u) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, u] \<cdot> (f \<star> \<nu>)"
using assms trnl\<^sub>\<epsilon>_def antipar strict_lunit comp_cod_arr hcomp_obj_arr by auto
end
locale adjunction_in_strict_bicategory =
strict_bicategory +
adjunction_in_normal_bicategory
begin
lemma triangle_left:
shows "(\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f"
using ide_left ide_right antipar triangle_left strict_assoc' comp_cod_arr
by (metis dom_eqI ideD(1) seqE)
lemma triangle_right:
shows "(g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g) = g"
using ide_left ide_right antipar triangle_right strict_assoc comp_cod_arr
by (metis ideD(1) ideD(2) seqE)
lemma trnr\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<mu> \<in> hom (v \<star> g) u"
shows "trnr\<^sub>\<eta> v \<mu> = (\<mu> \<star> f) \<cdot> (v \<star> \<eta>)"
using assms antipar trnr\<^sub>\<eta>_eq strict_assoc' comp_ide_arr [of "\<a>\<^sup>-\<^sup>1[v, g, f]" "v \<star> \<eta>"]
by force
lemma trnr\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src v = trg g" and "src u = trg f"
and "\<nu> \<in> hom v (u \<star> f)"
shows "trnr\<^sub>\<epsilon> u \<nu> = (u \<star> \<epsilon>) \<cdot> (\<nu> \<star> g)"
using assms antipar trnr\<^sub>\<epsilon>_eq strict_assoc comp_ide_arr [of "\<a>[u, f, g]" "\<nu> \<star> g"]
by force
lemma trnl\<^sub>\<eta>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<mu> \<in> hom (f \<star> v) u"
shows "trnl\<^sub>\<eta> v \<mu> = (g \<star> \<mu>) \<cdot> (\<eta> \<star> v)"
using assms antipar trnl\<^sub>\<eta>_eq strict_assoc comp_ide_arr [of "\<a>[g, f, v]" "\<eta> \<star> v"]
by force
lemma trnl\<^sub>\<epsilon>_eq:
assumes "ide u" and "ide v"
and "src f = trg v" and "src g = trg u"
and "\<nu> \<in> hom v (g \<star> u)"
shows "trnl\<^sub>\<epsilon> u \<nu> = (\<epsilon> \<star> u) \<cdot> (f \<star> \<nu>)"
using assms antipar trnl\<^sub>\<epsilon>_eq strict_assoc' comp_ide_arr [of "\<a>\<^sup>-\<^sup>1[f, g, u]" "f \<star> \<nu>"]
by fastforce
end
subsection "Preservation Properties for Adjunctions"
text \<open>
Here we show that adjunctions are preserved under isomorphisms of the
left and right adjoints.
\<close>
context bicategory
begin
interpretation E: self_evaluation_map V H \<a> \<i> src trg ..
notation E.eval ("\<lbrace>_\<rbrace>")
definition adjoint_pair
where "adjoint_pair f g \<equiv> \<exists>\<eta> \<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
(* These would normally be called "maps", but that name is too heavily used already. *)
abbreviation is_left_adjoint
where "is_left_adjoint f \<equiv> \<exists>g. adjoint_pair f g"
abbreviation is_right_adjoint
where "is_right_adjoint g \<equiv> \<exists>f. adjoint_pair f g"
lemma adjoint_pair_antipar:
assumes "adjoint_pair f g"
shows "ide f" and "ide g" and "src f = trg g" and "src g = trg f"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
show "ide f" by simp
show "ide g" by simp
show "src f = trg g" using A.antipar by simp
show "src g = trg f" using A.antipar by simp
qed
lemma left_adjoint_is_ide:
assumes "is_left_adjoint f"
shows "ide f"
using assms adjoint_pair_antipar by auto
lemma right_adjoint_is_ide:
assumes "is_right_adjoint f"
shows "ide f"
using assms adjoint_pair_antipar by auto
lemma adjunction_preserved_by_iso_right:
assumes "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f g' ((\<phi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
proof
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide f" by simp
show "ide g'"
using assms(2) isomorphic_def by auto
show "\<guillemotleft>(\<phi> \<star> f) \<cdot> \<eta> : src f \<Rightarrow> g' \<star> f\<guillemotright>"
using assms A.antipar by fastforce
show "\<guillemotleft>\<epsilon> \<cdot> (f \<star> inv \<phi>) : f \<star> g' \<Rightarrow> src g'\<guillemotright>"
using assms A.antipar A.counit_in_hom by auto
show "(\<epsilon> \<cdot> (f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f] \<cdot> (f \<star> (\<phi> \<star> f) \<cdot> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "(\<epsilon> \<cdot> (f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f] \<cdot> (f \<star> (\<phi> \<star> f) \<cdot> \<eta>) =
(\<epsilon> \<star> f) \<cdot> (((f \<star> inv \<phi>) \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g', f]) \<cdot> (f \<star> \<phi> \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar whisker_right whisker_left comp_assoc by auto
also have "... = (\<epsilon> \<star> f) \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> inv \<phi> \<star> f)) \<cdot> (f \<star> \<phi> \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar assoc'_naturality [of f "inv \<phi>" f] by auto
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> ((f \<star> inv \<phi> \<star> f) \<cdot> (f \<star> \<phi> \<star> f)) \<cdot> (f \<star> \<eta>)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> g \<star> f) \<cdot> (f \<star> \<eta>)"
using assms A.antipar comp_inv_arr inv_is_inverse whisker_left
whisker_right [of f "inv \<phi>" \<phi>]
by auto
also have "... = (\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)"
using assms A.antipar comp_cod_arr by simp
also have "... = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using A.triangle_left by simp
finally show ?thesis by simp
qed
show "(g' \<star> \<epsilon> \<cdot> (f \<star> inv \<phi>)) \<cdot> \<a>[g', f, g'] \<cdot> ((\<phi> \<star> f) \<cdot> \<eta> \<star> g') = \<r>\<^sup>-\<^sup>1[g'] \<cdot> \<l>[g']"
proof -
have "(g' \<star> \<epsilon> \<cdot> (f \<star> inv \<phi>)) \<cdot> \<a>[g', f, g'] \<cdot> ((\<phi> \<star> f) \<cdot> \<eta> \<star> g') =
(g' \<star> \<epsilon>) \<cdot> ((g' \<star> f \<star> inv \<phi>) \<cdot> \<a>[g', f, g']) \<cdot> ((\<phi> \<star> f) \<star> g') \<cdot> (\<eta> \<star> g')"
using assms A.antipar whisker_left whisker_right comp_assoc by auto
also have "... = (g' \<star> \<epsilon>) \<cdot> (\<a>[g', f, g] \<cdot> ((g' \<star> f) \<star> inv \<phi>)) \<cdot> ((\<phi> \<star> f) \<star> g') \<cdot> (\<eta> \<star> g')"
using assms A.antipar assoc_naturality [of g' f "inv \<phi>"] by auto
also have "... = (g' \<star> \<epsilon>) \<cdot> \<a>[g', f, g] \<cdot> (((g' \<star> f) \<star> inv \<phi>) \<cdot> ((\<phi> \<star> f) \<star> g')) \<cdot> (\<eta> \<star> g')"
using comp_assoc by simp
also have "... = (g' \<star> \<epsilon>) \<cdot> (\<a>[g', f, g] \<cdot> ((\<phi> \<star> f) \<star> g)) \<cdot> ((g \<star> f) \<star> inv \<phi>) \<cdot> (\<eta> \<star> g')"
proof -
have "((g' \<star> f) \<star> inv \<phi>) \<cdot> ((\<phi> \<star> f) \<star> g') = (\<phi> \<star> f) \<star> inv \<phi>"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "g' \<star> f" "\<phi> \<star> f" "inv \<phi>" g']
by auto
also have "... = ((\<phi> \<star> f) \<star> g) \<cdot> ((g \<star> f) \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "\<phi> \<star> f" "g \<star> f" g "inv \<phi>"]
by auto
finally show ?thesis
using comp_assoc by simp
qed
also have "... = ((g' \<star> \<epsilon>) \<cdot> (\<phi> \<star> f \<star> g)) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
proof -
have "\<a>[g', f, g] \<cdot> ((\<phi> \<star> f) \<star> g) = (\<phi> \<star> f \<star> g) \<cdot> \<a>[g, f, g]"
using assms A.antipar assoc_naturality [of \<phi> f g] by auto
moreover have "((g \<star> f) \<star> inv \<phi>) \<cdot> (\<eta> \<star> g') = (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "g \<star> f" \<eta> "inv \<phi>" g'] interchange [of \<eta> "trg g" g "inv \<phi>"]
by auto
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = ((\<phi> \<star> src g) \<cdot> (g \<star> \<epsilon>)) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of g' \<phi> \<epsilon> "f \<star> g"] interchange [of \<phi> g "src g" \<epsilon>]
by (metis A.counit_simps(1) A.counit_simps(2) A.counit_simps(3) in_homE)
also have "... = (\<phi> \<star> src g) \<cdot> ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)) \<cdot> (trg g \<star> inv \<phi>)"
using comp_assoc by simp
also have "... = ((\<phi> \<star> src g) \<cdot> \<r>\<^sup>-\<^sup>1[g]) \<cdot> \<l>[g] \<cdot> (trg g \<star> inv \<phi>)"
using assms A.antipar A.triangle_right comp_cod_arr comp_assoc
by simp
also have "... = (\<r>\<^sup>-\<^sup>1[g'] \<cdot> \<phi>) \<cdot> inv \<phi> \<cdot> \<l>[g']"
using assms A.antipar runit'_naturality [of \<phi>] lunit_naturality [of "inv \<phi>"]
by auto
also have "... = \<r>\<^sup>-\<^sup>1[g'] \<cdot> (\<phi> \<cdot> inv \<phi>) \<cdot> \<l>[g']"
using comp_assoc by simp
also have "... = \<r>\<^sup>-\<^sup>1[g'] \<cdot> \<l>[g']"
using assms comp_cod_arr comp_arr_inv' by auto
finally show ?thesis by simp
qed
qed
lemma adjunction_preserved_by_iso_left:
assumes "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f' g ((g \<star> \<phi>) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g))"
proof
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by auto
show "ide g" by simp
show "ide f'"
using assms(2) isomorphic_def by auto
show "\<guillemotleft>(g \<star> \<phi>) \<cdot> \<eta> : src f' \<Rightarrow> g \<star> f'\<guillemotright>"
using assms A.antipar A.unit_in_hom by force
show "\<guillemotleft>\<epsilon> \<cdot> (inv \<phi> \<star> g) : f' \<star> g \<Rightarrow> src g\<guillemotright>"
using assms A.antipar by force
show "(g \<star> \<epsilon> \<cdot> (inv \<phi> \<star> g)) \<cdot> \<a>[g, f', g] \<cdot> ((g \<star> \<phi>) \<cdot> \<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
proof -
have "(g \<star> \<epsilon> \<cdot> (inv \<phi> \<star> g)) \<cdot> \<a>[g, f', g] \<cdot> ((g \<star> \<phi>) \<cdot> \<eta> \<star> g) =
(g \<star> \<epsilon>) \<cdot> ((g \<star> inv \<phi> \<star> g) \<cdot> \<a>[g, f', g]) \<cdot> ((g \<star> \<phi>) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar whisker_left whisker_right comp_assoc by auto
also have "... = (g \<star> \<epsilon>) \<cdot> (\<a>[g, f, g] \<cdot> ((g \<star> inv \<phi>) \<star> g)) \<cdot> ((g \<star> \<phi>) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar assoc_naturality [of g "inv \<phi>" g] by auto
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (((g \<star> inv \<phi>) \<star> g) \<cdot> ((g \<star> \<phi>) \<star> g)) \<cdot> (\<eta> \<star> g)"
using comp_assoc by simp
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> ((g \<star> f) \<star> g) \<cdot> (\<eta> \<star> g)"
using assms A.antipar comp_inv_arr inv_is_inverse whisker_right
whisker_left [of g "inv \<phi>" \<phi>]
by auto
also have "... = (g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g)"
using assms A.antipar comp_cod_arr by simp
also have "... = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using A.triangle_right by simp
finally show ?thesis by simp
qed
show "(\<epsilon> \<cdot> (inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f'] \<cdot> (f' \<star> (g \<star> \<phi>) \<cdot> \<eta>) = \<l>\<^sup>-\<^sup>1[f'] \<cdot> \<r>[f']"
proof -
have "(\<epsilon> \<cdot> (inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f'] \<cdot> (f' \<star> (g \<star> \<phi>) \<cdot> \<eta>) =
(\<epsilon> \<star> f') \<cdot> (((inv \<phi> \<star> g) \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f', g, f']) \<cdot> (f' \<star> g \<star> \<phi>) \<cdot> (f' \<star> \<eta>)"
using assms A.antipar whisker_right whisker_left comp_assoc
by auto
also have "... = (\<epsilon> \<star> f') \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (inv \<phi> \<star> g \<star> f')) \<cdot> (f' \<star> g \<star> \<phi>) \<cdot> (f' \<star> \<eta>)"
using assms A.antipar assoc'_naturality [of "inv \<phi>" g f'] by auto
also have "... = (\<epsilon> \<star> f') \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> ((inv \<phi> \<star> g \<star> f') \<cdot> (f' \<star> g \<star> \<phi>)) \<cdot> (f' \<star> \<eta>)"
using comp_assoc by simp
also have "... = (\<epsilon> \<star> f') \<cdot> (\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (f \<star> g \<star> \<phi>)) \<cdot> (inv \<phi> \<star> g \<star> f) \<cdot> (f' \<star> \<eta>)"
proof -
have "(inv \<phi> \<star> g \<star> f') \<cdot> (f' \<star> g \<star> \<phi>) = (f \<star> g \<star> \<phi>) \<cdot> (inv \<phi> \<star> g \<star> f)"
using assms(2-3) A.antipar comp_arr_dom comp_cod_arr
interchange [of "inv \<phi>" f' "g \<star> f'" "g \<star> \<phi>"]
interchange [of f "inv \<phi>" "g \<star> \<phi>" "g \<star> f"]
by auto
thus ?thesis
using comp_assoc by simp
qed
also have "... = ((\<epsilon> \<star> f') \<cdot> ((f \<star> g) \<star> \<phi>)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
proof -
have "\<a>\<^sup>-\<^sup>1[f, g, f'] \<cdot> (f \<star> g \<star> \<phi>) = ((f \<star> g) \<star> \<phi>) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f]"
using assms A.antipar assoc'_naturality [of f g \<phi>] by auto
moreover have "(inv \<phi> \<star> g \<star> f) \<cdot> (f' \<star> \<eta>) = (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of "inv \<phi>" f' "g \<star> f" \<eta>] interchange [of f "inv \<phi>" \<eta> "src f"]
by auto
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = ((trg f \<star> \<phi>) \<cdot> (\<epsilon> \<star> f)) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar comp_arr_dom comp_cod_arr
interchange [of \<epsilon> "f \<star> g" f' \<phi>] interchange [of "trg f" \<epsilon> \<phi> f]
by (metis A.counit_simps(1) A.counit_simps(2) A.counit_simps(3) in_homE)
also have "... = (trg f \<star> \<phi>) \<cdot> ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) \<cdot> (inv \<phi> \<star> src f)"
using comp_assoc by simp
also have "... = ((trg f \<star> \<phi>) \<cdot> \<l>\<^sup>-\<^sup>1[f]) \<cdot> \<r>[f] \<cdot> (inv \<phi> \<star> src f)"
using assms A.antipar A.triangle_left comp_cod_arr comp_assoc
by simp
also have "... = (\<l>\<^sup>-\<^sup>1[f'] \<cdot> \<phi>) \<cdot> inv \<phi> \<cdot> \<r>[f']"
using assms A.antipar lunit'_naturality runit_naturality [of "inv \<phi>"] by auto
also have "... = \<l>\<^sup>-\<^sup>1[f'] \<cdot> (\<phi> \<cdot> inv \<phi>) \<cdot> \<r>[f']"
using comp_assoc by simp
also have "... = \<l>\<^sup>-\<^sup>1[f'] \<cdot> \<r>[f']"
using assms comp_cod_arr comp_arr_inv inv_is_inverse by auto
finally show ?thesis by simp
qed
qed
lemma adjoint_pair_preserved_by_iso:
assumes "adjoint_pair f g"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
and "\<guillemotleft>\<psi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<psi>"
shows "adjoint_pair f' g'"
using assms adjoint_pair_def adjunction_preserved_by_iso_left
adjunction_preserved_by_iso_right
by metis
lemma left_adjoint_preserved_by_iso:
assumes "is_left_adjoint f"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "is_left_adjoint f'"
proof -
obtain g where g: "adjoint_pair f g"
using assms by auto
have "adjoint_pair f' g"
using assms g adjoint_pair_preserved_by_iso [of f g \<phi> f' g g]
adjoint_pair_antipar [of f g]
by auto
thus ?thesis by auto
qed
lemma right_adjoint_preserved_by_iso:
assumes "is_right_adjoint g"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "is_right_adjoint g'"
proof -
obtain f where f: "adjoint_pair f g"
using assms by auto
have "adjoint_pair f g'"
using assms f adjoint_pair_preserved_by_iso [of f g f f \<phi> g']
adjoint_pair_antipar [of f g]
by auto
thus ?thesis by auto
qed
lemma left_adjoint_preserved_by_iso':
assumes "is_left_adjoint f" and "f \<cong> f'"
shows "is_left_adjoint f'"
using assms isomorphic_def left_adjoint_preserved_by_iso by blast
lemma right_adjoint_preserved_by_iso':
assumes "is_right_adjoint g" and "g \<cong> g'"
shows "is_right_adjoint g'"
using assms isomorphic_def right_adjoint_preserved_by_iso by blast
lemma obj_self_adjunction:
assumes "obj a"
shows "adjunction_in_bicategory V H \<a> \<i> src trg a a \<l>\<^sup>-\<^sup>1[a] \<r>[a]"
proof
show 1: "ide a"
using assms by auto
show "\<guillemotleft>\<l>\<^sup>-\<^sup>1[a] : src a \<Rightarrow> a \<star> a\<guillemotright>"
using assms 1 by auto
show "\<guillemotleft>\<r>[a] : a \<star> a \<Rightarrow> src a\<guillemotright>"
using assms 1 by fastforce
show "(\<r>[a] \<star> a) \<cdot> \<a>\<^sup>-\<^sup>1[a, a, a] \<cdot> (a \<star> \<l>\<^sup>-\<^sup>1[a]) = \<l>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a]"
using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3
whisker_can_right_1 whisker_can_left_1 can_Ide_self obj_simps
by simp
show "(a \<star> \<r>[a]) \<cdot> \<a>[a, a, a] \<cdot> (\<l>\<^sup>-\<^sup>1[a] \<star> a) = \<r>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a]"
using assms 1 canI_unitor_1 canI_associator_1(2) canI_associator_3
whisker_can_right_1 whisker_can_left_1 can_Ide_self
by simp
qed
lemma obj_is_self_adjoint:
assumes "obj a"
shows "adjoint_pair a a" and "is_left_adjoint a" and "is_right_adjoint a"
using assms obj_self_adjunction adjoint_pair_def by auto
end
subsection "Pseudofunctors and Adjunctions"
context pseudofunctor
begin
lemma preserves_adjunction:
assumes "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof -
interpret adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms by auto
interpret A: adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>F f\<close> \<open>F g\<close> \<open>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)\<close>
\<open>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)\<close>
using adjunction_data_in_bicategory_axioms preserves_adjunction_data by auto
show "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof
show "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
D.lunit' (F f) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
proof -
have 1: "D.iso (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
using antipar C.VV.ide_char C.VV.arr_char D.iso_is_arr FF_def
by (intro D.isos_compose D.seqI, simp_all)
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))"
proof -
have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))) =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
proof -
have "D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F f]) =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
using antipar assoc_coherence by simp
moreover
have "D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F f]) =
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
proof -
have "D.seq (\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))) \<a>\<^sub>D[F f, F g, F f]"
using antipar by fastforce
thus ?thesis
using 1 antipar D.comp_assoc
D.inv_comp [of "\<a>\<^sub>D[F f, F g, F f]" "\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f))"]
by auto
qed
ultimately show ?thesis by simp
qed
moreover have 2: "D.iso (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f))"
using antipar D.isos_compose C.VV.ide_char C.VV.arr_char cmp_simps(4)
by simp
ultimately have "\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] =
D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)))"
using 1 2 antipar D.invert_side_of_triangle(2) D.inv_inv D.iso_inv_iso D.arr_inv
by metis
moreover have "D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
proof -
have "D.inv (F \<a>\<^sub>C[f, g, f] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) =
D.inv (\<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using antipar D.isos_compose C.VV.arr_char cmp_simps(4)
preserves_inv D.inv_comp C.VV.cod_char
by simp
also have "... = (D.inv (\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using antipar D.inv_comp C.VV.ide_char C.VV.arr_char cmp_simps(4)
by simp
also have "... = ((D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using antipar C.VV.ide_char C.VV.arr_char by simp
also have "... = (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using D.comp_assoc by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "... = ((D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D
((F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)))"
proof -
have "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f =
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)"
using ide_left ide_right antipar D.whisker_right unit_char(2)
by (metis A.counit_simps(1) A.ide_left D.comp_assoc)
moreover have "F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) =
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
using antipar unit_char(2) D.whisker_left by simp
ultimately show ?thesis by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
(((\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D \<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(((F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f)))) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)))"
using D.comp_assoc by simp
also have "... = (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
proof -
have "((F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f)))) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
proof -
have "(F f \<star>\<^sub>D \<Phi> (g, f)) \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) = F f \<star>\<^sub>D F (g \<star>\<^sub>C f)"
using antipar unit_char(2) D.comp_arr_inv D.inv_is_inverse
D.whisker_left [of "F f" "\<Phi> (g, f)" "D.inv (\<Phi> (g, f))"]
by simp
moreover have "D.seq (F f \<star>\<^sub>D F (g \<star>\<^sub>C f)) (F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))"
using antipar by fastforce
ultimately show ?thesis
using D.comp_cod_arr by auto
qed
moreover have "((\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D (D.inv (\<Phi> (f, g)) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) =
D.inv (\<Phi> (f \<star>\<^sub>C g, f))"
using antipar D.comp_arr_inv D.inv_is_inverse D.comp_cod_arr
D.whisker_right [of "F f" "\<Phi> (f, g)" "D.inv (\<Phi> (f, g))"]
by simp
ultimately show ?thesis by simp
qed
finally show ?thesis by simp
qed
also have "... = (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D
((\<Phi> (f \<star>\<^sub>C g, f) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, f))) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]) \<cdot>\<^sub>D
((\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C f))) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
(F f \<star>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) =
((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F \<epsilon> \<star>\<^sub>D F f)) \<cdot>\<^sub>D
D.inv (\<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D
\<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D
((F f \<star>\<^sub>D F \<eta>) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)))"
using antipar D.comp_assoc D.whisker_left D.whisker_right unit_char(2)
by simp
moreover have "F \<epsilon> \<star>\<^sub>D F f = D.inv (\<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f)"
using antipar \<Phi>.naturality [of "(\<epsilon>, f)"] C.VV.arr_char FF_def
D.invert_side_of_triangle(1) C.VV.dom_char C.VV.cod_char
by simp
moreover have "F f \<star>\<^sub>D F \<eta> = D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f)"
using antipar \<Phi>.naturality [of "(f, \<eta>)"] C.VV.arr_char FF_def
D.invert_side_of_triangle(1) C.VV.dom_char C.VV.cod_char
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))) \<cdot>\<^sub>D
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D
(F ((f \<star>\<^sub>C g) \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C f)) \<cdot>\<^sub>D
F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
using antipar D.comp_arr_inv' D.comp_assoc by simp
also have "... = ((D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))) \<cdot>\<^sub>D
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "F ((f \<star>\<^sub>C g) \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C f) = F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f]"
using antipar D.comp_arr_dom D.comp_cod_arr by simp
thus ?thesis by simp
qed
also have "... = D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) \<cdot>\<^sub>D
\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f)) =
D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
proof -
have "D.iso (\<Phi> (trg\<^sub>C f, f))"
using antipar by simp
moreover have "D.iso (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using antipar unit_char(2) by simp
moreover have "D.seq (\<Phi> (trg\<^sub>C f, f)) (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using antipar D.iso_is_arr calculation(2)
apply (intro D.seqI D.hseqI) by auto
ultimately show ?thesis
using antipar D.inv_comp unit_char(2) by simp
qed
moreover have "F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) =
F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>))"
using antipar by simp
ultimately show ?thesis by simp
qed
also have "... = (D.lunit' (F f) \<cdot>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D
F (C.lunit' f \<cdot>\<^sub>C \<r>\<^sub>C[f]) \<cdot>\<^sub>D
(D.inv (F \<r>\<^sub>C[f]) \<cdot>\<^sub>D \<r>\<^sub>D[F f])"
proof -
have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) = F (C.lunit' f \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using triangle_left by simp
moreover have "D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) =
D.lunit' (F f) \<cdot>\<^sub>D F \<l>\<^sub>C[f]"
proof -
have 0: "D.iso (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
using unit_char(2)
apply (intro D.isos_compose D.seqI) by auto
show ?thesis
proof -
have 1: "D.iso (F \<l>\<^sub>C[f])"
using C.iso_lunit preserves_iso by auto
moreover have "D.iso (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
by (metis (no_types) A.ide_left D.iso_lunit ide_left lunit_coherence)
moreover have "D.inv (D.inv (F \<l>\<^sub>C[f])) = F \<l>\<^sub>C[f]"
using 1 D.inv_inv by blast
ultimately show ?thesis
by (metis 0 D.inv_comp D.invert_side_of_triangle(2) D.iso_inv_iso
D.iso_is_arr ide_left lunit_coherence)
qed
qed
moreover have "\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) = D.inv (F \<r>\<^sub>C[f]) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
using ide_left runit_coherence preserves_iso C.iso_runit D.invert_side_of_triangle(1)
by (metis A.ide_left D.runit_simps(1))
ultimately show ?thesis by simp
qed
also have "... = D.lunit' (F f) \<cdot>\<^sub>D
((F \<l>\<^sub>C[f] \<cdot>\<^sub>D F (C.lunit' f)) \<cdot>\<^sub>D (F \<r>\<^sub>C[f] \<cdot>\<^sub>D D.inv (F \<r>\<^sub>C[f]))) \<cdot>\<^sub>D
\<r>\<^sub>D[F f]"
using D.comp_assoc by simp
also have "... = D.lunit' (F f) \<cdot>\<^sub>D \<r>\<^sub>D[F f]"
using D.comp_cod_arr C.iso_runit C.iso_lunit preserves_iso D.comp_arr_inv'
preserves_inv
by force
finally show ?thesis by blast
qed
show "(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
D.runit' (F g) \<cdot>\<^sub>D \<l>\<^sub>D[F g]"
proof -
have "\<a>\<^sub>D[F g, F f, F g] =
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g)"
proof -
have "D.iso (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)))"
using antipar D.iso_is_arr
apply (intro D.isos_compose, auto)
by (metis C.iso_assoc D.comp_assoc D.seqE ide_left ide_right
preserves_assoc(1) preserves_iso)
moreover have "F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g) =
\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, F g]"
using antipar assoc_coherence by simp
moreover have "D.seq (F \<a>\<^sub>C[g, f, g]) (\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g))"
using antipar C.VV.arr_char C.VV.dom_char C.VV.cod_char FF_def
by (intro D.seqI D.hseqI') auto
ultimately show ?thesis
using D.invert_side_of_triangle(1) D.comp_assoc by auto
qed
hence "(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
(F g \<star>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D
\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using D.comp_assoc by simp
also have "... = ((F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
(\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D ((D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g))"
proof -
have "F g \<star>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D \<Phi> (f, g) =
(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))"
proof -
have "D.seq (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) (\<Phi> (f, g))"
using antipar D.comp_assoc by simp
thus ?thesis
using antipar D.whisker_left by simp
qed
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g =
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using antipar D.whisker_right by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D
(((F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g))) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
((\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "D.inv (\<Phi> (g, f \<star>\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) =
D.inv (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g))"
proof -
have "D.iso (\<Phi> (g, f \<star>\<^sub>C g))"
using antipar by simp
moreover have "D.iso (F g \<star>\<^sub>D \<Phi> (f, g))"
using antipar by simp
moreover have "D.seq (\<Phi> (g, f \<star>\<^sub>C g)) (F g \<star>\<^sub>D \<Phi> (f, g))"
using antipar cmp_in_hom A.ide_right D.iso_is_arr
by (intro D.seqI) auto
ultimately show ?thesis
using antipar D.inv_comp by simp
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g)) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "((\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g) =
(F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "(\<Phi> (g, f) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) = F (g \<star>\<^sub>C f) \<star>\<^sub>D F g"
using antipar D.comp_arr_inv'
D.whisker_right [of "F g" "\<Phi> (g, f)" "D.inv (\<Phi> (g, f))"]
by simp
thus ?thesis
using antipar D.comp_cod_arr D.whisker_right by simp
qed
moreover have "((F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D D.inv (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
D.inv (\<Phi> (g, f \<star>\<^sub>C g)) =
D.inv (\<Phi> (g, f \<star>\<^sub>C g))"
using antipar D.comp_arr_inv' D.comp_cod_arr
D.whisker_left [of "F g" "\<Phi> (f, g)" "D.inv (\<Phi> (f, g))"]
by simp
ultimately show ?thesis by simp
qed
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D
((F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g))) \<cdot>\<^sub>D
F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D
(\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(unit (src\<^sub>C f) \<star>\<^sub>D F g)"
using antipar D.whisker_left D.whisker_right unit_char(2) D.comp_assoc by simp
also have "... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D
(F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g)) \<cdot>\<^sub>D
\<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (src\<^sub>C f) \<star>\<^sub>D F g)"
proof -
have "(F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D D.inv (\<Phi> (g, f \<star>\<^sub>C g)) = D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (g \<star>\<^sub>C \<epsilon>)"
using antipar C.VV.arr_char \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def
D.invert_opposite_sides_of_square C.VV.dom_char C.VV.cod_char
by simp
moreover have "\<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g) = F (\<eta> \<star>\<^sub>C g) \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g)"
using antipar C.VV.arr_char \<Phi>.naturality [of "(\<eta>, g)"] FF_def
C.VV.dom_char C.VV.cod_char
by simp
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = ((F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D
F (C.runit' g)) \<cdot>\<^sub>D (F \<l>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) \<cdot>\<^sub>D (unit (src\<^sub>C f) \<star>\<^sub>D F g))"
proof -
have "F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g) = F (C.runit' g) \<cdot>\<^sub>D F \<l>\<^sub>C[g]"
using ide_left ide_right antipar triangle_right
by (metis C.comp_in_homE C.seqI' preserves_comp triangle_in_hom(2))
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = D.runit' (F g) \<cdot>\<^sub>D \<l>\<^sub>D[F g]"
proof -
have "D.inv \<r>\<^sub>D[F g] =
(F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.runit' (F g) = D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using runit_coherence by simp
also have
"... = (F g \<star>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g)) \<cdot>\<^sub>D F (C.runit' g)"
proof -
have "D.iso (F \<r>\<^sub>C[g])"
using preserves_iso by simp
moreover have 1: "D.iso (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using preserves_iso unit_char(2) D.arrI D.seqE ide_right runit_coherence
by (intro D.isos_compose D.seqI, auto)
moreover have "D.seq (F \<r>\<^sub>C[g]) (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g)))"
using ide_right A.ide_right D.runit_simps(1) runit_coherence by metis
ultimately have "D.inv (F \<r>\<^sub>C[g] \<cdot>\<^sub>D \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) \<cdot>\<^sub>D F (C.runit' g)"
using C.iso_runit preserves_inv D.inv_comp by simp
moreover have "D.inv (\<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (src\<^sub>C g))) =
D.inv (F g \<star>\<^sub>D unit (src\<^sub>C g)) \<cdot>\<^sub>D D.inv (\<Phi> (g, src\<^sub>C g))"
proof -
have "D.seq (\<Phi> (g, src\<^sub>C g)) (F g \<star>\<^sub>D unit (src\<^sub>C g))"
using 1 antipar preserves_iso unit_char(2) by fast
(*
* TODO: The fact that auto cannot do this step is probably what is blocking
* the whole thing from being done by auto.
*)
thus ?thesis
using 1 antipar preserves_iso unit_char(2) D.inv_comp by auto
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
thus ?thesis
using antipar unit_char(2) preserves_iso by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using antipar lunit_coherence by simp
qed
finally show ?thesis by simp
qed
qed
qed
lemma preserves_adjoint_pair:
assumes "C.adjoint_pair f g"
shows "D.adjoint_pair (F f) (F g)"
using assms C.adjoint_pair_def D.adjoint_pair_def preserves_adjunction by blast
lemma preserves_left_adjoint:
assumes "C.is_left_adjoint f"
shows "D.is_left_adjoint (F f)"
using assms preserves_adjoint_pair by auto
lemma preserves_right_adjoint:
assumes "C.is_right_adjoint g"
shows "D.is_right_adjoint (F g)"
using assms preserves_adjoint_pair by auto
end
context equivalence_pseudofunctor
begin
lemma reflects_adjunction:
assumes "C.ide f" and "C.ide g"
and "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>" and "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
and "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
shows "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
proof -
let ?\<eta>' = "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
let ?\<epsilon>' = "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close> ?\<eta>' ?\<epsilon>'
using assms(5) by auto
interpret A: adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms(1-4) by (unfold_locales, auto)
show ?thesis
proof
show "(\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>) = \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f]"
proof -
have 1: "C.par ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using assms A.antipar by simp
moreover have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) = F (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
proof -
have "F ((\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>C \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>C (f \<star>\<^sub>C \<eta>)) =
F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, f] \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>)"
using 1 by auto
also have "... =
(F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f)) \<cdot>\<^sub>D
(\<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>))"
using assms A.antipar preserves_assoc(2) D.comp_assoc by auto
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D ((F \<epsilon> \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F f)) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
((F f \<star>\<^sub>D D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta>)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F (\<epsilon> \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, f) = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (F \<epsilon> \<star>\<^sub>D F f)"
using assms \<Phi>.naturality [of "(\<epsilon>, f)"] FF_def C.VV.arr_char
C.VV.dom_char C.VV.cod_char
by simp
moreover have "D.inv (\<Phi> (f, g \<star>\<^sub>C f)) \<cdot>\<^sub>D F (f \<star>\<^sub>C \<eta>) =
(F f \<star>\<^sub>D F \<eta>) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F (f \<star>\<^sub>C \<eta>) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) = \<Phi> (f, g \<star>\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D F \<eta>)"
using assms \<Phi>.naturality [of "(f, \<eta>)"] FF_def C.VV.arr_char A.antipar
C.VV.dom_char C.VV.cod_char
by simp
thus ?thesis
using assms A.antipar cmp_components_are_iso C.VV.arr_char cmp_in_hom
FF_def C.VV.dom_simp C.VV.cod_simp
D.invert_opposite_sides_of_square
[of "\<Phi> (f, g \<star>\<^sub>C f)" "F f \<star>\<^sub>D F \<eta>" "F (f \<star>\<^sub>C \<eta>)" "\<Phi> (f, src\<^sub>C f)"]
by fastforce
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta>) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using assms A.antipar cmp_in_hom A.ide_left A.ide_right A'.ide_left A'.ide_right
D.whisker_left [of "F f" "D.inv (\<Phi> (g, f))" "F \<eta>"]
D.whisker_right [of "F f" "F \<epsilon>" "\<Phi> (f, g)"]
by (metis A'.counit_in_vhom A'.unit_simps(1)D.arrI D.comp_assoc
D.src.preserves_reflects_arr D.src_vcomp D.vseq_implies_hpar(1) cmp_simps(2))
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>' \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D
(F f \<star>\<^sub>D ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
proof -
have "F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) = unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>'"
proof -
have "D.iso (unit (trg\<^sub>C f))"
using A.ide_left C.ideD(1) unit_char(2) by blast
thus ?thesis
by (metis A'.counit_simps(1) D.comp_assoc D.comp_cod_arr D.inv_is_inverse
D.seqE D.comp_arr_inv)
qed
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> = ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using assms(2) unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.unit_simps(1) A.antipar(1) C.ideD(1) C.obj_trg
D.invert_side_of_triangle(2))
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D ((unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D
(?\<epsilon>' \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D ((F f \<star>\<^sub>D ?\<eta>') \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f)))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using assms A.antipar A'.antipar unit_char D.whisker_left D.whisker_right
by simp
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D
((?\<epsilon>' \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D ?\<eta>')) \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using D.comp_assoc by simp
also have "... = (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]) \<cdot>\<^sub>D
\<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using A'.triangle_left D.comp_assoc by simp
also have "... = F \<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>D F \<r>\<^sub>C[f]"
using assms A.antipar preserves_lunit(2) preserves_runit(1) by simp
also have "... = F (\<l>\<^sub>C\<^sup>-\<^sup>1[f] \<cdot>\<^sub>C \<r>\<^sub>C[f])"
using assms by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using is_faithful by blast
qed
show "(g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g) = \<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g]"
proof -
have 1: "C.par ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C g f g \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
using assms A.antipar by auto
moreover have "F ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) = F (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
proof -
have "F ((g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>C \<a>\<^sub>C g f g \<cdot>\<^sub>C (\<eta> \<star>\<^sub>C g)) =
F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D F \<a>\<^sub>C[g, f, g] \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g)"
using 1 by auto
also have "... = (F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D \<Phi> (g, f \<star>\<^sub>C g)) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (D.inv (\<Phi> (g \<star>\<^sub>C f, g)) \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g))"
using assms A.antipar preserves_assoc(1) [of g f g] D.comp_assoc by auto
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D ((F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g))) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
((D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F (g \<star>\<^sub>C \<epsilon>) \<cdot>\<^sub>D \<Phi> (g, f \<star>\<^sub>C g) = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<epsilon>)"
using assms \<Phi>.naturality [of "(g, \<epsilon>)"] FF_def C.VV.arr_char
C.VV.dom_simp C.VV.cod_simp
by auto
moreover have "D.inv (\<Phi> (g \<star>\<^sub>C f, g)) \<cdot>\<^sub>D F (\<eta> \<star>\<^sub>C g) =
(F \<eta> \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F (\<eta> \<star>\<^sub>C g) \<cdot>\<^sub>D \<Phi> (trg\<^sub>C g, g) = \<Phi> (g \<star>\<^sub>C f, g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g)"
using assms \<Phi>.naturality [of "(\<eta>, g)"] FF_def C.VV.arr_char A.antipar
C.VV.dom_simp C.VV.cod_simp
by auto
thus ?thesis
using assms A.antipar cmp_components_are_iso C.VV.arr_char FF_def
C.VV.dom_simp C.VV.cod_simp
D.invert_opposite_sides_of_square
[of "\<Phi> (g \<star>\<^sub>C f, g)" "F \<eta> \<star>\<^sub>D F g" "F (\<eta> \<star>\<^sub>C g)" "\<Phi> (trg\<^sub>C g, g)"]
by fastforce
qed
ultimately show ?thesis
using D.comp_assoc by simp
qed
also have " ... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "(F g \<star>\<^sub>D F \<epsilon>) \<cdot>\<^sub>D (F g \<star>\<^sub>D \<Phi> (f, g)) = F g \<star>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
using assms A.antipar D.whisker_left
by (metis A'.counit_simps(1) A'.ide_right D.seqE)
moreover have "(D.inv (\<Phi> (g, f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D (F \<eta> \<star>\<^sub>D F g) =
D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<star>\<^sub>D F g"
using assms A.antipar D.whisker_right by simp
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>') \<cdot>\<^sub>D
\<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D
(?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
proof -
have "F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) = unit (trg\<^sub>C f) \<cdot>\<^sub>D ?\<epsilon>'"
using unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.counit_simps(1) C.ideD(1) C.obj_trg D.seqE assms(1))
moreover have "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> = ?\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using unit_char D.comp_arr_inv D.inv_is_inverse D.comp_assoc D.comp_cod_arr
by (metis A'.unit_simps(1) A.unit_simps(1) A.unit_simps(5)
C.obj_trg D.invert_side_of_triangle(2))
ultimately show ?thesis by simp
qed
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D
((F g \<star>\<^sub>D ?\<epsilon>') \<cdot>\<^sub>D \<a>\<^sub>D[F g, F f, F g] \<cdot>\<^sub>D (?\<eta>' \<star>\<^sub>D F g)) \<cdot>\<^sub>D
(D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using assms A.antipar unit_char D.whisker_left D.whisker_right D.comp_assoc
by simp
also have "... = \<Phi> (g, src\<^sub>C g) \<cdot>\<^sub>D (F g \<star>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F g] \<cdot>\<^sub>D
\<l>\<^sub>D[F g] \<cdot>\<^sub>D (D.inv (unit (src\<^sub>C f)) \<star>\<^sub>D F g) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C g, g))"
using A'.triangle_right D.comp_assoc by simp
also have "... = F \<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>D F \<l>\<^sub>C[g]"
using assms A.antipar preserves_lunit(1) preserves_runit(2) D.comp_assoc
by simp
also have "... = F (\<r>\<^sub>C\<^sup>-\<^sup>1[g] \<cdot>\<^sub>C \<l>\<^sub>C[g])"
using assms by simp
finally show ?thesis by simp
qed
ultimately show ?thesis
using is_faithful by blast
qed
qed
qed
lemma reflects_adjoint_pair:
assumes "C.ide f" and "C.ide g"
and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C f"
and "D.adjoint_pair (F f) (F g)"
shows "C.adjoint_pair f g"
proof -
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g) \<eta>' \<epsilon>'"
using assms D.adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close> \<eta>' \<epsilon>'
using A' by auto
have 1: "\<guillemotleft>\<Phi> (g, f) \<cdot>\<^sub>D \<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f)) : F (src\<^sub>C f) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
using assms unit_char [of "src\<^sub>C f"] A'.unit_in_hom
by (intro D.comp_in_homI, auto)
have 2: "\<guillemotleft>unit (trg\<^sub>C f) \<cdot>\<^sub>D \<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g)): F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using assms cmp_in_hom [of f g] unit_char [of "trg\<^sub>C f"] A'.counit_in_hom
by (intro D.comp_in_homI, auto)
obtain \<eta> where \<eta>: "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright> \<and>
F \<eta> = \<Phi> (g, f) \<cdot>\<^sub>D \<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
using assms 1 A'.unit_in_hom cmp_in_hom locally_full by fastforce
have \<eta>': "\<eta>' = D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)"
using assms 1 \<eta> cmp_in_hom D.iso_inv_iso cmp_components_are_iso unit_char(2)
D.invert_side_of_triangle(1) [of "F \<eta>" "\<Phi> (g, f)" "\<eta>' \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"]
D.invert_side_of_triangle(2) [of "D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta>" \<eta>' "D.inv (unit (src\<^sub>C f))"]
by (metis (no_types, lifting) C.ideD(1) C.obj_trg D.arrI D.comp_assoc D.inv_inv)
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C trg\<^sub>C f\<guillemotright> \<and>
F \<epsilon> = unit (trg\<^sub>C f) \<cdot>\<^sub>D \<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g))"
using assms 2 A'.counit_in_hom cmp_in_hom locally_full by fastforce
have \<epsilon>': "\<epsilon>' = D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)"
using assms 2 \<epsilon> cmp_in_hom D.iso_inv_iso unit_char(2) D.comp_assoc
D.invert_side_of_triangle(1) [of "F \<epsilon>" "unit (trg\<^sub>C f)" "\<epsilon>' \<cdot>\<^sub>D D.inv (\<Phi> (f, g))"]
D.invert_side_of_triangle(2) [of "D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon>" \<epsilon>' "D.inv (\<Phi> (f, g))"]
by (metis (no_types, lifting) C.arrI C.ideD(1) C.obj_trg D.inv_inv cmp_components_are_iso
preserves_arr)
have "adjunction_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
using A'.adjunction_in_bicategory_axioms \<eta>' \<epsilon>' by simp
hence "adjunction_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms \<eta> \<epsilon> reflects_adjunction by simp
thus ?thesis
using C.adjoint_pair_def by auto
qed
lemma reflects_left_adjoint:
assumes "C.ide f" and "D.is_left_adjoint (F f)"
shows "C.is_left_adjoint f"
proof -
obtain g' where g': "D.adjoint_pair (F f) g'"
using assms D.adjoint_pair_def by auto
obtain g where g: "\<guillemotleft>g : trg\<^sub>C f \<rightarrow>\<^sub>C src\<^sub>C f\<guillemotright> \<and> C.ide g \<and> D.isomorphic (F g) g'"
using assms g' locally_essentially_surjective [of "trg\<^sub>C f" "src\<^sub>C f" g']
D.adjoint_pair_antipar [of "F f" g']
by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : g' \<Rightarrow>\<^sub>D F g\<guillemotright> \<and> D.iso \<phi>"
using g D.isomorphic_def D.isomorphic_symmetric by metis
have "D.adjoint_pair (F f) (F g)"
using assms g g' \<phi> D.adjoint_pair_preserved_by_iso [of "F f" g' "F f" "F f" \<phi> "F g"]
by auto
thus ?thesis
using assms g reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def
by auto
qed
lemma reflects_right_adjoint:
assumes "C.ide g" and "D.is_right_adjoint (F g)"
shows "C.is_right_adjoint g"
proof -
obtain f' where f': "D.adjoint_pair f' (F g)"
using assms D.adjoint_pair_def by auto
obtain f where f: "\<guillemotleft>f : trg\<^sub>C g \<rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) f'"
using assms f' locally_essentially_surjective [of "trg\<^sub>C g" "src\<^sub>C g" f']
D.adjoint_pair_antipar [of f' "F g"]
by auto
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : f' \<Rightarrow>\<^sub>D F f\<guillemotright> \<and> D.iso \<phi>"
using f D.isomorphic_def D.isomorphic_symmetric by metis
have "D.adjoint_pair (F f) (F g)"
using assms f f' \<phi> D.adjoint_pair_preserved_by_iso [of f' "F g" \<phi> "F f" "F g" "F g"]
by auto
thus ?thesis
using assms f reflects_adjoint_pair [of f g] D.adjoint_pair_antipar C.in_hhom_def
by auto
qed
end
subsection "Composition of Adjunctions"
text \<open>
We first consider the strict case, then extend to all bicategories using strictification.
\<close>
locale composite_adjunction_in_strict_bicategory =
strict_bicategory V H \<a> \<i> src trg +
fg: adjunction_in_strict_bicategory V H \<a> \<i> src trg f g \<zeta> \<xi> +
hk: adjunction_in_strict_bicategory V H \<a> \<i> src trg h k \<sigma> \<tau>
for V :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<cdot>" 55)
and H :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<star>" 53)
and \<a> :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src :: "'a \<Rightarrow> 'a"
and trg :: "'a \<Rightarrow> 'a"
and f :: "'a"
and g :: "'a"
and \<zeta> :: "'a"
and \<xi> :: "'a"
and h :: "'a"
and k :: "'a"
and \<sigma> :: "'a"
and \<tau> :: "'a" +
assumes composable: "src h = trg f"
begin
abbreviation \<eta>
where "\<eta> \<equiv> (g \<star> \<sigma> \<star> f) \<cdot> \<zeta>"
abbreviation \<epsilon>
where "\<epsilon> \<equiv> \<tau> \<cdot> (h \<star> \<xi> \<star> k)"
interpretation adjunction_data_in_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "ide (h \<star> f)"
using composable by simp
show "ide (g \<star> k)"
using fg.antipar hk.antipar composable by simp
show "\<guillemotleft>\<eta> : src (h \<star> f) \<Rightarrow> (g \<star> k) \<star> h \<star> f\<guillemotright>"
proof
show "\<guillemotleft>\<zeta> : src (h \<star> f) \<Rightarrow> g \<star> f\<guillemotright>"
using fg.antipar hk.antipar composable \<open>ide (h \<star> f)\<close> by auto
show "\<guillemotleft>g \<star> \<sigma> \<star> f : g \<star> f \<Rightarrow> (g \<star> k) \<star> h \<star> f\<guillemotright>"
proof -
have "\<guillemotleft>g \<star> \<sigma> \<star> f : g \<star> trg f \<star> f \<Rightarrow> g \<star> (k \<star> h) \<star> f\<guillemotright>"
using fg.antipar hk.antipar composable hk.unit_in_hom
apply (intro hcomp_in_vhom) by auto
thus ?thesis
using hcomp_obj_arr hcomp_assoc by fastforce
qed
qed
show "\<guillemotleft>\<epsilon> : (h \<star> f) \<star> g \<star> k \<Rightarrow> src (g \<star> k)\<guillemotright>"
proof
show "\<guillemotleft>h \<star> \<xi> \<star> k : (h \<star> f) \<star> g \<star> k \<Rightarrow> h \<star> k\<guillemotright>"
proof -
have "\<guillemotleft>h \<star> \<xi> \<star> k : h \<star> (f \<star> g) \<star> k \<Rightarrow> h \<star> trg f \<star> k\<guillemotright>"
using composable fg.antipar(1-2) hk.antipar(1) by fastforce
thus ?thesis
using fg.antipar hk.antipar composable hk.unit_in_hom hcomp_obj_arr hcomp_assoc
by simp
qed
show "\<guillemotleft>\<tau> : h \<star> k \<Rightarrow> src (g \<star> k)\<guillemotright>"
using fg.antipar hk.antipar composable hk.unit_in_hom by auto
qed
qed
sublocale adjunction_in_strict_bicategory V H \<a> \<i> src trg \<open>h \<star> f\<close> \<open>g \<star> k\<close> \<eta> \<epsilon>
proof
show "(\<epsilon> \<star> h \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[h \<star> f, g \<star> k, h \<star> f] \<cdot> ((h \<star> f) \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[h \<star> f] \<cdot> \<r>[h \<star> f]"
proof -
have "(\<epsilon> \<star> h \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[h \<star> f, g \<star> k, h \<star> f] \<cdot> ((h \<star> f) \<star> \<eta>) =
(\<tau> \<cdot> (h \<star> \<xi> \<star> k) \<star> h \<star> f) \<cdot> ((h \<star> f) \<star> (g \<star> \<sigma> \<star> f) \<cdot> \<zeta>)"
using fg.antipar hk.antipar composable strict_assoc comp_ide_arr
ide_left ide_right antipar(1) antipar(2)
by (metis arrI seqE strict_assoc' triangle_in_hom(1))
also have "... = (\<tau> \<star> h \<star> f) \<cdot> ((h \<star> \<xi> \<star> (k \<star> h) \<star> f) \<cdot> (h \<star> (f \<star> g) \<star> \<sigma> \<star> f)) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable whisker_left [of "h \<star> f"] whisker_right
comp_assoc hcomp_assoc
by simp
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> (\<xi> \<star> (k \<star> h)) \<cdot> ((f \<star> g) \<star> \<sigma>) \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> (trg f \<star> \<sigma>) \<cdot> (\<xi> \<star> trg f) \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable comp_arr_dom comp_cod_arr
interchange [of \<xi> "f \<star> g" "k \<star> h" \<sigma>] interchange [of "trg f" \<xi> \<sigma> "trg f"]
by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3)
hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3))
also have "... = (\<tau> \<star> h \<star> f) \<cdot> (h \<star> \<sigma> \<cdot> \<xi> \<star> f) \<cdot> (h \<star> f \<star> \<zeta>)"
using fg.antipar hk.antipar composable hcomp_obj_arr hcomp_arr_obj
by (metis fg.counit_simps(1) fg.counit_simps(4) hk.unit_simps(1) hk.unit_simps(5)
obj_src)
also have "... = ((\<tau> \<star> h \<star> f) \<cdot> (h \<star> \<sigma> \<star> f)) \<cdot> ((h \<star> \<xi> \<star> f) \<cdot> (h \<star> f \<star> \<zeta>))"
using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc
by simp
also have "... = ((\<tau> \<star> h) \<cdot> (h \<star> \<sigma>) \<star> f) \<cdot> (h \<star> (\<xi> \<star> f) \<cdot> (f \<star> \<zeta>))"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = h \<star> f"
using fg.antipar hk.antipar composable fg.triangle_left hk.triangle_left
by simp
also have "... = \<l>\<^sup>-\<^sup>1[h \<star> f] \<cdot> \<r>[h \<star> f]"
using fg.antipar hk.antipar composable strict_lunit' strict_runit by simp
finally show ?thesis by simp
qed
show "((g \<star> k) \<star> \<epsilon>) \<cdot> \<a>[g \<star> k, h \<star> f, g \<star> k] \<cdot> (\<eta> \<star> g \<star> k) = \<r>\<^sup>-\<^sup>1[g \<star> k] \<cdot> \<l>[g \<star> k]"
proof -
have "((g \<star> k) \<star> \<epsilon>) \<cdot> \<a>[g \<star> k, h \<star> f, g \<star> k] \<cdot> (\<eta> \<star> g \<star> k) =
((g \<star> k) \<star> \<tau> \<cdot> (h \<star> \<xi> \<star> k)) \<cdot> ((g \<star> \<sigma> \<star> f) \<cdot> \<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable strict_assoc comp_ide_arr
ide_left ide_right
by (metis antipar(1) antipar(2) arrI seqE triangle_in_hom(2))
also have "... = (g \<star> k \<star> \<tau>) \<cdot> ((g \<star> (k \<star> h) \<star> \<xi> \<star> k) \<cdot> (g \<star> \<sigma> \<star> (f \<star> g) \<star> k)) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left [of "g \<star> k"] whisker_right
comp_assoc hcomp_assoc
by simp
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> ((k \<star> h) \<star> \<xi>) \<cdot> (\<sigma> \<star> f \<star> g) \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> (\<sigma> \<star> src g) \<cdot> (src g \<star> \<xi>) \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable interchange [of "k \<star> h" \<sigma> \<xi> "f \<star> g"]
interchange [of \<sigma> "src g" "src g" \<xi>] comp_arr_dom comp_cod_arr
by (metis fg.counit_simps(1) fg.counit_simps(2) fg.counit_simps(3)
hk.unit_simps(1) hk.unit_simps(2) hk.unit_simps(3))
also have "... = (g \<star> k \<star> \<tau>) \<cdot> (g \<star> \<sigma> \<cdot> \<xi> \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable hcomp_obj_arr [of "src g" \<xi>]
hcomp_arr_obj [of \<sigma> "src g"]
by simp
also have "... = ((g \<star> k \<star> \<tau>) \<cdot> (g \<star> \<sigma> \<star> k)) \<cdot> (g \<star> \<xi> \<star> k) \<cdot> (\<zeta> \<star> g \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right comp_assoc
by simp
also have "... = (g \<star> (k \<star> \<tau>) \<cdot> (\<sigma> \<star> k)) \<cdot> ((g \<star> \<xi>) \<cdot> (\<zeta> \<star> g) \<star> k)"
using fg.antipar hk.antipar composable whisker_left whisker_right hcomp_assoc
by simp
also have "... = g \<star> k"
using fg.antipar hk.antipar composable fg.triangle_right hk.triangle_right
by simp
also have "... = \<r>\<^sup>-\<^sup>1[g \<star> k] \<cdot> \<l>[g \<star> k]"
using fg.antipar hk.antipar composable strict_lunit strict_runit' by simp
finally show ?thesis by simp
qed
qed
lemma is_adjunction_in_strict_bicategory:
shows "adjunction_in_strict_bicategory V H \<a> \<i> src trg (h \<star> f) (g \<star> k) \<eta> \<epsilon>"
..
end
context strict_bicategory
begin
lemma left_adjoints_compose:
assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f' = trg f"
shows "is_left_adjoint (f' \<star> f)"
proof -
obtain g \<eta> \<epsilon> where fg: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by auto
obtain g' \<eta>' \<epsilon>' where f'g': "adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret f'g': adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'
using f'g' by auto
interpret f'fgg': composite_adjunction_in_strict_bicategory V H \<a> \<i> src trg
f g \<eta> \<epsilon> f' g' \<eta>' \<epsilon>'
using assms apply unfold_locales by simp
have "adjoint_pair (f' \<star> f) (g \<star> g')"
using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto
thus ?thesis by auto
qed
lemma right_adjoints_compose:
assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g = trg g'"
shows "is_right_adjoint (g \<star> g')"
proof -
obtain f \<eta> \<epsilon> where fg: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by auto
obtain f' \<eta>' \<epsilon>' where f'g': "adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret f'g': adjunction_in_bicategory V H \<a> \<i> src trg f' g' \<eta>' \<epsilon>'
using f'g' by auto
interpret f'fgg': composite_adjunction_in_strict_bicategory V H \<a> \<i> src trg
f g \<eta> \<epsilon> f' g' \<eta>' \<epsilon>'
using assms fg.antipar f'g'.antipar apply unfold_locales by simp
have "adjoint_pair (f' \<star> f) (g \<star> g')"
using adjoint_pair_def f'fgg'.adjunction_in_bicategory_axioms by auto
thus ?thesis by auto
qed
end
text \<open>
We now use strictification to extend the preceding results to an arbitrary bicategory.
We only prove that ``left adjoints compose'' and ``right adjoints compose'';
I did not work out formulas for the unit and counit of the composite adjunction in the
non-strict case.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
using S.UP_is_equivalence_pseudofunctor by auto
lemma left_adjoints_compose:
assumes "is_left_adjoint f" and "is_left_adjoint f'" and "src f = trg f'"
shows "is_left_adjoint (f \<star> f')"
proof -
have "S.is_left_adjoint (S.UP f) \<and> S.is_left_adjoint (S.UP f')"
using assms UP.preserves_left_adjoint by simp
moreover have "S.src (S.UP f) = S.trg (S.UP f')"
using assms left_adjoint_is_ide by simp
ultimately have "S.is_left_adjoint (S.hcomp (S.UP f) (S.UP f'))"
using S.left_adjoints_compose by simp
moreover have "S.isomorphic (S.hcomp (S.UP f) (S.UP f')) (S.UP (f \<star> f'))"
proof -
have "\<guillemotleft>S.cmp\<^sub>U\<^sub>P (f, f') : S.hcomp (S.UP f) (S.UP f') \<Rightarrow>\<^sub>S S.UP (f \<star> f')\<guillemotright>"
using assms left_adjoint_is_ide UP.cmp_in_hom by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f, f'))"
using assms left_adjoint_is_ide by simp
ultimately show ?thesis
using S.isomorphic_def by blast
qed
ultimately have "S.is_left_adjoint (S.UP (f \<star> f'))"
using S.left_adjoint_preserved_by_iso S.isomorphic_def by blast
thus "is_left_adjoint (f \<star> f')"
using assms left_adjoint_is_ide UP.reflects_left_adjoint by simp
qed
lemma right_adjoints_compose:
assumes "is_right_adjoint g" and "is_right_adjoint g'" and "src g' = trg g"
shows "is_right_adjoint (g' \<star> g)"
proof -
have "S.is_right_adjoint (S.UP g) \<and> S.is_right_adjoint (S.UP g')"
using assms UP.preserves_right_adjoint by simp
moreover have "S.src (S.UP g') = S.trg (S.UP g)"
using assms right_adjoint_is_ide by simp
ultimately have "S.is_right_adjoint (S.hcomp (S.UP g') (S.UP g))"
using S.right_adjoints_compose by simp
moreover have "S.isomorphic (S.hcomp (S.UP g') (S.UP g)) (S.UP (g' \<star> g))"
proof -
have "\<guillemotleft>S.cmp\<^sub>U\<^sub>P (g', g) : S.hcomp (S.UP g') (S.UP g) \<Rightarrow>\<^sub>S S.UP (g' \<star> g)\<guillemotright>"
using assms right_adjoint_is_ide UP.cmp_in_hom by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (g', g))"
using assms right_adjoint_is_ide by simp
ultimately show ?thesis
using S.isomorphic_def by blast
qed
ultimately have "S.is_right_adjoint (S.UP (g' \<star> g))"
using S.right_adjoint_preserved_by_iso S.isomorphic_def by blast
thus "is_right_adjoint (g' \<star> g)"
using assms right_adjoint_is_ide UP.reflects_right_adjoint by simp
qed
end
subsection "Choosing Right Adjoints"
text \<open>
It will be useful in various situations to suppose that we have made a choice of
right adjoint for each left adjoint ({\it i.e.} each ``map'') in a bicategory.
\<close>
locale chosen_right_adjoints =
bicategory
begin
(* Global notation is evil! *)
no_notation Transitive_Closure.rtrancl ("(_\<^sup>*)" [1000] 999)
definition some_right_adjoint ("_\<^sup>*" [1000] 1000)
where "f\<^sup>* \<equiv> SOME g. adjoint_pair f g"
definition some_unit
where "some_unit f \<equiv> SOME \<eta>. \<exists>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* \<eta> \<epsilon>"
definition some_counit
where "some_counit f \<equiv>
SOME \<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) \<epsilon>"
lemma left_adjoint_extends_to_adjunction:
assumes "is_left_adjoint f"
shows "adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) (some_counit f)"
using assms some_right_adjoint_def adjoint_pair_def some_unit_def some_counit_def
someI_ex [of "\<lambda>g. adjoint_pair f g"]
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjunction_in_bicategory V H \<a> \<i> src trg f f\<^sup>* (some_unit f) \<epsilon>"]
by auto
lemma left_adjoint_extends_to_adjoint_pair:
assumes "is_left_adjoint f"
shows "adjoint_pair f f\<^sup>*"
using assms adjoint_pair_def left_adjoint_extends_to_adjunction by blast
lemma right_adjoint_in_hom [intro]:
assumes "is_left_adjoint f"
shows "\<guillemotleft>f\<^sup>* : trg f \<rightarrow> src f\<guillemotright>"
and "\<guillemotleft>f\<^sup>* : f\<^sup>* \<Rightarrow> f\<^sup>*\<guillemotright>"
using assms left_adjoint_extends_to_adjoint_pair adjoint_pair_antipar [of f "f\<^sup>*"]
by auto
lemma right_adjoint_simps [simp]:
assumes "is_left_adjoint f"
shows "ide f\<^sup>*"
and "src f\<^sup>* = trg f" and "trg f\<^sup>* = src f"
and "dom f\<^sup>* = f\<^sup>*" and "cod f\<^sup>* = f\<^sup>*"
using assms right_adjoint_in_hom left_adjoint_extends_to_adjoint_pair apply auto
using assms right_adjoint_is_ide [of "f\<^sup>*"] by blast
end
locale map_in_bicategory =
bicategory + chosen_right_adjoints +
fixes f :: 'a
assumes is_map: "is_left_adjoint f"
begin
abbreviation \<eta>
where "\<eta> \<equiv> some_unit f"
abbreviation \<epsilon>
where "\<epsilon> \<equiv> some_counit f"
sublocale adjunction_in_bicategory V H \<a> \<i> src trg f \<open>f\<^sup>*\<close> \<eta> \<epsilon>
using is_map left_adjoint_extends_to_adjunction by simp
end
subsection "Equivalences Refine to Adjoint Equivalences"
text \<open>
In this section, we show that, just as an equivalence between categories can always
be refined to an adjoint equivalence, an internal equivalence in a bicategory can also
always be so refined.
The proof, which follows that of Theorem 3.3 from \cite{nlab-adjoint-equivalence},
makes use of the fact that if an internal equivalence satisfies one of the triangle
identities, then it also satisfies the other.
\<close>
locale adjoint_equivalence_in_bicategory =
equivalence_in_bicategory +
adjunction_in_bicategory
begin
lemma dual_adjoint_equivalence:
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg g f (inv \<epsilon>) (inv \<eta>)"
proof -
interpret gf: equivalence_in_bicategory V H \<a> \<i> src trg g f \<open>inv \<epsilon>\<close> \<open>inv \<eta>\<close>
using dual_equivalence by simp
show ?thesis
proof
show "(inv \<eta> \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[g, f, g] \<cdot> (g \<star> inv \<epsilon>) = \<l>\<^sup>-\<^sup>1[g] \<cdot> \<r>[g]"
proof -
have "(inv \<eta> \<star> g) \<cdot> \<a>\<^sup>-\<^sup>1[g, f, g] \<cdot> (g \<star> inv \<epsilon>) =
inv ((g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g))"
using antipar inv_comp isos_compose comp_assoc by simp
also have "... = inv (\<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g])"
using triangle_right by simp
also have "... = \<l>\<^sup>-\<^sup>1[g] \<cdot> \<r>[g]"
using inv_comp by simp
finally show ?thesis
by blast
qed
show "(f \<star> inv \<eta>) \<cdot> \<a>[f, g, f] \<cdot> (inv \<epsilon> \<star> f) = \<r>\<^sup>-\<^sup>1[f] \<cdot> \<l>[f]"
proof -
have "(f \<star> inv \<eta>) \<cdot> \<a>[f, g, f] \<cdot> (inv \<epsilon> \<star> f) =
inv ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>))"
using antipar inv_comp isos_compose comp_assoc by simp
also have "... = inv (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
using triangle_left by simp
also have "... = \<r>\<^sup>-\<^sup>1[f] \<cdot> \<l>[f]"
using inv_comp by simp
finally show ?thesis by blast
qed
qed
qed
end
context bicategory
begin
lemma adjoint_equivalence_preserved_by_iso_right:
assumes "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : g \<Rightarrow> g'\<guillemotright>" and "iso \<phi>"
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g' ((\<phi> \<star> f) \<cdot> \<eta>) (\<epsilon> \<cdot> (f \<star> inv \<phi>))"
proof -
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by simp
interpret fg': adjunction_in_bicategory V H \<a> \<i> src trg f g' \<open>(\<phi> \<star> f) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (f \<star> inv \<phi>)\<close>
using assms fg.adjunction_in_bicategory_axioms adjunction_preserved_by_iso_right
by simp
interpret fg': equivalence_in_bicategory V H \<a> \<i> src trg f g' \<open>(\<phi> \<star> f) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (f \<star> inv \<phi>)\<close>
using assms fg.equivalence_in_bicategory_axioms equivalence_preserved_by_iso_right
by simp
show ?thesis ..
qed
lemma adjoint_equivalence_preserved_by_iso_left:
assumes "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
and "\<guillemotleft>\<phi> : f \<Rightarrow> f'\<guillemotright>" and "iso \<phi>"
shows "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f' g ((g \<star> \<phi>) \<cdot> \<eta>) (\<epsilon> \<cdot> (inv \<phi> \<star> g))"
proof -
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms by simp
interpret fg': adjunction_in_bicategory V H \<a> \<i> src trg f' g \<open>(g \<star> \<phi>) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (inv \<phi> \<star> g)\<close>
using assms fg.adjunction_in_bicategory_axioms adjunction_preserved_by_iso_left
by simp
interpret fg': equivalence_in_bicategory V H \<a> \<i> src trg f' g \<open>(g \<star> \<phi>) \<cdot> \<eta>\<close> \<open>\<epsilon> \<cdot> (inv \<phi> \<star> g)\<close>
using assms fg.equivalence_in_bicategory_axioms equivalence_preserved_by_iso_left
by simp
show ?thesis ..
qed
end
context strict_bicategory
begin
notation isomorphic (infix "\<cong>" 50)
lemma equivalence_refines_to_adjoint_equivalence:
assumes "equivalence_map f" and "\<guillemotleft>g : trg f \<rightarrow> src f\<guillemotright>" and "ide g"
and "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>" and "iso \<eta>"
shows "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
obtain g' \<eta>' \<epsilon>' where E': "equivalence_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'"
using assms equivalence_map_def by auto
interpret E': equivalence_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'
using E' by auto
let ?a = "src f" and ?b = "trg f"
(* TODO: in_homE cannot be applied automatically to a conjunction. Must break down! *)
have f_in_hhom: "\<guillemotleft>f : ?a \<rightarrow> ?b\<guillemotright>" and ide_f: "ide f"
using assms equivalence_map_def by auto
have g_in_hhom: "\<guillemotleft>g : ?b \<rightarrow> ?a\<guillemotright>" and ide_g: "ide g"
using assms by auto
have g'_in_hhom: "\<guillemotleft>g' : ?b \<rightarrow> ?a\<guillemotright>" and ide_g': "ide g'"
using assms f_in_hhom E'.antipar by auto
have \<eta>_in_hom: "\<guillemotleft>\<eta> : ?a \<Rightarrow> g \<star> f\<guillemotright>" and iso_\<eta>: "iso \<eta>"
using assms by auto
have a: "obj ?a" and b: "obj ?b"
using f_in_hhom by auto
have \<eta>_in_hhom: "\<guillemotleft>\<eta> : ?a \<rightarrow> ?a\<guillemotright>"
using a src_dom trg_dom \<eta>_in_hom by fastforce
text \<open>
The following is quoted from \cite{nlab-adjoint-equivalence}:
\begin{quotation}
``Since \<open>g \<cong> gfg' \<cong> g'\<close>, the isomorphism \<open>fg' \<cong> 1\<close> also induces an isomorphism \<open>fg \<cong> 1\<close>,
which we denote \<open>\<xi>\<close>. Now \<open>\<eta>\<close> and \<open>\<xi>\<close> may not satisfy the zigzag identities, but if we
define \<open>\<epsilon>\<close> by \<open>\<xi> \<cdot> (f \<star> \<eta>\<^sup>-\<^sup>1) \<cdot> (f \<star> g \<star> \<xi>\<^sup>-\<^sup>1) : f \<star> g \<Rightarrow> 1\<close>, then we can verify,
using string diagram notation as above,
that \<open>\<epsilon>\<close> satisfies one zigzag identity, and hence (by the previous lemma) also the other.
Finally, if \<open>\<epsilon>': fg \<Rightarrow> 1\<close> is any other isomorphism satisfying the zigzag identities
with \<open>\<eta>\<close>, then we have:
\[
\<open>\<epsilon>' = \<epsilon>' \<cdot> (\<epsilon> f g) \<cdot> (f \<eta> g) = \<epsilon> \<cdot> (f g \<epsilon>') \<cdot> (f \<eta> g) = \<epsilon>\<close>
\]
using the interchange law and two zigzag identities. This shows uniqueness.''
\end{quotation}
\<close>
have 1: "g \<cong> g'"
proof -
have "g \<cong> g \<star> ?b"
using assms hcomp_arr_obj isomorphic_reflexive by auto
also have "... \<cong> g \<star> f \<star> g'"
using assms f_in_hhom g_in_hhom g'_in_hhom E'.counit_in_vhom E'.counit_is_iso
isomorphic_def hcomp_ide_isomorphic isomorphic_symmetric
by (metis E'.counit_simps(5) in_hhomE trg_trg)
also have "... \<cong> ?a \<star> g'"
using assms f_in_hhom g_in_hhom g'_in_hhom ide_g' E'.unit_in_vhom E'.unit_is_iso
isomorphic_def hcomp_isomorphic_ide isomorphic_symmetric
by (metis hcomp_assoc hcomp_isomorphic_ide in_hhomE src_src)
also have "... \<cong> g'"
using assms
by (simp add: E'.antipar(1) hcomp_obj_arr isomorphic_reflexive)
finally show ?thesis by blast
qed
have "f \<star> g' \<cong> ?b"
using E'.counit_is_iso isomorphicI [of \<epsilon>'] by auto
hence 2: "f \<star> g \<cong> ?b"
using assms 1 ide_f hcomp_ide_isomorphic [of f g g'] isomorphic_transitive
isomorphic_symmetric
by (metis in_hhomE)
obtain \<xi> where \<xi>: "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> iso \<xi>"
using 2 by auto
have \<xi>_in_hom: "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright>" and iso_\<xi>: "iso \<xi>"
using \<xi> by auto
have \<xi>_in_hhom: "\<guillemotleft>\<xi> : ?b \<rightarrow> ?b\<guillemotright>"
using b src_cod trg_cod \<xi>_in_hom by fastforce
text \<open>
At the time of this writing, the definition of \<open>\<epsilon>\<close> given on nLab
\cite{nlab-adjoint-equivalence} had an apparent typo:
the expression \<open>f \<star> g \<star> \<xi>\<^sup>-\<^sup>1\<close> should read \<open>\<xi>\<^sup>-\<^sup>1 \<star> f \<star> g\<close>, as we have used here.
\<close>
let ?\<epsilon> = "\<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g)"
have \<epsilon>_in_hom: "\<guillemotleft>?\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
proof (intro comp_in_homI)
show "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> g \<star> f \<star> g \<Rightarrow> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> (g \<star> f) \<star> g \<Rightarrow> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>f \<star> inv \<eta> \<star> g : f \<star> (g \<star> f) \<star> g \<Rightarrow> f \<star> ?a \<star> g\<guillemotright>"
using assms \<eta>_in_hom iso_\<eta> by (intro hcomp_in_vhom, auto)
thus ?thesis
using assms f_in_hhom hcomp_obj_arr by (metis in_hhomE)
qed
moreover have "f \<star> (g \<star> f) \<star> g = f \<star> g \<star> f \<star> g"
using hcomp_assoc by simp
ultimately show ?thesis by simp
qed
show "\<guillemotleft>inv \<xi> \<star> f \<star> g : f \<star> g \<Rightarrow> f \<star> g \<star> f \<star> g\<guillemotright>"
proof -
have "\<guillemotleft>inv \<xi> \<star> f \<star> g : ?b \<star> f \<star> g \<Rightarrow> (f \<star> g) \<star> f \<star> g\<guillemotright>"
using assms \<xi>_in_hom iso_\<xi> by (intro hcomp_in_vhom, auto)
thus ?thesis
using hcomp_assoc f_in_hhom g_in_hhom b hcomp_obj_arr [of ?b "f \<star> g"]
by fastforce
qed
show "\<guillemotleft>\<xi> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
using \<xi>_in_hom by blast
qed
have "iso ?\<epsilon>"
using f_in_hhom g_in_hhom \<eta>_in_hhom ide_f ide_g \<eta>_in_hom iso_\<eta> \<xi>_in_hhom \<xi>_in_hom iso_\<xi>
iso_inv_iso isos_compose
by (metis \<epsilon>_in_hom arrI hseqE ide_is_iso iso_hcomp seqE)
have 4: "\<guillemotleft>inv \<xi> \<star> f : ?b \<star> f \<Rightarrow> f \<star> g \<star> f\<guillemotright>"
proof -
have "\<guillemotleft>inv \<xi> \<star> f : ?b \<star> f \<Rightarrow> (f \<star> g) \<star> f\<guillemotright>"
using \<xi>_in_hom iso_\<xi> f_in_hhom
by (intro hcomp_in_vhom, auto)
thus ?thesis
using hcomp_assoc by simp
qed
text \<open>
First show \<open>?\<epsilon>\<close> and \<open>\<eta>\<close> satisfy the ``left'' triangle identity.
\<close>
have triangle_left: "(?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = f"
proof -
have "(?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) = (\<xi> \<star> f) \<cdot> (f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>)"
proof -
have "f \<star> \<eta> = ?b \<star> f \<star> \<eta>"
using b \<eta>_in_hhom hcomp_obj_arr [of ?b "f \<star> \<eta>"] by fastforce
moreover have "\<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g) \<star> f =
(\<xi> \<star> f) \<cdot> ((f \<star> inv \<eta> \<star> g) \<star> f) \<cdot> ((inv \<xi> \<star> f \<star> g) \<star> f)"
using ide_f ide_g \<xi>_in_hhom \<xi>_in_hom iso_\<xi> \<eta>_in_hhom \<eta>_in_hom iso_\<eta> whisker_right
by (metis \<epsilon>_in_hom arrI seqE)
moreover have "... = (\<xi> \<star> f) \<cdot> (f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f \<star> g \<star> f)"
using hcomp_assoc by simp
ultimately show ?thesis
using comp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (f \<star> g \<star> f \<star> \<eta>)) \<cdot> (inv \<xi> \<star> f)"
proof -
have "(inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>) = ((inv \<xi> \<star> f) \<star> (g \<star> f)) \<cdot> ((?b \<star> f) \<star> \<eta>)"
using hcomp_assoc by simp
also have "... = (inv \<xi> \<star> f) \<cdot> (?b \<star> f) \<star> (g \<star> f) \<cdot> \<eta>"
proof -
have "seq (inv \<xi> \<star> f) (?b \<star> f)"
using a b 4 ide_f ide_g \<xi>_in_hhom \<xi>_in_hom iso_\<xi> by blast
moreover have "seq (g \<star> f) \<eta>"
using f_in_hhom g_in_hhom ide_g ide_f \<eta>_in_hom by fast
ultimately show ?thesis
using interchange [of "inv \<xi> \<star> f" "?b \<star> f" "g \<star> f" \<eta>] by simp
qed
also have "... = inv \<xi> \<star> f \<star> \<eta>"
using 4 comp_arr_dom comp_cod_arr \<eta>_in_hom hcomp_assoc by (metis in_homE)
also have "... = (f \<star> g) \<cdot> inv \<xi> \<star> (f \<star> \<eta>) \<cdot> (f \<star> ?a)"
proof -
have "(f \<star> g) \<cdot> inv \<xi> = inv \<xi>"
using \<xi>_in_hom iso_\<xi> comp_cod_arr by auto
moreover have "(f \<star> \<eta>) \<cdot> (f \<star> ?a) = f \<star> \<eta>"
proof -
have "\<guillemotleft>f \<star> \<eta> : f \<star> ?a \<Rightarrow> f \<star> g \<star> f\<guillemotright>"
using \<eta>_in_hom by fastforce
thus ?thesis
using comp_arr_dom by blast
qed
ultimately show ?thesis by argo
qed
also have "... = ((f \<star> g) \<star> (f \<star> \<eta>)) \<cdot> (inv \<xi> \<star> (f \<star> ?a))"
proof -
have "seq (f \<star> g) (inv \<xi>)"
using \<xi>_in_hom iso_\<xi> comp_cod_arr by auto
moreover have "seq (f \<star> \<eta>) (f \<star> ?a)"
using f_in_hhom \<eta>_in_hom by force
ultimately show ?thesis
using interchange by simp
qed
also have "... = (f \<star> g \<star> f \<star> \<eta>) \<cdot> (inv \<xi> \<star> f)"
using hcomp_arr_obj hcomp_assoc by auto
finally have "(inv \<xi> \<star> f \<star> g \<star> f) \<cdot> (?b \<star> f \<star> \<eta>) = (f \<star> g \<star> f \<star> \<eta>) \<cdot> (inv \<xi> \<star> f)"
by simp
thus ?thesis
using comp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> ?a \<star> \<eta>) \<cdot> (f \<star> inv \<eta> \<star> ?a)) \<cdot> (inv \<xi> \<star> f)"
proof -
have "(f \<star> inv \<eta> \<star> g \<star> f) \<cdot> (f \<star> (g \<star> f) \<star> \<eta>) = f \<star> (inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>)"
proof -
have "(f \<star> (inv \<eta> \<star> g) \<star> f) \<cdot> (f \<star> (g \<star> f) \<star> \<eta>) = f \<star> ((inv \<eta> \<star> g) \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>)"
proof -
have "seq ((inv \<eta> \<star> g) \<star> f) ((g \<star> f) \<star> \<eta>)"
proof -
have "seq (inv \<eta> \<star> g \<star> f) ((g \<star> f) \<star> \<eta>)"
using f_in_hhom ide_f g_in_hhom ide_g \<eta>_in_hhom \<eta>_in_hom iso_\<eta>
apply (intro seqI hseqI')
apply auto
by fastforce+
thus ?thesis
using hcomp_assoc by simp
qed
thus ?thesis
using whisker_left by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = f \<star> (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
proof -
have "(inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>) = (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
proof -
have "(inv \<eta> \<star> g \<star> f) \<cdot> ((g \<star> f) \<star> \<eta>) = inv \<eta> \<cdot> (g \<star> f) \<star> (g \<star> f) \<cdot> \<eta>"
using g_in_hhom ide_g \<eta>_in_hom iso_\<eta>
interchange [of "inv \<eta>" "g \<star> f" "g \<star> f" \<eta>]
by force
also have "... = inv \<eta> \<star> \<eta>"
using \<eta>_in_hom iso_\<eta> comp_arr_dom comp_cod_arr by auto
also have "... = ?a \<cdot> inv \<eta> \<star> \<eta> \<cdot> ?a"
using \<eta>_in_hom iso_\<eta> comp_arr_dom comp_cod_arr by auto
also have "... = (?a \<star> \<eta>) \<cdot> (inv \<eta> \<star> ?a)"
using a \<eta>_in_hom iso_\<eta> interchange [of ?a "inv \<eta>" \<eta> ?a] by blast
finally show ?thesis by simp
qed
thus ?thesis by argo
qed
also have "... = (f \<star> ?a \<star> \<eta>) \<cdot> (f \<star> inv \<eta> \<star> ?a)"
proof -
have "seq (?a \<star> \<eta>) (inv \<eta> \<star> ?a)"
proof (intro seqI')
show "\<guillemotleft>inv \<eta> \<star> ?a : (g \<star> f) \<star> ?a \<Rightarrow> ?a \<star> ?a\<guillemotright>"
using a g_in_hhom \<eta>_in_hom iso_\<eta> hseqI ide_f ide_g
by (intro hcomp_in_vhom) auto
show "\<guillemotleft>?a \<star> \<eta> : ?a \<star> ?a \<Rightarrow> ?a \<star> g \<star> f\<guillemotright>"
using a \<eta>_in_hom hseqI by (intro hcomp_in_vhom) auto
qed
thus ?thesis
using whisker_left by simp
qed
finally show ?thesis
using hcomp_assoc by simp
qed
also have "... = (\<xi> \<star> f) \<cdot> ((f \<star> \<eta>) \<cdot> (f \<star> inv \<eta>)) \<cdot> (inv \<xi> \<star> f)"
using a \<eta>_in_hhom iso_\<eta> hcomp_obj_arr [of ?a \<eta>] hcomp_arr_obj [of "inv \<eta>" ?a] by auto
also have "... = (\<xi> \<star> f) \<cdot> (inv \<xi> \<star> f)"
proof -
have "((f \<star> \<eta>) \<cdot> (f \<star> inv \<eta>)) \<cdot> (inv \<xi> \<star> f) = (f \<star> \<eta> \<cdot> inv \<eta>) \<cdot> (inv \<xi> \<star> f)"
using \<eta>_in_hhom iso_\<eta> whisker_left inv_in_hom by auto
also have "... = (f \<star> g \<star> f) \<cdot> (inv \<xi> \<star> f)"
using \<eta>_in_hom iso_\<eta> comp_arr_inv inv_is_inverse by auto
also have "... = inv \<xi> \<star> f"
using 4 comp_cod_arr by blast
ultimately show ?thesis by simp
qed
also have "... = f"
proof -
have "(\<xi> \<star> f) \<cdot> (inv \<xi> \<star> f) = \<xi> \<cdot> inv \<xi> \<star> f"
using \<xi>_in_hhom iso_\<xi> whisker_right by auto
also have "... = ?b \<star> f"
using \<xi>_in_hom iso_\<xi> comp_arr_inv' by auto
also have "... = f"
using hcomp_obj_arr by auto
finally show ?thesis by blast
qed
finally show ?thesis by blast
qed
(* TODO: Putting this earlier breaks some steps in the proof. *)
interpret E: equivalence_in_strict_bicategory V H \<a> \<i> src trg f g \<eta> ?\<epsilon>
- using ide_g \<eta>_in_hom \<epsilon>_in_hom g_in_hhom `iso \<eta>` `iso ?\<epsilon>`
+ using ide_g \<eta>_in_hom \<epsilon>_in_hom g_in_hhom \<open>iso \<eta>\<close> \<open>iso ?\<epsilon>\<close>
by (unfold_locales, auto)
text \<open>
Apply ``triangle left if and only iff right'' to show the ``right'' triangle identity.
\<close>
have triangle_right: "((g \<star> \<xi> \<cdot> (f \<star> inv \<eta> \<star> g) \<cdot> (inv \<xi> \<star> f \<star> g)) \<cdot> (\<eta> \<star> g) = g)"
using triangle_left E.triangle_left_iff_right by simp
text \<open>
Use the two triangle identities to establish an adjoint equivalence and show that
there is only one choice for the counit.
\<close>
show "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
have "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> ?\<epsilon>"
proof
show "(?\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "(?\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = (?\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>)"
proof -
have "seq \<a>\<^sup>-\<^sup>1[f, g, f] (f \<star> \<eta>)"
using E.antipar
by (intro seqI, auto)
hence "\<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = f \<star> \<eta>"
using ide_f ide_g E.antipar triangle_right strict_assoc' comp_ide_arr
by presburger
thus ?thesis by simp
qed
also have "... = f"
using triangle_left by simp
also have "... = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
using strict_lunit strict_runit by simp
finally show ?thesis by simp
qed
show "(g \<star> ?\<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
proof -
have "(g \<star> ?\<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = (g \<star> ?\<epsilon>) \<cdot> (\<eta> \<star> g)"
proof -
have "seq \<a>[g, f, g] (\<eta> \<star> g)"
using E.antipar
by (intro seqI, auto)
hence "\<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<eta> \<star> g"
using ide_f ide_g E.antipar triangle_right strict_assoc comp_ide_arr
by presburger
thus ?thesis by simp
qed
also have "... = g"
using triangle_right by simp
also have "... = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using strict_lunit strict_runit by simp
finally show ?thesis by blast
qed
qed
moreover have "\<And>\<epsilon> \<epsilon>'. \<lbrakk> adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>;
adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>' \<rbrakk>
\<Longrightarrow> \<epsilon> = \<epsilon>'"
using adjunction_unit_determines_counit
by (meson adjoint_equivalence_in_bicategory.axioms(2))
ultimately show ?thesis by auto
qed
qed
end
text \<open>
We now apply strictification to generalize the preceding result to an arbitrary bicategory.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
using S.UP_is_equivalence_pseudofunctor by auto
interpretation UP: pseudofunctor_into_strict_bicategory V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
..
lemma equivalence_refines_to_adjoint_equivalence:
assumes "equivalence_map f" and "\<guillemotleft>g : trg f \<rightarrow> src f\<guillemotright>" and "ide g"
and "\<guillemotleft>\<eta> : src f \<Rightarrow> g \<star> f\<guillemotright>" and "iso \<eta>"
shows "\<exists>!\<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
text \<open>
To unpack the consequences of the assumptions, we need to obtain an
interpretation of @{locale equivalence_in_bicategory}, even though we don't
need the associated data other than \<open>f\<close>, \<open>a\<close>, and \<open>b\<close>.
\<close>
obtain g' \<phi> \<psi> where E: "equivalence_in_bicategory V H \<a> \<i> src trg f g' \<phi> \<psi>"
using assms equivalence_map_def by auto
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g' \<phi> \<psi>
using E by auto
let ?a = "src f" and ?b = "trg f"
have ide_f: "ide f" by simp
have f_in_hhom: "\<guillemotleft>f : ?a \<rightarrow> ?b\<guillemotright>" by simp
have a: "obj ?a" and b: "obj ?b" by auto
have 1: "S.equivalence_map (S.UP f)"
using assms UP.preserves_equivalence_maps by simp
let ?\<eta>' = "S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a"
have 2: "\<guillemotleft>S.UP \<eta> : S.UP ?a \<Rightarrow>\<^sub>S S.UP (g \<star> f)\<guillemotright>"
using assms UP.preserves_hom [of \<eta> "src f" "g \<star> f"] by auto
have 3: "\<guillemotleft>?\<eta>' : UP.map\<^sub>0 ?a \<Rightarrow>\<^sub>S S.UP g \<star>\<^sub>S S.UP f\<guillemotright> \<and> S.iso ?\<eta>'"
proof (intro S.comp_in_homI conjI)
show "\<guillemotleft>S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) : S.UP (g \<star> f) \<Rightarrow>\<^sub>S S.UP g \<star>\<^sub>S S.UP f\<guillemotright>"
using assms UP.cmp_components_are_iso UP.cmp_in_hom(2) [of g f] UP.FF_def
by auto
moreover show "\<guillemotleft>UP.unit ?a : UP.map\<^sub>0 ?a \<Rightarrow>\<^sub>S S.UP ?a\<guillemotright>" by auto
moreover show "\<guillemotleft>S.UP \<eta> : S.UP ?a \<Rightarrow>\<^sub>S S.UP (g \<star> f)\<guillemotright>"
using 2 by simp
ultimately show "S.iso (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a)"
using assms UP.cmp_components_are_iso UP.unit_char(2)
by (intro S.isos_compose) auto
qed
have ex_un_\<xi>': "\<exists>!\<xi>'. adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
(S.UP f) (S.UP g) ?\<eta>' \<xi>'"
proof -
have "\<guillemotleft>S.UP g : S.trg (S.UP f) \<rightarrow>\<^sub>S S.src (S.UP f)\<guillemotright>"
using assms(2) by auto
moreover have "S.ide (S.UP g)"
by (simp add: assms(3))
ultimately show ?thesis
using 1 3 S.equivalence_refines_to_adjoint_equivalence S.UP_map\<^sub>0_obj by simp
qed
obtain \<xi>' where \<xi>': "adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
(S.UP f) (S.UP g) ?\<eta>' \<xi>'"
using ex_un_\<xi>' by auto
interpret E': adjoint_equivalence_in_bicategory S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg
\<open>S.UP f\<close> \<open>S.UP g\<close> ?\<eta>' \<xi>'
using \<xi>' by auto
let ?\<epsilon>' = "UP.unit ?b \<cdot>\<^sub>S \<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g))"
have \<epsilon>': "\<guillemotleft>?\<epsilon>' : S.UP (f \<star> g) \<Rightarrow>\<^sub>S S.UP ?b\<guillemotright>"
using assms(2-3) S.UP_map\<^sub>0_obj
by (intro S.in_homI) auto
have ex_un_\<epsilon>: "\<exists>!\<epsilon>. \<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
proof -
have "\<exists>\<epsilon>. \<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
proof -
have "src (f \<star> g) = src ?b \<and> trg (f \<star> g) = trg ?b"
using assms(2) f_in_hhom by auto
moreover have "ide (f \<star> g)"
using assms(2-3) by auto
ultimately show ?thesis
using \<epsilon>' UP.locally_full by auto
qed
moreover have
"\<And>\<mu> \<nu>. \<lbrakk> \<guillemotleft>\<mu> : f \<star> g \<Rightarrow> ?b\<guillemotright>; S.UP \<mu> = ?\<epsilon>'; \<guillemotleft>\<nu> : f \<star> g \<Rightarrow> ?b\<guillemotright>; S.UP \<nu> = ?\<epsilon>' \<rbrakk>
\<Longrightarrow> \<mu> = \<nu>"
proof -
fix \<mu> \<nu>
assume \<mu>: "\<guillemotleft>\<mu> : f \<star> g \<Rightarrow> ?b\<guillemotright>" and \<nu>: "\<guillemotleft>\<nu> : f \<star> g \<Rightarrow> ?b\<guillemotright>"
and 1: "S.UP \<mu> = ?\<epsilon>'" and 2: "S.UP \<nu> = ?\<epsilon>'"
have "par \<mu> \<nu>"
using \<mu> \<nu> by fastforce
thus "\<mu> = \<nu>"
using 1 2 UP.is_faithful [of \<mu> \<nu>] by simp
qed
ultimately show ?thesis by auto
qed
have iso_\<epsilon>': "S.iso ?\<epsilon>'"
proof (intro S.isos_compose)
show "S.iso (S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))"
using assms UP.cmp_components_are_iso by auto
show "S.iso \<xi>'"
using E'.counit_is_iso by blast
show "S.iso (UP.unit ?b)"
using b UP.unit_char(2) by simp
show "S.seq (UP.unit ?b) (\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))"
proof (intro S.seqI')
show "\<guillemotleft>UP.unit ?b : UP.map\<^sub>0 ?b \<Rightarrow>\<^sub>S S.UP ?b\<guillemotright>"
using b UP.unit_char by simp
show "\<guillemotleft>\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) : S.UP (f \<star> g) \<Rightarrow>\<^sub>S UP.map\<^sub>0 ?b\<guillemotright>"
using assms UP.cmp_components_are_iso VV.arr_char S.cmp\<^sub>U\<^sub>P_in_hom [of "(f, g)"]
E'.counit_in_hom S.UP_map\<^sub>0_obj
by (intro S.comp_in_homI) auto
qed
thus "S.seq \<xi>' (S.inv (S.cmp\<^sub>U\<^sub>P (f, g)))" by auto
qed
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star> g \<Rightarrow> ?b\<guillemotright> \<and> S.UP \<epsilon> = ?\<epsilon>'"
using ex_un_\<epsilon> by auto
interpret E'': equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using assms(1,3-5)
apply unfold_locales
apply simp_all
using assms(2) \<epsilon>
apply auto[1]
using \<epsilon> iso_\<epsilon>' UP.reflects_iso [of \<epsilon>]
by auto
interpret E'': adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
proof
show "(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) = \<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f]"
proof -
have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) =
S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (S.UP \<epsilon> \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
using E''.UP_triangle(3) by simp
also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S
(UP.unit ?b \<cdot>\<^sub>S \<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
using \<epsilon> S.comp_assoc by simp
also have "... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<cdot>\<^sub>S \<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "\<xi>' \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) = \<xi>'"
proof -
have "S.iso (S.cmp\<^sub>U\<^sub>P (f, g))"
using assms by auto
moreover have "S.dom (S.cmp\<^sub>U\<^sub>P (f, g)) = S.UP f \<star>\<^sub>S S.UP g"
using assms by auto
ultimately have "S.inv (S.cmp\<^sub>U\<^sub>P (f, g)) \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, g) = S.UP f \<star>\<^sub>S S.UP g"
using S.comp_inv_arr' by simp
thus ?thesis
using S.comp_arr_dom E'.counit_in_hom(2) by simp
qed
thus ?thesis by argo
qed
also have
"... = S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S
((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "UP.unit ?b \<cdot>\<^sub>S \<xi>' \<star>\<^sub>S S.UP f = (UP.unit ?b \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (\<xi>' \<star>\<^sub>S S.UP f)"
using assms b UP.unit_char S.whisker_right S.UP_map\<^sub>0_obj by auto
moreover have "S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> =
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)"
using assms S.whisker_left S.comp_assoc by auto
ultimately show ?thesis
using S.comp_assoc by simp
qed
also have "... = (S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>)
= (S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>) =
S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>"
using assms S.whisker_left by auto
hence "((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f))) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.UP \<eta>))
= ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>))"
by simp
also have "... = ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) = \<xi>' \<star>\<^sub>S S.UP f"
proof -
have "\<guillemotleft>\<xi>' \<star>\<^sub>S S.UP f :
(S.UP f \<star>\<^sub>S S.UP g) \<star>\<^sub>S S.UP f \<Rightarrow>\<^sub>S S.trg (S.UP f) \<star>\<^sub>S S.UP f\<guillemotright>"
using assms by (intro S.hcomp_in_vhom, auto)
moreover have "\<guillemotleft>S.\<a>' (S.UP f) (S.UP g) (S.UP f) :
S.UP f \<star>\<^sub>S S.UP g \<star>\<^sub>S S.UP f \<Rightarrow>\<^sub>S (S.UP f \<star>\<^sub>S S.UP g) \<star>\<^sub>S S.UP f\<guillemotright>"
using assms S.assoc'_in_hom by auto
ultimately show ?thesis
using assms S.strict_assoc' S.iso_assoc S.hcomp_assoc E'.antipar
S.comp_arr_ide S.seqI'
by (metis (no_types, lifting) E'.ide_left E'.ide_right)
qed
thus ?thesis
using S.comp_assoc by simp
qed
also have "... = ((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>))"
using S.comp_assoc by simp
also have "... = (S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) =
(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f)))"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
S.lunit' (S.UP f) \<cdot>\<^sub>S S.runit (S.UP f)"
proof -
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a)"
proof -
have "S.seq (S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) (UP.unit ?a)"
using assms UP.unit_char UP.cmp_components_are_iso
E'.unit_simps(1) S.comp_assoc
by presburger
hence "(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) =
S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta> \<cdot>\<^sub>S UP.unit ?a"
using assms UP.unit_char UP.cmp_components_are_iso S.comp_assoc
S.whisker_left [of "S.UP f" "S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>" "UP.unit ?a"]
by simp
thus ?thesis by simp
qed
thus ?thesis
using assms E'.triangle_left UP.cmp_components_are_iso UP.unit_char
by argo
qed
also have "... = S.UP f"
using S.strict_lunit' S.strict_runit by simp
finally have 1: "((\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S UP.unit ?a) = S.UP f"
using S.comp_assoc by simp
have "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>) =
S.UP f \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S S.inv (UP.unit ?a))"
proof -
have "S.arr (S.UP f)"
using assms by simp
moreover have "S.iso (S.UP f \<star>\<^sub>S UP.unit ?a)"
using assms UP.unit_char S.UP_map\<^sub>0_obj by auto
moreover have "S.inv (S.UP f \<star>\<^sub>S UP.unit ?a) =
S.UP f \<star>\<^sub>S S.inv (UP.unit ?a)"
using assms a UP.unit_char S.UP_map\<^sub>0_obj by auto
ultimately show ?thesis
using assms 1 UP.unit_char UP.cmp_components_are_iso
S.invert_side_of_triangle(2)
[of "S.UP f" "(\<xi>' \<star>\<^sub>S S.UP f) \<cdot>\<^sub>S S.\<a>' (S.UP f) (S.UP g) (S.UP f) \<cdot>\<^sub>S
(S.UP f \<star>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (g, f)) \<cdot>\<^sub>S S.UP \<eta>)"
"S.UP f \<star>\<^sub>S UP.unit ?a"]
by presburger
qed
also have "... = S.UP f \<star>\<^sub>S S.inv (UP.unit ?a)"
proof -
have "\<guillemotleft>S.UP f \<star>\<^sub>S S.inv (UP.unit ?a) :
S.UP f \<star>\<^sub>S S.UP ?a \<Rightarrow>\<^sub>S S.UP f \<star>\<^sub>S UP.map\<^sub>0 ?a\<guillemotright>"
using assms ide_f f_in_hhom UP.unit_char [of ?a] S.inv_in_hom
apply (intro S.hcomp_in_vhom)
apply auto[1]
apply blast
by auto
moreover have "S.UP f \<star>\<^sub>S UP.map\<^sub>0 ?a = S.UP f"
using a S.hcomp_arr_obj S.UP_map\<^sub>0_obj by auto
finally show ?thesis
using S.comp_cod_arr by blast
qed
finally show ?thesis by auto
qed
thus ?thesis
using S.comp_assoc by simp
qed
finally show ?thesis by simp
qed
thus ?thesis
using S.comp_assoc by simp
qed
also have "... = S.UP \<l>\<^sup>-\<^sup>1[f] \<cdot>\<^sub>S S.UP \<r>[f]"
proof -
have "S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit ?b \<star>\<^sub>S S.UP f) = S.UP \<l>\<^sup>-\<^sup>1[f]"
proof -
have "S.UP f = S.UP \<l>[f] \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (trg f, f) \<cdot>\<^sub>S (UP.unit (trg f) \<star>\<^sub>S S.UP f)"
using UP.lunit_coherence iso_lunit S.strict_lunit by simp
thus ?thesis
using UP.image_of_unitor(3) ide_f by presburger
qed
moreover have "(S.UP f \<star>\<^sub>S S.inv (UP.unit (src f))) \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f))
= S.UP \<r>[f]"
proof -
have "S.UP \<r>[f] \<cdot>\<^sub>S S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)) = S.UP f"
using UP.runit_coherence [of f] S.strict_runit by simp
moreover have "S.iso (S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)))"
using UP.unit_char UP.cmp_components_are_iso VV.arr_char S.UP_map\<^sub>0_obj
by (intro S.isos_compose S.seqI) auto
ultimately have
"S.UP \<r>[f] = S.UP f \<cdot>\<^sub>S S.inv (S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f)))"
using S.invert_side_of_triangle(2)
[of "S.UP f" "S.UP \<r>[f]" "S.cmp\<^sub>U\<^sub>P (f, src f) \<cdot>\<^sub>S (S.UP f \<star>\<^sub>S UP.unit (src f))"]
ideD(1) ide_f by blast
thus ?thesis
using ide_f UP.image_of_unitor(2) [of f] by argo
qed
ultimately show ?thesis
using S.comp_assoc by simp
qed
also have "... = S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
by simp
finally have "S.UP ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) = S.UP (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
by simp
moreover have "par ((\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>)) (\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f])"
proof -
have "\<guillemotleft>(\<epsilon> \<star> f) \<cdot> \<a>\<^sup>-\<^sup>1[f, g, f] \<cdot> (f \<star> \<eta>) : f \<star> src f \<Rightarrow> trg f \<star> f\<guillemotright>"
using E''.triangle_in_hom(1) by simp
moreover have "\<guillemotleft>\<l>\<^sup>-\<^sup>1[f] \<cdot> \<r>[f] : f \<star> src f \<Rightarrow> trg f \<star> f\<guillemotright>" by auto
ultimately show ?thesis
by (metis in_homE)
qed
ultimately show ?thesis
using UP.is_faithful by blast
qed
thus "(g \<star> \<epsilon>) \<cdot> \<a>[g, f, g] \<cdot> (\<eta> \<star> g) = \<r>\<^sup>-\<^sup>1[g] \<cdot> \<l>[g]"
using E''.triangle_left_implies_right by simp
qed
show ?thesis
using E''.adjoint_equivalence_in_bicategory_axioms E''.adjunction_in_bicategory_axioms
adjunction_unit_determines_counit adjoint_equivalence_in_bicategory_def
by metis
qed
lemma equivalence_map_extends_to_adjoint_equivalence:
assumes "equivalence_map f"
shows "\<exists>g \<eta> \<epsilon>. adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
proof -
obtain g \<eta> \<epsilon>' where E: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using assms equivalence_map_def by auto
interpret E: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using E by auto
obtain \<epsilon> where A: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
E.antipar E.unit_is_iso E.unit_in_hom by auto
show ?thesis
using E A by blast
qed
end
subsection "Uniqueness of Adjoints"
text \<open>
Left and right adjoints determine each other up to isomorphism.
\<close>
context strict_bicategory
begin
lemma left_adjoint_determines_right_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f g'"
shows "g \<cong> g'"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
interpret A: adjunction_in_strict_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon> ..
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>'
using A' by auto
interpret A': adjunction_in_strict_bicategory V H \<a> \<i> src trg f g' \<eta>' \<epsilon>' ..
let ?\<phi> = "A'.trnl\<^sub>\<eta> g \<epsilon>"
have "\<guillemotleft>?\<phi>: g \<Rightarrow> g'\<guillemotright>"
using A'.trnl\<^sub>\<eta>_eq A'.adjoint_transpose_left(1) [of "trg f" g] A.antipar A'.antipar
hcomp_arr_obj
by auto
moreover have "iso ?\<phi>"
proof (intro isoI)
let ?\<psi> = "A.trnl\<^sub>\<eta> g' \<epsilon>'"
show "inverse_arrows ?\<phi> ?\<psi>"
proof
show "ide (?\<phi> \<cdot> ?\<psi>)"
proof -
have 1: "ide (trg f) \<and> trg (trg f) = trg f"
by simp
have "?\<phi> \<cdot> ?\<psi> = (g' \<star> \<epsilon>) \<cdot> ((\<eta>' \<star> g) \<cdot> (g \<star> \<epsilon>')) \<cdot> (\<eta> \<star> g')"
using 1 A.antipar A'.antipar A.trnl\<^sub>\<eta>_eq [of "trg f" g' \<epsilon>']
A'.trnl\<^sub>\<eta>_eq [of "trg f" g \<epsilon>] comp_assoc A.counit_in_hom A'.counit_in_hom
by simp
also have "... = ((g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>')) \<cdot> ((\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g'))"
proof -
have "(\<eta>' \<star> g) \<cdot> (g \<star> \<epsilon>') = (\<eta>' \<star> g \<star> trg f) \<cdot> (src f \<star> g \<star> \<epsilon>')"
using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr [of "src f" "g \<star> \<epsilon>'"]
hseqI'
by (metis A'.counit_simps(1) A'.counit_simps(5) A.ide_right ideD(1)
obj_trg trg_hcomp)
also have "... = \<eta>' \<star> g \<star> \<epsilon>'"
using A.antipar A'.antipar interchange [of \<eta>' "src f" "g \<star> trg f" "g \<star> \<epsilon>'"]
whisker_left comp_arr_dom comp_cod_arr
by simp
also have "... = ((g' \<star> f) \<star> g \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g \<star> (f \<star> g'))"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
A'.unit_in_hom A'.counit_in_hom interchange whisker_left
comp_arr_dom comp_cod_arr
by (metis A'.counit_simps(1-2,5) A'.unit_simps(1,3) hseqI' ide_char)
also have "... = (g' \<star> f \<star> g \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g \<star> f \<star> g')"
using hcomp_assoc by simp
finally show ?thesis
using comp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> ((g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g')) \<cdot> (\<eta>' \<star> g')"
proof -
have "(g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>') = (g' \<star> \<epsilon>') \<cdot> (g' \<star> \<epsilon> \<star> f \<star> g')"
proof -
have "(g' \<star> \<epsilon>) \<cdot> (g' \<star> f \<star> g \<star> \<epsilon>') = g' \<star> \<epsilon> \<star> \<epsilon>'"
proof -
have "\<epsilon> \<cdot> (f \<star> g \<star> \<epsilon>') = \<epsilon> \<star> \<epsilon>'"
using A.ide_left A.ide_right A.antipar A'.antipar hcomp_arr_obj comp_arr_dom
comp_cod_arr interchange obj_src trg_src
by (metis A'.counit_simps(1,3) A.counit_simps(1-2,4) hcomp_assoc)
thus ?thesis
using A.antipar A'.antipar whisker_left [of g' \<epsilon> "f \<star> g \<star> \<epsilon>'"]
by (simp add: hcomp_assoc)
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> (g' \<star> \<epsilon> \<star> f \<star> g')"
proof -
have "\<epsilon> \<star> \<epsilon>' = \<epsilon>' \<cdot> (\<epsilon> \<star> f \<star> g')"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
hcomp_obj_arr hcomp_arr_obj comp_arr_dom comp_cod_arr interchange
obj_src trg_src
by (metis A'.counit_simps(1-2,5) A.counit_simps(1,3-4) arr_cod
not_arr_null seq_if_composable)
thus ?thesis
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
whisker_left
by (metis A'.counit_simps(1,5) A.counit_simps(1,4) hseqI')
qed
finally show ?thesis by simp
qed
moreover have "(\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g') = (g' \<star> f \<star> \<eta> \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "(\<eta>' \<star> g \<star> f \<star> g') \<cdot> (\<eta> \<star> g') = \<eta>' \<star> \<eta> \<star> g'"
proof -
have "(\<eta>' \<star> g \<star> f) \<cdot> \<eta> = \<eta>' \<star> \<eta>"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
interchange comp_arr_dom comp_cod_arr
by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3,5) hcomp_obj_arr obj_trg)
thus ?thesis
using A.antipar A'.antipar whisker_right [of g' "\<eta>' \<star> g \<star> f" \<eta>]
by (simp add: hcomp_assoc)
qed
also have "... = (g' \<star> f \<star> \<eta> \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "\<eta>' \<star> \<eta> = (g' \<star> f \<star> \<eta>) \<cdot> \<eta>'"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
comp_arr_dom comp_cod_arr hcomp_assoc interchange
by (metis A'.unit_simps(1,3-4) A.unit_simps(1-2) obj_src)
thus ?thesis
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom hcomp_arr_obj
whisker_right [of g' "g' \<star> f \<star> \<eta>" \<eta>']
by (metis A'.ide_right A'.unit_simps(1,4) A.unit_simps(1,5)
hseqI' hcomp_assoc)
qed
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc hcomp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> ((g' \<star> f) \<star> g') \<cdot> (\<eta>' \<star> g')"
proof -
have "(g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g') = g' \<star> f \<star> g'"
proof -
have "(g' \<star> (\<epsilon> \<star> f) \<star> g') \<cdot> (g' \<star> (f \<star> \<eta>) \<star> g') =
g' \<star> ((\<epsilon> \<star> f) \<star> g') \<cdot> ((f \<star> \<eta>) \<star> g')"
using A.ide_left A.ide_right A.antipar A'.antipar A'.unit_in_hom
A'.counit_in_hom whisker_left [of g' "(\<epsilon> \<star> f) \<star> g'" "(f \<star> \<eta>) \<star> g'"]
by (metis A'.ide_right A.triangle_left hseqI' ideD(1) whisker_right)
also have "... = g' \<star> (\<epsilon> \<star> f) \<cdot> (f \<star> \<eta>) \<star> g'"
using A.antipar A'.antipar whisker_right [of g' "\<epsilon> \<star> f" "f \<star> \<eta>"]
by (simp add: A.triangle_left)
also have "... = g' \<star> f \<star> g'"
using A.triangle_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = (g' \<star> \<epsilon>') \<cdot> (\<eta>' \<star> g')"
using A.antipar A'.antipar A'.unit_in_hom A'.counit_in_hom comp_cod_arr
by (metis A'.ide_right A'.triangle_in_hom(2) A.ide_left arrI assoc_is_natural_2
ide_char seqE strict_assoc)
also have "... = g'"
using A'.triangle_right by simp
finally have "?\<phi> \<cdot> ?\<psi> = g'" by simp
thus ?thesis by simp
qed
show "ide (?\<psi> \<cdot> ?\<phi>)"
proof -
have 1: "ide (trg f) \<and> trg (trg f) = trg f"
by simp
have "?\<psi> \<cdot> ?\<phi> = (g \<star> \<epsilon>') \<cdot> ((\<eta> \<star> g') \<cdot> (g' \<star> \<epsilon>)) \<cdot> (\<eta>' \<star> g)"
using A.antipar A'.antipar A'.trnl\<^sub>\<eta>_eq [of "trg f" g \<epsilon>]
A.trnl\<^sub>\<eta>_eq [of "trg f" g' \<epsilon>'] comp_assoc A.counit_in_hom A'.counit_in_hom
by simp
also have "... = ((g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>)) \<cdot> ((\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g))"
proof -
have "(\<eta> \<star> g') \<cdot> (g' \<star> \<epsilon>) = (\<eta> \<star> g' \<star> trg f) \<cdot> (src f \<star> g' \<star> \<epsilon>)"
using A.antipar A'.antipar hcomp_arr_obj hcomp_obj_arr hseqI'
by (metis A'.ide_right A.unit_simps(1,4) hcomp_assoc hcomp_obj_arr
ideD(1) obj_src)
also have "... = \<eta> \<star> g' \<star> \<epsilon>"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A.counit_in_hom interchange
by (metis "1" A.counit_simps(5) A.unit_simps(4) hseqI' ide_def ide_in_hom(2)
not_arr_null seqI' src.preserves_ide)
also have "... = ((g \<star> f) \<star> g' \<star> \<epsilon>) \<cdot> (\<eta> \<star> g' \<star> (f \<star> g))"
using A'.ide_right A'.antipar interchange ide_char comp_arr_dom comp_cod_arr hseqI'
by (metis A.counit_simps(1-2,5) A.unit_simps(1,3))
also have "... = (g \<star> f \<star> g' \<star> \<epsilon>) \<cdot> (\<eta> \<star> g' \<star> f \<star> g)"
using hcomp_assoc by simp
finally show ?thesis
using comp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> ((g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g)) \<cdot> (\<eta> \<star> g)"
proof -
have "(g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>) = (g \<star> \<epsilon>) \<cdot> (g \<star> \<epsilon>' \<star> f \<star> g)"
proof -
have "(g \<star> \<epsilon>') \<cdot> (g \<star> f \<star> g' \<star> \<epsilon>) = g \<star> \<epsilon>' \<star> \<epsilon>"
proof -
have "\<epsilon>' \<cdot> (f \<star> g' \<star> \<epsilon>) = \<epsilon>' \<star> \<epsilon>"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_arr_obj
comp_arr_dom comp_cod_arr interchange obj_src trg_src hcomp_assoc
by (metis A.counit_simps(1,3) A'.counit_simps(1-2,4))
thus ?thesis
using A.antipar A'.antipar whisker_left [of g \<epsilon>' "f \<star> g' \<star> \<epsilon>"]
by (simp add: hcomp_assoc)
qed
also have "... = (g \<star> \<epsilon>) \<cdot> (g \<star> \<epsilon>' \<star> f \<star> g)"
proof -
have "\<epsilon>' \<star> \<epsilon> = \<epsilon> \<cdot> (\<epsilon>' \<star> f \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar hcomp_obj_arr
hcomp_arr_obj comp_arr_dom comp_cod_arr interchange obj_src trg_src
by (metis A.counit_simps(1-2,5) A'.counit_simps(1,3-4)
arr_cod not_arr_null seq_if_composable)
thus ?thesis
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar
whisker_left
by (metis A.counit_simps(1,5) A'.counit_simps(1,4) hseqI')
qed
finally show ?thesis by simp
qed
moreover have "(\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g) = (g \<star> f \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> g)"
proof -
have "(\<eta> \<star> g' \<star> f \<star> g) \<cdot> (\<eta>' \<star> g) = \<eta> \<star> \<eta>' \<star> g"
proof -
have "(\<eta> \<star> g' \<star> f) \<cdot> \<eta>' = \<eta> \<star> \<eta>'"
using A.antipar A'.antipar A.unit_in_hom hcomp_arr_obj
comp_arr_dom comp_cod_arr hcomp_obj_arr interchange
by (metis A'.unit_simps(1,3,5) A.unit_simps(1-2,4) obj_trg)
thus ?thesis
using A.antipar A'.antipar whisker_right [of g "\<eta> \<star> g' \<star> f" \<eta>']
by (simp add: hcomp_assoc)
qed
also have "... = ((g \<star> f) \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> src f \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A'.unit_in_hom comp_arr_dom comp_cod_arr interchange
by (metis A'.unit_simps(1-2,4) A.unit_simps(1,3) hseqI' ide_char)
also have "... = (g \<star> f \<star> \<eta>' \<star> g) \<cdot> (\<eta> \<star> g)"
using A.antipar A'.antipar hcomp_assoc
by (simp add: hcomp_obj_arr)
finally show ?thesis by simp
qed
ultimately show ?thesis
using comp_assoc hcomp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> ((g \<star> f) \<star> g) \<cdot> (\<eta> \<star> g)"
proof -
have "(g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g) = g \<star> f \<star> g"
proof -
have "(g \<star> (\<epsilon>' \<star> f) \<star> g) \<cdot> (g \<star> (f \<star> \<eta>') \<star> g) =
g \<star> ((\<epsilon>' \<star> f) \<star> g) \<cdot> ((f \<star> \<eta>') \<star> g)"
using A.ide_left A.ide_right A'.ide_right A.antipar A'.antipar A.unit_in_hom
A.counit_in_hom whisker_left
by (metis A'.triangle_left hseqI' ideD(1) whisker_right)
also have "... = g \<star> (\<epsilon>' \<star> f) \<cdot> (f \<star> \<eta>') \<star> g"
using A.antipar A'.antipar whisker_right [of g "\<epsilon>' \<star> f" "f \<star> \<eta>'"]
by (simp add: A'.triangle_left)
also have "... = g \<star> f \<star> g"
using A'.triangle_left by simp
finally show ?thesis by simp
qed
thus ?thesis
using hcomp_assoc by simp
qed
also have "... = (g \<star> \<epsilon>) \<cdot> (\<eta> \<star> g)"
using A.antipar A'.antipar A.unit_in_hom A.counit_in_hom comp_cod_arr
by (metis A.ide_left A.ide_right A.triangle_in_hom(2) arrI assoc_is_natural_2
ide_char seqE strict_assoc)
also have "... = g"
using A.triangle_right by simp
finally have "?\<psi> \<cdot> ?\<phi> = g" by simp
moreover have "ide g"
by simp
ultimately show ?thesis by simp
qed
qed
qed
ultimately show ?thesis
using isomorphic_def by auto
qed
end
text \<open>
We now use strictification to extend to arbitrary bicategories.
\<close>
context bicategory
begin
interpretation S: strictified_bicategory V H \<a> \<i> src trg ..
notation S.vcomp (infixr "\<cdot>\<^sub>S" 55)
notation S.hcomp (infixr "\<star>\<^sub>S" 53)
notation S.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>S _\<guillemotright>")
notation S.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
interpretation UP: equivalence_pseudofunctor V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
using S.UP_is_equivalence_pseudofunctor by auto
interpretation UP: pseudofunctor_into_strict_bicategory V H \<a> \<i> src trg
S.vcomp S.hcomp S.\<a> S.\<i> S.src S.trg S.UP S.cmp\<^sub>U\<^sub>P
..
interpretation UP: fully_faithful_functor V S.vcomp S.UP
using S.UP_is_fully_faithful_functor by auto
lemma left_adjoint_determines_right_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f g'"
shows "g \<cong> g'"
proof -
have 0: "ide g \<and> ide g'"
using assms adjoint_pair_def adjunction_in_bicategory_def
adjunction_data_in_bicategory_def adjunction_data_in_bicategory_axioms_def
by metis
have 1: "S.adjoint_pair (S.UP f) (S.UP g) \<and> S.adjoint_pair (S.UP f) (S.UP g')"
using assms UP.preserves_adjoint_pair by simp
obtain \<nu> where \<nu>: "\<guillemotleft>\<nu> : S.UP g \<Rightarrow>\<^sub>S S.UP g'\<guillemotright> \<and> S.iso \<nu>"
using 1 S.left_adjoint_determines_right_up_to_iso S.isomorphic_def by blast
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : g \<Rightarrow> g'\<guillemotright> \<and> S.UP \<mu> = \<nu>"
using 0 \<nu> UP.is_full [of g' g \<nu>] by auto
have "\<guillemotleft>\<mu> : g \<Rightarrow> g'\<guillemotright> \<and> iso \<mu>"
using \<mu> \<nu> UP.reflects_iso by auto
thus ?thesis
using isomorphic_def by auto
qed
lemma right_adjoint_determines_left_up_to_iso:
assumes "adjoint_pair f g" and "adjoint_pair f' g"
shows "f \<cong> f'"
proof -
obtain \<eta> \<epsilon> where A: "adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret A: adjunction_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using A by auto
obtain \<eta>' \<epsilon>' where A': "adjunction_in_bicategory V H \<a> \<i> src trg f' g \<eta>' \<epsilon>'"
using assms adjoint_pair_def by auto
interpret A': adjunction_in_bicategory V H \<a> \<i> src trg f' g \<eta>' \<epsilon>'
using A' by auto
interpret Cop: op_bicategory V H \<a> \<i> src trg ..
interpret Aop: adjunction_in_bicategory V Cop.H Cop.\<a> \<i> Cop.src Cop.trg g f \<eta> \<epsilon>
using A.antipar A.triangle_left A.triangle_right Cop.assoc_ide_simp
Cop.lunit_ide_simp Cop.runit_ide_simp
by (unfold_locales, auto)
interpret Aop': adjunction_in_bicategory V Cop.H Cop.\<a> \<i> Cop.src Cop.trg g f' \<eta>' \<epsilon>'
using A'.antipar A'.triangle_left A'.triangle_right Cop.assoc_ide_simp
Cop.lunit_ide_simp Cop.runit_ide_simp
by (unfold_locales, auto)
show ?thesis
using Aop.adjunction_in_bicategory_axioms Aop'.adjunction_in_bicategory_axioms
Cop.left_adjoint_determines_right_up_to_iso Cop.adjoint_pair_def
by blast
qed
end
context chosen_right_adjoints
begin
lemma isomorphic_to_left_adjoint_implies_isomorphic_right_adjoint:
assumes "is_left_adjoint f" and "f \<cong> h"
shows "f\<^sup>* \<cong> h\<^sup>*"
proof -
have 1: "adjoint_pair f f\<^sup>*"
using assms left_adjoint_extends_to_adjoint_pair by blast
moreover have "adjoint_pair h f\<^sup>*"
using assms 1 adjoint_pair_preserved_by_iso isomorphic_symmetric isomorphic_reflexive
by (meson isomorphic_def right_adjoint_simps(1))
thus ?thesis
using left_adjoint_determines_right_up_to_iso left_adjoint_extends_to_adjoint_pair
by blast
qed
end
context bicategory
begin
lemma equivalence_is_adjoint:
assumes "equivalence_map f"
shows equivalence_is_left_adjoint: "is_left_adjoint f"
and equivalence_is_right_adjoint: "is_right_adjoint f"
proof -
obtain g \<eta> \<epsilon> where fg: "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms equivalence_map_extends_to_adjoint_equivalence by blast
interpret fg: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by simp
interpret gf: adjoint_equivalence_in_bicategory V H \<a> \<i> src trg g f \<open>inv \<epsilon>\<close> \<open>inv \<eta>\<close>
using fg.dual_adjoint_equivalence by simp
show "is_left_adjoint f"
using fg.adjunction_in_bicategory_axioms adjoint_pair_def by auto
show "is_right_adjoint f"
using gf.adjunction_in_bicategory_axioms adjoint_pair_def by auto
qed
lemma right_adjoint_to_equivalence_is_equivalence:
assumes "equivalence_map f" and "adjoint_pair f g"
shows "equivalence_map g"
proof -
obtain \<eta> \<epsilon> where fg: "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret fg: adjunction_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g \<eta> \<epsilon>
using fg by simp
obtain g' \<phi> \<psi> where fg': "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g' \<phi> \<psi>"
using assms equivalence_map_def by auto
interpret fg': equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g' \<phi> \<psi>
using fg' by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g' \<phi> \<psi>'"
using assms equivalence_refines_to_adjoint_equivalence [of f g' \<phi>]
fg'.antipar fg'.unit_in_hom fg'.unit_is_iso
by auto
interpret \<psi>': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg f g' \<phi> \<psi>'
using \<psi>' by simp
have 1: "g \<cong> g'"
using fg.adjunction_in_bicategory_axioms \<psi>'.adjunction_in_bicategory_axioms
left_adjoint_determines_right_up_to_iso adjoint_pair_def
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : g' \<Rightarrow> g\<guillemotright> \<and> iso \<gamma>"
using 1 isomorphic_def isomorphic_symmetric by metis
have "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg f g ((\<gamma> \<star> f) \<cdot> \<phi>) (\<psi>' \<cdot> (f \<star> inv \<gamma>))"
using \<gamma> equivalence_preserved_by_iso_right \<psi>'.equivalence_in_bicategory_axioms by simp
hence "quasi_inverses f g"
using quasi_inverses_def by blast
thus ?thesis
using equivalence_mapI quasi_inverses_symmetric by blast
qed
lemma left_adjoint_to_equivalence_is_equivalence:
assumes "equivalence_map f" and "adjoint_pair g f"
shows "equivalence_map g"
proof -
obtain \<eta> \<epsilon> where gf: "adjunction_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g f \<eta> \<epsilon>"
using assms adjoint_pair_def by auto
interpret gf: adjunction_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g f \<eta> \<epsilon>
using gf by simp
obtain g' where 1: "quasi_inverses g' f"
using assms equivalence_mapE quasi_inverses_symmetric by blast
obtain \<phi> \<psi> where g'f: "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g' f \<phi> \<psi>"
using assms 1 quasi_inverses_def by auto
interpret g'f: equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g' f \<phi> \<psi>
using g'f by auto
obtain \<psi>' where \<psi>': "adjoint_equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g' f \<phi> \<psi>'"
using assms 1 equivalence_refines_to_adjoint_equivalence [of g' f \<phi>]
g'f.antipar g'f.unit_in_hom g'f.unit_is_iso quasi_inverses_def
equivalence_map_def
by auto
interpret \<psi>': adjoint_equivalence_in_bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg g' f \<phi> \<psi>'
using \<psi>' by simp
have 1: "g \<cong> g'"
using gf.adjunction_in_bicategory_axioms \<psi>'.adjunction_in_bicategory_axioms
right_adjoint_determines_left_up_to_iso adjoint_pair_def
by blast
obtain \<gamma> where \<gamma>: "\<guillemotleft>\<gamma> : g' \<Rightarrow> g\<guillemotright> \<and> iso \<gamma>"
using 1 isomorphic_def isomorphic_symmetric by metis
have "equivalence_in_bicategory (\<cdot>) (\<star>) \<a> \<i> src trg g f ((f \<star> \<gamma>) \<cdot> \<phi>) (\<psi>' \<cdot> (inv \<gamma> \<star> f))"
using \<gamma> equivalence_preserved_by_iso_left \<psi>'.equivalence_in_bicategory_axioms by simp
hence "quasi_inverses g f"
using quasi_inverses_def by auto
thus ?thesis
using quasi_inverses_symmetric quasi_inverses_def equivalence_map_def by blast
qed
lemma quasi_inverses_are_adjoint_pair:
assumes "quasi_inverses f g"
shows "adjoint_pair f g"
proof -
obtain \<eta> \<epsilon> where \<eta>\<epsilon>: "equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>"
using assms quasi_inverses_def by auto
interpret \<eta>\<epsilon>: equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>
using \<eta>\<epsilon> by auto
obtain \<epsilon>' where \<eta>\<epsilon>': "adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'"
using \<eta>\<epsilon> equivalence_map_def \<eta>\<epsilon>.antipar \<eta>\<epsilon>.unit_in_hom \<eta>\<epsilon>.unit_is_iso
\<eta>\<epsilon>.ide_right equivalence_refines_to_adjoint_equivalence [of f g \<eta>]
by force
interpret \<eta>\<epsilon>': adjoint_equivalence_in_bicategory V H \<a> \<i> src trg f g \<eta> \<epsilon>'
using \<eta>\<epsilon>' by auto
show ?thesis
using \<eta>\<epsilon>' adjoint_pair_def \<eta>\<epsilon>'.adjunction_in_bicategory_axioms by auto
qed
lemma quasi_inverses_isomorphic_right:
assumes "quasi_inverses f g"
shows "quasi_inverses f g' \<longleftrightarrow> g \<cong> g'"
proof
show "g \<cong> g' \<Longrightarrow> quasi_inverses f g'"
using assms quasi_inverses_def isomorphic_def equivalence_preserved_by_iso_right
by metis
assume g': "quasi_inverses f g'"
show "g \<cong> g'"
using assms g' quasi_inverses_are_adjoint_pair left_adjoint_determines_right_up_to_iso
by blast
qed
lemma quasi_inverses_isomorphic_left:
assumes "quasi_inverses f g"
shows "quasi_inverses f' g \<longleftrightarrow> f \<cong> f'"
proof
show "f \<cong> f' \<Longrightarrow> quasi_inverses f' g"
using assms quasi_inverses_def isomorphic_def equivalence_preserved_by_iso_left
by metis
assume f': "quasi_inverses f' g"
show "f \<cong> f'"
using assms f' quasi_inverses_are_adjoint_pair right_adjoint_determines_left_up_to_iso
by blast
qed
end
end
diff --git a/thys/Bicategory/Pseudofunctor.thy b/thys/Bicategory/Pseudofunctor.thy
--- a/thys/Bicategory/Pseudofunctor.thy
+++ b/thys/Bicategory/Pseudofunctor.thy
@@ -1,3138 +1,3138 @@
(* Title: Pseudofunctor
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Pseudofunctors"
theory Pseudofunctor
imports MonoidalCategory.MonoidalFunctor Bicategory Subbicategory InternalEquivalence Coherence
begin
text \<open>
The traditional definition of a pseudofunctor \<open>F : C \<rightarrow> D\<close> between bicategories \<open>C\<close> and \<open>D\<close>
is in terms of two maps: an ``object map'' \<open>F\<^sub>o\<close> that takes objects of \<open>C\<close> to objects of \<open>D\<close>
and an ``arrow map'' \<open>F\<^sub>a\<close> that assigns to each pair of objects \<open>a\<close> and \<open>b\<close> of \<open>C\<close>
a functor \<open>F\<^sub>a a b\<close> from the hom-category \<open>hom\<^sub>C a b\<close> to the hom-category \<open>hom\<^sub>D (F\<^sub>o a) (F\<^sub>o b)\<close>.
In addition, there is assigned to each object \<open>a\<close> of \<open>C\<close> an invertible 2-cell
\<open>\<guillemotleft>\<Psi> a : F\<^sub>o a \<Rightarrow>\<^sub>D (F\<^sub>a a a) a\<guillemotright>\<close>, and to each pair \<open>(f, g)\<close> of composable 1-cells of C there
is assigned an invertible 2-cell \<open>\<guillemotleft>\<Phi> (f, g) : F g \<star> F f \<Rightarrow> F (g \<star> f)\<guillemotright>\<close>, all subject to
naturality and coherence conditions.
In keeping with the ``object-free'' style in which we have been working, we do not wish
to adopt a definition of pseudofunctor that distinguishes between objects and other
arrows. Instead, we would like to understand a pseudofunctor as an ordinary functor between
(vertical) categories that weakly preserves horizontal composition in a suitable sense.
So, we take as a starting point that a pseudofunctor \<open>F : C \<rightarrow> D\<close> is a functor from
\<open>C\<close> to \<open>D\<close>, when these are regarded as ordinary categories with respect to vertical
composition. Next, \<open>F\<close> should preserve source and target, but only ``weakly''
(up to isomorphism, rather than ``on the nose'').
Weak preservation of horizontal composition is expressed by specifying, for each horizontally
composable pair of vertical identities \<open>(f, g)\<close> of \<open>C\<close>, a ``compositor''
\<open>\<guillemotleft>\<Phi> (f, g) : F g \<star> F f \<Rightarrow> F (g \<star> f)\<guillemotright>\<close> in \<open>D\<close>, such that the \<open>\<Phi> (f, g)\<close> are the components
of a natural isomorphism.
Associators must also be weakly preserved by F; this is expressed by a coherence
condition that relates an associator \<open>\<a>\<^sub>C[f, g, h]\<close> in \<open>C\<close>, its image \<open>F \<a>\<^sub>C[f, g, h]\<close>,
the associator \<open>\<a>\<^sub>D[F f, F g, F h]\<close> in \<open>D\<close> and compositors involving \<open>f\<close>, \<open>g\<close>, and \<open>h\<close>.
As regards the weak preservation of unitors, just as for monoidal functors,
which are in fact pseudofunctors between one-object bicategories, it is only necessary
to assume that \<open>F \<i>\<^sub>C[a]\<close> and \<open>\<i>\<^sub>D[F a]\<close> are isomorphic in \<open>D\<close> for each object \<open>a\<close> of \<open>C\<close>,
for there is then a canonical way to obtain, for each \<open>a\<close>, an isomorphism
\<open>\<guillemotleft>\<Psi> a : src (F a) \<rightarrow> F a\<guillemotright>\<close> that satisfies the usual coherence conditions relating the
unitors and the associators. Note that the map \<open>a \<mapsto> src (F a)\<close> amounts to the traditional
``object map'' \<open>F\<^sub>o\<close>, so that this becomes a derived notion, rather than a primitive one.
\<close>
subsection "Weak Arrows of Homs"
text \<open>
We begin with a locale that defines a functor between ``horizontal homs'' that preserves
source and target up to isomorphism.
\<close>
locale weak_arrow_of_homs =
C: horizontal_homs C src\<^sub>C trg\<^sub>C +
D: horizontal_homs D src\<^sub>D trg\<^sub>D +
"functor" C D F
for C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd" +
assumes weakly_preserves_src: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D.isomorphic (F (src\<^sub>C \<mu>)) (src\<^sub>D (F \<mu>))"
and weakly_preserves_trg: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D.isomorphic (F (trg\<^sub>C \<mu>)) (trg\<^sub>D (F \<mu>))"
begin
lemma isomorphic_src:
assumes "C.obj a"
shows "D.isomorphic (src\<^sub>D (F a)) (F a)"
using assms weakly_preserves_src [of a] D.isomorphic_symmetric by auto
lemma isomorphic_trg:
assumes "C.obj a"
shows "D.isomorphic (trg\<^sub>D (F a)) (F a)"
using assms weakly_preserves_trg [of a] D.isomorphic_symmetric by auto
abbreviation (input) hseq\<^sub>C
where "hseq\<^sub>C \<mu> \<nu> \<equiv> C.arr \<mu> \<and> C.arr \<nu> \<and> src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
abbreviation (input) hseq\<^sub>D
where "hseq\<^sub>D \<mu> \<nu> \<equiv> D.arr \<mu> \<and> D.arr \<nu> \<and> src\<^sub>D \<mu> = trg\<^sub>D \<nu>"
lemma preserves_hseq:
assumes "hseq\<^sub>C \<mu> \<nu>"
shows "hseq\<^sub>D (F \<mu>) (F \<nu>)"
proof -
have "src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
using assms by auto
hence "D.isomorphic (F (src\<^sub>C \<mu>)) (trg\<^sub>D (F \<nu>))"
using assms weakly_preserves_trg by auto
moreover have "D.isomorphic (src\<^sub>D (F \<mu>)) (F (src\<^sub>C \<mu>))"
using assms weakly_preserves_src D.isomorphic_symmetric by blast
ultimately have "D.isomorphic (src\<^sub>D (F \<mu>)) (trg\<^sub>D (F \<nu>))"
using D.isomorphic_transitive by blast
hence "src\<^sub>D (F \<mu>) = trg\<^sub>D (F \<nu>)"
using assms D.isomorphic_objects_are_equal by auto
thus ?thesis
using assms by auto
qed
text \<open>
Though \<open>F\<close> does not preserve objects ``on the nose'', we can recover from it the
usual ``object map'', which does.
It is slightly confusing at first to get used to the idea that applying the
object map of a weak arrow of homs to an object does not give the same thing
as applying the underlying functor, but rather only something isomorphic to it.
The following defines the object map associated with \<open>F\<close>.
\<close>
definition map\<^sub>0
where "map\<^sub>0 a \<equiv> src\<^sub>D (F a)"
lemma map\<^sub>0_simps [simp]:
assumes "C.obj a"
shows "D.obj (map\<^sub>0 a)"
and "src\<^sub>D (map\<^sub>0 a) = map\<^sub>0 a" and "trg\<^sub>D (map\<^sub>0 a) = map\<^sub>0 a"
and "D.dom (map\<^sub>0 a) = map\<^sub>0 a" and "D.cod (map\<^sub>0 a) = map\<^sub>0 a"
using assms map\<^sub>0_def by auto
lemma preserves_src [simp]:
assumes "C.arr \<mu>"
shows "src\<^sub>D (F \<mu>) = map\<^sub>0 (src\<^sub>C \<mu>)"
using assms
by (metis C.src.preserves_arr C.src_src C.trg_src map\<^sub>0_def preserves_hseq)
lemma preserves_trg [simp]:
assumes "C.arr \<mu>"
shows "trg\<^sub>D (F \<mu>) = map\<^sub>0 (trg\<^sub>C \<mu>)"
using assms map\<^sub>0_def preserves_hseq C.src_trg C.trg.preserves_arr by presburger
lemma preserves_hhom [intro]:
assumes "C.arr \<mu>"
shows "D.in_hhom (F \<mu>) (map\<^sub>0 (src\<^sub>C \<mu>)) (map\<^sub>0 (trg\<^sub>C \<mu>))"
using assms by simp
text \<open>
We define here the lifting of \<open>F\<close> to a functor \<open>FF: CC \<rightarrow> DD\<close>.
We need this to define the domains and codomains of the compositors.
\<close>
definition FF
where "FF \<equiv> \<lambda>\<mu>\<nu>. if C.VV.arr \<mu>\<nu> then (F (fst \<mu>\<nu>), F (snd \<mu>\<nu>)) else D.VV.null"
sublocale FF: "functor" C.VV.comp D.VV.comp FF
proof -
have 1: "\<And>\<mu>\<nu>. C.VV.arr \<mu>\<nu> \<Longrightarrow> D.VV.arr (FF \<mu>\<nu>)"
unfolding FF_def using C.VV.arr_char D.VV.arr_char preserves_hseq by simp
show "functor C.VV.comp D.VV.comp FF"
proof
fix \<mu>\<nu>
show "\<not> C.VV.arr \<mu>\<nu> \<Longrightarrow> FF \<mu>\<nu> = D.VV.null"
using FF_def by simp
show "C.VV.arr \<mu>\<nu> \<Longrightarrow> D.VV.arr (FF \<mu>\<nu>)"
using 1 by simp
assume \<mu>\<nu>: "C.VV.arr \<mu>\<nu>"
show "D.VV.dom (FF \<mu>\<nu>) = FF (C.VV.dom \<mu>\<nu>)"
using \<mu>\<nu> 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.dom_simp D.VV.dom_simp
by simp
show "D.VV.cod (FF \<mu>\<nu>) = FF (C.VV.cod \<mu>\<nu>)"
using \<mu>\<nu> 1 FF_def C.VV.arr_char D.VV.arr_char C.VV.cod_simp D.VV.cod_simp
by simp
next
fix \<mu>\<nu> \<tau>\<pi>
assume 2: "C.VV.seq \<mu>\<nu> \<tau>\<pi>"
show "FF (C.VV.comp \<mu>\<nu> \<tau>\<pi>) = D.VV.comp (FF \<mu>\<nu>) (FF \<tau>\<pi>)"
proof -
have "FF (C.VV.comp \<mu>\<nu> \<tau>\<pi>) = (F (fst \<mu>\<nu>) \<cdot>\<^sub>D F (fst \<tau>\<pi>), F (snd \<mu>\<nu>) \<cdot>\<^sub>D F (snd \<tau>\<pi>))"
using 1 2 FF_def C.VV.comp_char C.VxV.comp_char C.VV.arr_char
by (metis (no_types, lifting) C.VV.seq_char C.VxV.seqE fst_conv preserves_comp_2
snd_conv)
also have "... = D.VV.comp (FF \<mu>\<nu>) (FF \<tau>\<pi>)"
using 1 2 FF_def D.VV.comp_char D.VxV.comp_char C.VV.arr_char D.VV.arr_char
C.VV.seq_char C.VxV.seqE preserves_seq
by (simp, meson)
finally show ?thesis by simp
qed
qed
qed
lemma functor_FF:
shows "functor C.VV.comp D.VV.comp FF"
..
end
subsection "Definition of Pseudofunctors"
text \<open>
I don't much like the term "pseudofunctor", which is suggestive of something that
is ``not really'' a functor. In the development here we can see that a pseudofunctor
is really a \emph{bona fide} functor with respect to vertical composition,
which happens to have in addition a weak preservation property with respect to
horizontal composition.
This weak preservation of horizontal composition is captured by extra structure,
the ``compositors'', which are the components of a natural transformation.
So ``pseudofunctor'' is really a misnomer; it's an actual functor that has been equipped
with additional structure relating to horizontal composition. I would use the term
``bifunctor'' for such a thing, but it seems to not be generally accepted and also tends
to conflict with the usage of that term to refer to an ordinary functor of two
arguments; which I have called a ``binary functor''. Sadly, there seem to be no other
plausible choices of terminology, other than simply ``functor''
(recommended on n-Lab @{url \<open>https://ncatlab.org/nlab/show/pseudofunctor\<close>}),
but that is not workable here because we need a name that does not clash with that
used for an ordinary functor between categories.
\<close>
locale pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
weak_arrow_of_homs V\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D src\<^sub>D trg\<^sub>D F +
FoH\<^sub>C: composite_functor C.VV.comp V\<^sub>C V\<^sub>D \<open>\<lambda>\<mu>\<nu>. H\<^sub>C (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> F +
H\<^sub>DoFF: composite_functor C.VV.comp D.VV.comp V\<^sub>D
FF \<open>\<lambda>\<mu>\<nu>. H\<^sub>D (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> +
\<Phi>: natural_isomorphism C.VV.comp V\<^sub>D H\<^sub>DoFF.map FoH\<^sub>C.map \<Phi>
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd" +
assumes assoc_coherence:
"\<lbrakk> C.ide f; C.ide g; C.ide h; src\<^sub>C f = trg\<^sub>C g; src\<^sub>C g = trg\<^sub>C h \<rbrakk> \<Longrightarrow>
F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
begin
no_notation C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
no_notation D.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
notation C.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>C _\<guillemotright>")
notation C.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>C _\<guillemotright>")
notation D.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>D _\<guillemotright>")
notation D.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>D _\<guillemotright>")
notation C.lunit ("\<l>\<^sub>C[_]")
notation C.runit ("\<r>\<^sub>C[_]")
notation C.lunit' ("\<l>\<^sub>C\<^sup>-\<^sup>1[_]")
notation C.runit' ("\<r>\<^sub>C\<^sup>-\<^sup>1[_]")
notation C.\<a>' ("\<a>\<^sub>C\<^sup>-\<^sup>1[_, _, _]")
notation D.lunit ("\<l>\<^sub>D[_]")
notation D.runit ("\<r>\<^sub>D[_]")
notation D.lunit' ("\<l>\<^sub>D\<^sup>-\<^sup>1[_]")
notation D.runit' ("\<r>\<^sub>D\<^sup>-\<^sup>1[_]")
notation D.\<a>' ("\<a>\<^sub>D\<^sup>-\<^sup>1[_, _, _]")
lemma weakly_preserves_objects:
assumes "C.obj a"
shows "D.isomorphic (map\<^sub>0 a) (F a)"
using assms weakly_preserves_src [of a] D.isomorphic_symmetric by auto
lemma cmp_in_hom [intro]:
assumes "C.ide a" and "C.ide b" and "src\<^sub>C a = trg\<^sub>C b"
shows "\<guillemotleft>\<Phi> (a, b) : map\<^sub>0 (src\<^sub>C b) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C a)\<guillemotright>"
and "\<guillemotleft>\<Phi> (a, b) : F a \<star>\<^sub>D F b \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C b)\<guillemotright>"
proof -
show "\<guillemotleft>\<Phi> (a, b) : F a \<star>\<^sub>D F b \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C b)\<guillemotright>"
proof -
have "H\<^sub>DoFF.map (a, b) = F a \<star>\<^sub>D F b"
using assms C.VV.ide_char FF_def by auto
moreover have "FoH\<^sub>C.map (a, b) = F (a \<star>\<^sub>C b)"
using assms C.VV.ide_char by simp
ultimately show ?thesis
using assms C.VV.ide_char \<Phi>.preserves_hom
apply simp
by (metis (no_types, lifting) C.VV.ideI C.VV.ide_in_hom C.VxV.ide_Ide C.ideD(1)
fst_conv snd_conv)
qed
show "\<guillemotleft>\<Phi> (a, b) : map\<^sub>0 (src\<^sub>C b) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C a)\<guillemotright>"
proof -
have "C.hseq a b"
by (simp add: assms(1-3))
thus ?thesis
by (metis C.src_hcomp C.trg_hcomp D.in_hhom_def D.in_homE D.src_cod D.trg_cod
\<open>\<guillemotleft>\<Phi> (a, b) : F a \<star>\<^sub>D F b \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C b)\<guillemotright>\<close> preserves_src preserves_trg)
qed
qed
lemma cmp_simps [simp]:
assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g"
shows "D.arr (\<Phi> (f, g))"
and "src\<^sub>D (\<Phi> (f, g)) = src\<^sub>D (F g)" and "trg\<^sub>D (\<Phi> (f, g)) = trg\<^sub>D (F f)"
and "D.dom (\<Phi> (f, g)) = F f \<star>\<^sub>D F g" and "D.cod (\<Phi> (f, g)) = F (f \<star>\<^sub>C g)"
proof -
show "D.arr (\<Phi> (f, g))"
using assms cmp_in_hom by auto
show "src\<^sub>D (\<Phi> (f, g)) = src\<^sub>D (F g)"
using assms cmp_in_hom by auto
show "trg\<^sub>D (\<Phi> (f, g)) = trg\<^sub>D (F f)"
using assms cmp_in_hom by auto
show "D.dom (\<Phi> (f, g)) = F f \<star>\<^sub>D F g"
using assms cmp_in_hom by blast
show "D.cod (\<Phi> (f, g)) = F (f \<star>\<^sub>C g)"
using assms cmp_in_hom by blast
qed
lemma cmp_in_hom':
assumes "C.arr \<mu>" and "C.arr \<nu>" and "src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
shows "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : map\<^sub>0 (src\<^sub>C \<nu>) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>"
and "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>) \<Rightarrow>\<^sub>D F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)\<guillemotright>"
proof -
show 1: "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>) \<Rightarrow>\<^sub>D F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)\<guillemotright>"
using assms C.VV.arr_char \<Phi>.preserves_hom FF_def C.VV.dom_simp C.VV.cod_simp
by auto
show "\<guillemotleft>\<Phi> (\<mu>, \<nu>) : map\<^sub>0 (src\<^sub>C \<nu>) \<rightarrow>\<^sub>D map\<^sub>0 (trg\<^sub>C \<mu>)\<guillemotright>"
using assms 1 D.src_dom [of "\<Phi> (\<mu>, \<nu>)"] D.trg_dom [of "\<Phi> (\<mu>, \<nu>)"]
C.VV.dom_simp C.VV.cod_simp
by auto
qed
lemma cmp_simps':
assumes "C.arr \<mu>" and "C.arr \<nu>" and "src\<^sub>C \<mu> = trg\<^sub>C \<nu>"
shows "D.arr (\<Phi> (\<mu>, \<nu>))"
and "src\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (src\<^sub>C \<nu>)" and "trg\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (trg\<^sub>C \<mu>)"
and "D.dom (\<Phi> (\<mu>, \<nu>)) = F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>)"
and "D.cod (\<Phi> (\<mu>, \<nu>)) = F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)"
proof -
show "D.arr (\<Phi> (\<mu>, \<nu>))"
using assms cmp_in_hom by auto
show "src\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (src\<^sub>C \<nu>)"
using assms cmp_in_hom' by auto
show "trg\<^sub>D (\<Phi> (\<mu>, \<nu>)) = map\<^sub>0 (trg\<^sub>C \<mu>)"
using assms cmp_in_hom' by auto
show "D.dom (\<Phi> (\<mu>, \<nu>)) = F (C.dom \<mu>) \<star>\<^sub>D F (C.dom \<nu>)"
using assms cmp_in_hom' by blast
show "D.cod (\<Phi> (\<mu>, \<nu>)) = F (C.cod \<mu> \<star>\<^sub>C C.cod \<nu>)"
using assms cmp_in_hom' by blast
qed
lemma cmp_components_are_iso [simp]:
assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g"
shows "D.iso (\<Phi> (f, g))"
using assms C.VV.ide_char C.VV.arr_char by simp
lemma weakly_preserves_hcomp:
assumes "C.ide f" and "C.ide g" and "src\<^sub>C f = trg\<^sub>C g"
shows "D.isomorphic (F f \<star>\<^sub>D F g) (F (f \<star>\<^sub>C g))"
using assms D.isomorphic_def by auto
end
context pseudofunctor
begin
text \<open>
The following defines the image of the unit isomorphism \<open>\<i>\<^sub>C[a]\<close> under \<open>F\<close>.
We will use \<open>(F a, \<i>[a])\<close> as an ``alternate unit'', to substitute for
\<open>(src\<^sub>D (F a), \<i>\<^sub>D[src\<^sub>D (F a)])\<close>.
\<close>
abbreviation (input) \<i> ("\<i>[_]")
where "\<i>[a] \<equiv> F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a)"
lemma \<i>_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>\<i>[a] : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F a\<guillemotright>"
proof (unfold map\<^sub>0_def)
show "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F a\<guillemotright>"
using assms preserves_hom cmp_in_hom
by (intro D.comp_in_homI, auto)
show "\<guillemotleft>F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a) : src\<^sub>D (F a) \<rightarrow>\<^sub>D src\<^sub>D (F a)\<guillemotright>"
using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
by (intro D.vcomp_in_hhom D.seqI, auto)
qed
lemma \<i>_simps [simp]:
assumes "C.obj a"
shows "D.arr (\<i> a)"
and "src\<^sub>D \<i>[a] = map\<^sub>0 a" and "trg\<^sub>D \<i>[a] = map\<^sub>0 a"
and "D.dom \<i>[a] = F a \<star>\<^sub>D F a" and "D.cod \<i>[a] = F a"
proof -
show "src\<^sub>D \<i>[a] = map\<^sub>0 a"
unfolding map\<^sub>0_def
using assms \<i>_in_hom D.src_cod [of "F a"]
by (metis C.unit_simps(1) C.unit_simps(5) D.arrI D.src_vcomp D.vseq_implies_hpar(1)
is_natural_2 preserves_arr)
show "trg\<^sub>D \<i>[a] = map\<^sub>0 a"
unfolding map\<^sub>0_def
using assms \<i>_in_hom D.trg_cod [of "F a"]
by (metis C.obj_def C.unit_simps(1) C.unit_simps(3) D.arrI D.trg_vcomp preserves_hseq)
show "D.arr \<i>[a]"
using assms \<i>_in_hom by auto
show "D.dom \<i>[a] = F a \<star>\<^sub>D F a"
using assms \<i>_in_hom by auto
show "D.cod \<i>[a] = F a"
using assms \<i>_in_hom by auto
qed
lemma iso_\<i>:
assumes "C.obj a"
shows "D.iso \<i>[a]"
proof -
have "D.iso (\<Phi> (a, a))"
using assms by auto
moreover have "D.iso (F \<i>\<^sub>C[a])"
using assms C.iso_unit by simp
moreover have "D.seq (F \<i>\<^sub>C[a]) (\<Phi> (a, a))"
using assms C.obj_self_composable(1) C.seq_if_composable by simp
ultimately show ?thesis by auto
qed
text \<open>
If \<open>a\<close> is an object of \<open>C\<close> and we have an isomorphism \<open>\<guillemotleft>\<Phi> (a, a) : F a \<star>\<^sub>D F a \<Rightarrow>\<^sub>D F (a \<star>\<^sub>C a)\<guillemotright>\<close>,
then there is a canonical way to define a compatible isomorphism \<open>\<guillemotleft>\<Psi> a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>\<close>.
Specifically, we take \<open>\<Psi> a\<close> to be the unique isomorphism \<open>\<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>\<close> such that
\<open>\<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)\<close>.
\<close>
definition unit
where "unit a \<equiv> THE \<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and>
\<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
lemma unit_char:
assumes "C.obj a"
shows "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" and "D.iso (unit a)"
and "unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (unit a \<star>\<^sub>D unit a)"
and "\<exists>!\<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and> \<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
proof -
let ?P = "\<lambda>\<psi>. \<guillemotleft>\<psi> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<psi> \<and> \<psi> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<psi> \<star>\<^sub>D \<psi>)"
show "\<exists>!\<psi>. ?P \<psi>"
proof -
have "D.obj (map\<^sub>0 a)"
using assms by simp
moreover have "D.isomorphic (map\<^sub>0 a) (F a)"
unfolding map\<^sub>0_def
using assms isomorphic_src by simp
ultimately show ?thesis
using assms D.unit_unique_upto_unique_iso \<Phi>.preserves_hom \<i>_in_hom iso_\<i> by simp
qed
hence 1: "?P (unit a)"
using assms unit_def the1I2 [of ?P ?P] by simp
show "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" using 1 by simp
show "D.iso (unit a)" using 1 by simp
show "unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (unit a \<star>\<^sub>D unit a)"
using 1 by simp
qed
lemma unit_simps [simp]:
assumes "C.obj a"
shows "D.arr (unit a)"
and "src\<^sub>D (unit a) = map\<^sub>0 a" and "trg\<^sub>D (unit a) = map\<^sub>0 a"
and "D.dom (unit a) = map\<^sub>0 a" and "D.cod (unit a) = F a"
proof -
show "D.arr (unit a)"
using assms unit_char(1) by auto
show 1: "D.dom (unit a) = map\<^sub>0 a"
using assms unit_char(1) by auto
show 2: "D.cod (unit a) = F a"
using assms unit_char(1) by auto
show "src\<^sub>D (unit a) = map\<^sub>0 a"
using assms 1 D.src_dom
unfolding map\<^sub>0_def
by (metis C.obj_def D.arr_dom_iff_arr D.src.preserves_reflects_arr D.src_src preserves_arr)
show "trg\<^sub>D (unit a) = map\<^sub>0 a"
unfolding map\<^sub>0_def
using assms 2 unit_char
by (metis "1" D.trg_dom map\<^sub>0_def map\<^sub>0_simps(3) \<open>D.arr (unit a)\<close>)
qed
lemma unit_in_hom [intro]:
assumes "C.obj a"
shows "\<guillemotleft>unit a : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a\<guillemotright>"
and "\<guillemotleft>unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>"
using assms by auto
lemma unit_eqI:
assumes "C.obj a" and "\<guillemotleft>\<mu>: map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright>" and "D.iso \<mu>"
and "\<mu> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i> a \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D \<mu>)"
shows "\<mu> = unit a"
using assms unit_def unit_char
the1_equality [of "\<lambda>\<mu>. \<guillemotleft>\<mu> : map\<^sub>0 a \<Rightarrow>\<^sub>D F a\<guillemotright> \<and> D.iso \<mu> \<and>
\<mu> \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = \<i>[a] \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D \<mu>)" \<mu>]
by simp
text \<open>
The following defines the unique isomorphism satisfying the characteristic conditions
for the left unitor \<open>\<l>\<^sub>D[trg\<^sub>D (F f)]\<close>, but using the ``alternate unit'' \<open>\<i>[trg\<^sub>C f]\<close>
instead of \<open>\<i>\<^sub>D[trg\<^sub>D (F f)]\<close>, which is used to define \<open>\<l>\<^sub>D[trg\<^sub>D (F f)]\<close>.
\<close>
definition lF
where "lF f \<equiv> THE \<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> =(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
lemma lF_char:
assumes "C.ide f"
shows "\<guillemotleft>lF f : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright>"
and "F (trg\<^sub>C f) \<star>\<^sub>D lF f = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
and "\<exists>!\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
proof -
let ?P = "\<lambda>\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<and>
F (trg\<^sub>C f) \<star>\<^sub>D \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
show "\<exists>!\<mu>. ?P \<mu>"
proof -
interpret Df: prebicategory \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D
using D.is_prebicategory by simp
interpret S: subcategory \<open>(\<cdot>\<^sub>D)\<close> \<open>Df.left (F (trg\<^sub>C f))\<close>
using assms Df.left_hom_is_subcategory by simp
interpret Df: left_hom \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<open>F (trg\<^sub>C f)\<close>
using assms D.weak_unit_char
apply unfold_locales by simp
interpret Df: left_hom_with_unit \<open>(\<cdot>\<^sub>D)\<close> \<open>(\<star>\<^sub>D)\<close> \<a>\<^sub>D \<open>\<i>[trg\<^sub>C f]\<close> \<open>F (trg\<^sub>C f)\<close>
proof
show "Df.weak_unit (F (trg\<^sub>C f))"
using assms D.weak_unit_char
by (metis C.ideD(1) C.trg.preserves_reflects_arr C.trg_trg weakly_preserves_trg)
show "\<guillemotleft>\<i>[trg\<^sub>C f] : F (trg\<^sub>C f) \<star>\<^sub>D F (trg\<^sub>C f) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using assms \<i>_in_hom by simp
show "D.iso \<i>[trg\<^sub>C f]"
using assms iso_\<i> by simp
qed
have "\<exists>!\<mu>. \<guillemotleft>\<mu> : Df.L (F f) \<Rightarrow>\<^sub>S F f\<guillemotright> \<and>
Df.L \<mu> = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>S \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
proof -
have "Df.left (F (trg\<^sub>C f)) (F f)"
using assms weakly_preserves_src D.isomorphic_def D.hseq_char D.hseq_char'
Df.left_def
by fastforce
thus ?thesis
using assms Df.lunit_char(3) S.ide_char S.arr_char by simp
qed
moreover have "Df.L (F f) = F (trg\<^sub>C f) \<star>\<^sub>D F f"
using assms by (simp add: Df.H\<^sub>L_def)
moreover have "\<And>\<mu>. Df.L \<mu> = F (trg\<^sub>C f) \<star>\<^sub>D \<mu>"
using Df.H\<^sub>L_def by simp
moreover have "(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>S \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] =
(\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
by (metis (no_types, lifting) D.assoc'_eq_inv_assoc D.ext D.hseqE D.seqE
D.vconn_implies_hpar(1) D.vconn_implies_hpar(3) Df.characteristic_iso(4)
Df.equivalence_functor_L Df.iso_unit(2) S.comp_simp S.ext S.ide_char
S.in_hom_char S.iso_is_arr S.null_char calculation(1) category.ide_cod
category.in_homE equivalence_functor_def)
moreover have "\<And>\<mu>. \<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright> \<longleftrightarrow>
\<guillemotleft>\<mu> : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>S F f\<guillemotright>"
using assms S.in_hom_char S.arr_char
by (metis D.in_homE Df.hom_connected(2) Df.left_def calculation(1) calculation(2))
ultimately show ?thesis by simp
qed
hence 1: "?P (lF f)"
using lF_def the1I2 [of ?P ?P] by simp
show "\<guillemotleft>lF f : F (trg\<^sub>C f) \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F f\<guillemotright>"
using 1 by simp
show "F (trg\<^sub>C f) \<star>\<^sub>D lF f = (\<i>[trg\<^sub>C f] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F (trg\<^sub>C f), F (trg\<^sub>C f), F f]"
using 1 by simp
qed
lemma lF_simps [simp]:
assumes "C.ide f"
shows "D.arr (lF f)"
and "src\<^sub>D (lF f) = map\<^sub>0 (src\<^sub>C f)" and "trg\<^sub>D (lF f) = map\<^sub>0 (trg\<^sub>C f)"
and "D.dom (lF f) = F (trg\<^sub>C f) \<star>\<^sub>D F f" and "D.cod (lF f) = F f"
proof -
show "D.arr (lF f)"
using assms lF_char(1) by auto
show "D.dom (lF f) = F (trg\<^sub>C f) \<star>\<^sub>D F f"
using assms lF_char(1) by auto
show "D.cod (lF f) = F f"
using assms lF_char(1) by auto
show "src\<^sub>D (lF f) = map\<^sub>0 (src\<^sub>C f)"
unfolding map\<^sub>0_def
using assms \<open>D.arr (lF f)\<close> \<open>D.cod (lF f) = F f\<close> D.src_cod by fastforce
show "trg\<^sub>D (lF f) = map\<^sub>0 (trg\<^sub>C f)"
using assms \<open>D.arr (lF f)\<close> \<open>D.cod (lF f) = F f\<close> D.trg_cod by fastforce
qed
text \<open>
\sloppypar
The next two lemmas generalize the eponymous results from
@{theory MonoidalCategory.MonoidalFunctor}. See the proofs of those results for diagrams.
\<close>
lemma lunit_coherence1:
assumes "C.ide f"
shows "\<l>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (unit (trg\<^sub>C f) \<star>\<^sub>D F f) = lF f"
proof -
let ?b = "trg\<^sub>C f"
have 1: "trg\<^sub>D (F f) = map\<^sub>0 ?b"
using assms by simp
have "lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f) = \<l>\<^sub>D[F f]"
proof -
have "D.par (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) \<l>\<^sub>D[F f]"
using assms 1 D.lunit_in_hom unit_char(1-2) lF_char(1) D.ideD(1) by auto
moreover have "map\<^sub>0 ?b \<star>\<^sub>D (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) = map\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f]"
proof -
have "map\<^sub>0 ?b \<star>\<^sub>D (lF f \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D F f)) =
(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (map\<^sub>0 ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
using assms D.objE [of "map\<^sub>0 (trg\<^sub>C f)"]
D.whisker_left [of "map\<^sub>0 ?b" "lF f" "unit ?b \<star>\<^sub>D F f"]
by auto
also have "... = (map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D
(D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
proof -
have "(D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f) =
D.inv (unit ?b) \<cdot>\<^sub>D unit ?b \<star>\<^sub>D F ?b \<cdot>\<^sub>D unit ?b \<star>\<^sub>D F f \<cdot>\<^sub>D F f"
using assms unit_char(1-2)
D.interchange [of "F ?b" "unit ?b" "F f" "F f"]
D.interchange [of "D.inv (unit ?b)" "unit ?b" "F ?b \<star>\<^sub>D F f" "unit ?b \<star>\<^sub>D F f"]
by simp
also have "... = map\<^sub>0 ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f"
using assms unit_char(1-2) [of ?b] D.comp_arr_dom D.comp_cod_arr D.comp_inv_arr
by (simp add: D.inv_is_inverse)
finally show ?thesis by simp
qed
also have "... =
(D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
proof -
have "(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) =
(D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
proof -
have "(map\<^sub>0 ?b \<star>\<^sub>D lF f) \<cdot>\<^sub>D (D.inv (unit ?b) \<star>\<^sub>D F ?b \<star>\<^sub>D F f) =
D.inv (unit ?b) \<star>\<^sub>D lF f"
using assms unit_char(1-2) lF_char(1) D.comp_arr_dom D.comp_cod_arr
D.interchange [of "map\<^sub>0 ?b" "D.inv (unit ?b)" "lF f" "F ?b \<star>\<^sub>D F f"]
by simp
also have "... = D.inv (unit ?b) \<cdot>\<^sub>D F ?b \<star>\<^sub>D F f \<cdot>\<^sub>D lF f"
using assms unit_char(1-2) lF_char(1) D.comp_arr_dom D.comp_cod_arr
D.inv_in_hom
by auto
also have "... = (D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
using assms unit_char(1-2) lF_char(1) D.inv_in_hom
D.interchange [of "D.inv (unit ?b)" "F ?b" "F f" "lF f"]
by simp
finally show ?thesis by simp
qed
thus ?thesis
using assms unit_char(1-2) D.inv_in_hom D.comp_assoc by metis
qed
also have "... = (D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f] \<cdot>\<^sub>D
(unit ?b \<star>\<^sub>D unit ?b \<star>\<^sub>D F f)"
using assms unit_char(1-2) lF_char(2) D.comp_assoc by auto
also have "... = ((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D
((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 ?b, map\<^sub>0 ?b, F f]"
using assms unit_char(1-2) D.assoc'_naturality [of "unit ?b" "unit ?b" "F f"] D.comp_assoc
by (simp add: \<open>trg\<^sub>D (F f) = map\<^sub>0 (trg\<^sub>C f)\<close>)
also have "... = (\<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[map\<^sub>0 ?b, map\<^sub>0 ?b, F f]"
proof -
have "((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D ((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) =
\<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f"
proof -
have "((D.inv (unit ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D ((unit ?b \<star>\<^sub>D unit ?b) \<star>\<^sub>D F f)) =
D.inv (unit ?b) \<cdot>\<^sub>D unit ?b \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f"
using assms 1 D.unit_in_hom D.whisker_right [of "F f"] unit_char(2-3)
D.invert_side_of_triangle(1)
by (metis C.ideD(1) C.obj_trg D.seqI' map\<^sub>0_simps(1) unit_in_hom(2) preserves_ide)
also have "... = \<i>\<^sub>D[map\<^sub>0 ?b] \<star>\<^sub>D F f"
proof -
have "(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D unit (trg\<^sub>C f)) \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 ?b] = \<i>\<^sub>D[map\<^sub>0 ?b]"
by (simp add: D.comp_cod_arr D.comp_inv_arr D.inv_is_inverse unit_char(2)
assms)
thus ?thesis
by (simp add: D.comp_assoc)
qed
finally show ?thesis by blast
qed
thus ?thesis by simp
qed
also have "... = map\<^sub>0 ?b \<star>\<^sub>D \<l>\<^sub>D[F f]"
using assms D.lunit_char [of "F f"] \<open>trg\<^sub>D (F f) = map\<^sub>0 ?b\<close> by simp
finally show ?thesis by blast
qed
ultimately show ?thesis
using assms D.L.is_faithful
by (metis D.trg_cod D.trg_vcomp D.vseq_implies_hpar(2) lF_simps(3))
qed
thus ?thesis
using assms 1 unit_char(1-2) C.ideD(1) C.obj_trg D.inverse_arrows_hcomp(1)
D.invert_side_of_triangle(2) D.lunit_simps(1) unit_simps(2) preserves_ide
D.iso_hcomp components_are_iso
by metis
qed
lemma lunit_coherence2:
assumes "C.ide f"
shows "lF f = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)"
proof -
let ?b = "trg\<^sub>C f"
have "D.par (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f)) (lF f)"
using assms cmp_in_hom [of ?b f] lF_simps by auto
moreover have "F ?b \<star>\<^sub>D F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f) = F ?b \<star>\<^sub>D lF f"
proof -
have "F ?b \<star>\<^sub>D F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (?b, f) = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D \<Phi> (?b, f))"
using assms cmp_in_hom D.whisker_left [of "F ?b" "F \<l>\<^sub>C[f]" "\<Phi> (?b, f)"]
by (simp add: calculation)
also have "... = F ?b \<star>\<^sub>D lF f"
proof -
have "(F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D \<Phi> (?b, f))
= (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D
\<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have 1: "D.seq (F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f])
(\<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D (\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms by fastforce
hence 2: "D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)
= (F ?b \<star>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F ?b, F ?b, F f]"
using assms cmp_in_hom assoc_coherence cmp_components_are_iso
D.invert_side_of_triangle(1)
[of "F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f)"
"\<Phi> (?b, ?b \<star>\<^sub>C f)"
"(F ?b \<star>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F ?b, F ?b, F f]"]
C.ideD(1) C.ide_hcomp C.trg_hcomp C.trg_trg C.src_trg C.trg.preserves_ide
by metis
hence "F ?b \<star>\<^sub>D \<Phi> (?b, f)
= (D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have "D.seq (D.inv (\<Phi> (trg\<^sub>C f, trg\<^sub>C f \<star>\<^sub>C f)))
(F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D
(\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms 1 D.hseq_char by auto
moreover have "(F (trg\<^sub>C f) \<star>\<^sub>D \<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D \<a>\<^sub>D[F (trg\<^sub>C f), F (trg\<^sub>C f), F f] =
D.inv (\<Phi> (trg\<^sub>C f, trg\<^sub>C f \<star>\<^sub>C f)) \<cdot>\<^sub>D
F \<a>\<^sub>C[trg\<^sub>C f, trg\<^sub>C f, f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f \<star>\<^sub>C trg\<^sub>C f, f) \<cdot>\<^sub>D
(\<Phi> (trg\<^sub>C f, trg\<^sub>C f) \<star>\<^sub>D F f)"
using assms 2 by simp
ultimately show ?thesis
using assms
D.invert_side_of_triangle(2)
[of "D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f)"
"F ?b \<star>\<^sub>D \<Phi> (?b, f)" "\<a>\<^sub>D[F ?b, F ?b, F f]"]
by fastforce
qed
thus ?thesis
using D.comp_assoc by simp
qed
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
(D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)) \<cdot>\<^sub>D
\<Phi> (?b \<star>\<^sub>C ?b, f) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
proof -
have 1: "F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) = F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (F \<a>\<^sub>C[?b, ?b, f])"
using assms C.lunit_char(1-2) C.unit_in_hom preserves_inv by auto
have "F \<a>\<^sub>C[?b, ?b, f] = D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)"
proof -
have "F \<a>\<^sub>C[?b, ?b, f] \<cdot>\<^sub>D D.inv (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f)) =
D.inv (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D D.inv (F \<a>\<^sub>C[?b, ?b, f]))"
using assms by (simp add: C.iso_unit D.inv_comp)
thus ?thesis
using assms 1 D.invert_side_of_triangle D.iso_inv_iso
by (metis C.iso_hcomp C.ideD(1) C.ide_is_iso C.iso_lunit C.iso_unit
C.lunit_simps(3) C.obj_trg C.src_trg C.trg.components_are_iso
C.unit_simps(2) D.arr_inv D.inv_inv preserves_iso)
qed
thus ?thesis by argo
qed
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D (F (\<i>\<^sub>C[?b] \<star>\<^sub>C f) \<cdot>\<^sub>D \<Phi> (?b \<star>\<^sub>C ?b, f)) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using D.comp_assoc by auto
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D (\<Phi> (?b, f) \<cdot>\<^sub>D (F \<i>\<^sub>C[?b] \<star>\<^sub>D F f)) \<cdot>\<^sub>D
(\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using assms \<Phi>.naturality [of "(\<i>\<^sub>C[?b], f)"] FF_def C.VV.arr_char C.VV.cod_char
C.VV.dom_char
by simp
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f) \<cdot>\<^sub>D
((F \<i>\<^sub>C[?b] \<star>\<^sub>D F f)) \<cdot>\<^sub>D (\<Phi> (?b, ?b) \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using D.comp_assoc by auto
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D
D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f) \<cdot>\<^sub>D (\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D
\<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f]"
using assms by (simp add: D.comp_assoc D.whisker_right)
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D
(D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)) \<cdot>\<^sub>D
(F ?b \<star>\<^sub>D lF f)"
proof -
have "(\<i> ?b \<star>\<^sub>D F f) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F ?b, F ?b, F f] = F ?b \<star>\<^sub>D lF f"
using assms lF_char by auto
thus ?thesis
using assms D.inv_is_inverse \<i>_in_hom cmp_in_hom D.invert_side_of_triangle(2)
by (simp add: D.comp_assoc)
qed
also have "... = (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
proof -
have "D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) =
D.inv (((F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f))) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
using assms D.comp_inv_arr D.comp_inv_arr' cmp_simps(4)
D.comp_arr_dom D.comp_assoc
by simp
also have "... = D.inv (D.inv (\<Phi> (?b, f)) \<cdot>\<^sub>D F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
proof -
have 1: "\<Phi> (?b, f) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) = F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f)"
using assms \<Phi>.naturality [of "(?b, \<l>\<^sub>C[f])"] FF_def C.VV.arr_char
C.VV.cod_char D.VV.null_char C.VV.dom_simp
by simp
have "(F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) =
D.inv (\<Phi> (?b, f)) \<cdot>\<^sub>D F (?b \<star>\<^sub>C \<l>\<^sub>C[f])"
proof -
have "D.seq (\<Phi> (?b, f)) (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])"
using assms cmp_in_hom(2) [of ?b f] by auto
moreover have "D.iso (\<Phi> (?b, f)) \<and> D.iso (\<Phi> (?b, ?b \<star>\<^sub>C f))"
using assms by simp
ultimately show ?thesis
using 1 D.invert_opposite_sides_of_square by simp
qed
thus ?thesis
using D.comp_assoc by auto
qed
also have "... = D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D \<Phi> (?b, f)"
proof -
have "D.iso (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
using assms D.isos_compose C.VV.arr_char C.iso_lunit C.VV.dom_simp
C.VV.cod_simp
by simp
moreover have "D.iso (D.inv (\<Phi> (?b, f)))"
using assms by simp
moreover have "D.seq (D.inv (\<Phi> (?b, f))) (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f))"
using assms C.VV.arr_char C.VV.dom_simp C.VV.cod_simp by simp
ultimately show ?thesis
using assms D.inv_comp by simp
qed
also have "... = D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)"
proof -
have "D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]) \<cdot>\<^sub>D \<Phi> (?b, ?b \<star>\<^sub>C f)) =
D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f]))"
using assms D.isos_compose C.VV.arr_char C.iso_lunit D.inv_comp
C.VV.dom_simp C.VV.cod_simp
by simp
thus ?thesis using D.comp_assoc by simp
qed
finally have "D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])
= D.inv (\<Phi> (?b, ?b \<star>\<^sub>C f)) \<cdot>\<^sub>D D.inv (F (?b \<star>\<^sub>C \<l>\<^sub>C[f])) \<cdot>\<^sub>D \<Phi> (?b, f)"
by blast
thus ?thesis by argo
qed
also have "... = ((F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]) \<cdot>\<^sub>D D.inv (F ?b \<star>\<^sub>D F \<l>\<^sub>C[f])) \<cdot>\<^sub>D (F ?b \<star>\<^sub>D lF f)"
using D.comp_assoc by simp
also have "... = F ?b \<star>\<^sub>D lF f"
using assms D.comp_arr_inv' [of "F ?b \<star>\<^sub>D F \<l>\<^sub>C[f]"] D.comp_cod_arr by simp
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
ultimately show ?thesis
using assms D.L.is_faithful
by (metis D.in_homI lF_char(2-3) lF_simps(4-5))
qed
lemma lunit_coherence:
assumes "C.ide f"
shows "\<l>\<^sub>D[F f] = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
proof -
have 1: "\<l>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (unit (trg\<^sub>C f) \<star>\<^sub>D F f) = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)"
using assms lunit_coherence1 lunit_coherence2 by simp
have "\<l>\<^sub>D[F f] = (F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
proof -
have "D.seq (F \<l>\<^sub>C[f]) (\<Phi> (trg\<^sub>C f, f))"
using assms cmp_in_hom [of "trg\<^sub>C f" f] C.unit_in_vhom by auto
moreover have "D.iso (D.inv (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms unit_char unit_char(1-2)
by (simp add: preserves_hseq)
ultimately show ?thesis
using assms 1 unit_char(2) cmp_in_hom D.inv_inv
D.invert_side_of_triangle(2)
[of "F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f)" "\<l>\<^sub>D[F f]" "D.inv (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"]
by auto
qed
thus ?thesis
using assms unit_char(1) D.comp_assoc by auto
qed
text \<open>
We postpone proving the dual version of this result until after we have developed
the notion of the ``op bicategory'' in the next section.
\<close>
end
subsection "Pseudofunctors and Opposite Bicategories"
text \<open>
There are three duals to a bicategory:
\begin{enumerate}
\item ``op'': sources and targets are exchanged;
\item ``co'': domains and codomains are exchanged;
\item ``co-op'': both sources and targets and domains and codomains are exchanged.
\end{enumerate}
Here we consider the "op" case.
\<close>
locale op_bicategory =
B: bicategory V H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
for V :: "'a comp" (infixr "\<cdot>" 55)
and H\<^sub>B :: "'a comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'a \<Rightarrow> 'a" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
begin
abbreviation H (infixr "\<star>" 53)
where "H f g \<equiv> H\<^sub>B g f"
abbreviation \<i> ("\<i>[_]")
where "\<i> \<equiv> \<i>\<^sub>B"
abbreviation src
where "src \<equiv> trg\<^sub>B"
abbreviation trg
where "trg \<equiv> src\<^sub>B"
interpretation horizontal_homs V src trg
by (unfold_locales, auto)
interpretation H: "functor" VV.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
using VV.arr_char VV.dom_simp VV.cod_simp
apply unfold_locales
apply (metis (no_types, lifting) B.hseqE B.hseq_char')
apply auto[3]
using VV.comp_char VV.seq_char VV.arr_char B.VxV.comp_char B.interchange
by (metis (no_types, lifting) B.VxV.seqE fst_conv snd_conv)
interpretation horizontal_composition V H src trg
by (unfold_locales, auto)
abbreviation UP
where "UP \<mu>\<nu>\<tau> \<equiv> if B.VVV.arr \<mu>\<nu>\<tau> then
(snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)
else VVV.null"
abbreviation DN
where "DN \<mu>\<nu>\<tau> \<equiv> if VVV.arr \<mu>\<nu>\<tau> then
(snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)
else B.VVV.null"
lemma VVV_arr_char:
shows "VVV.arr \<mu>\<nu>\<tau> \<longleftrightarrow> B.VVV.arr (DN \<mu>\<nu>\<tau>)"
using VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char B.VVV.not_arr_null
by auto
lemma VVV_ide_char:
shows "VVV.ide \<mu>\<nu>\<tau> \<longleftrightarrow> B.VVV.ide (DN \<mu>\<nu>\<tau>)"
proof -
have "VVV.ide \<mu>\<nu>\<tau> \<longleftrightarrow> VVV.arr \<mu>\<nu>\<tau> \<and> B.VxVxV.ide \<mu>\<nu>\<tau>"
using VVV.ide_char by simp
also have "... \<longleftrightarrow> B.VVV.arr (DN \<mu>\<nu>\<tau>) \<and> B.VxVxV.ide (DN \<mu>\<nu>\<tau>)"
using VVV_arr_char B.VxVxV.ide_char by auto
also have "... \<longleftrightarrow> B.VVV.ide (DN \<mu>\<nu>\<tau>)"
using B.VVV.ide_char [of "DN \<mu>\<nu>\<tau>"] by blast
finally show ?thesis by fast
qed
lemma VVV_dom_char:
shows "VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
proof (cases "VVV.arr \<mu>\<nu>\<tau>")
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
using VVV.dom_def VVV.has_domain_iff_arr VVV_arr_char B.VVV.dom_null
by auto
show "VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.dom \<mu>\<nu>\<tau> = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
proof -
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have "VVV.dom \<mu>\<nu>\<tau> =
(B.dom (fst \<mu>\<nu>\<tau>), B.dom (fst (snd \<mu>\<nu>\<tau>)), B.dom (snd (snd \<mu>\<nu>\<tau>)))"
using \<mu>\<nu>\<tau> VVV.dom_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
by simp
also have "... = UP (B.dom (snd (snd \<mu>\<nu>\<tau>)), B.dom (fst (snd \<mu>\<nu>\<tau>)), B.dom (fst \<mu>\<nu>\<tau>))"
using \<mu>\<nu>\<tau> VVV_arr_char B.VV.arr_char B.VVV.arr_char B.arr_dom_iff_arr src_dom
trg_dom
apply simp
by (metis (no_types, lifting) src_dom trg_dom VV.arrE VVV.arrE)
also have "... = UP (B.VVV.dom (DN \<mu>\<nu>\<tau>))"
using \<mu>\<nu>\<tau> B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
by simp
finally show ?thesis by blast
qed
qed
lemma VVV_cod_char:
shows "VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
proof (cases "VVV.arr \<mu>\<nu>\<tau>")
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
using VVV.cod_def VVV.has_codomain_iff_arr VVV_arr_char B.VVV.cod_null
by auto
show "VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> VVV.cod \<mu>\<nu>\<tau> = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
proof -
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have "VVV.cod \<mu>\<nu>\<tau> = (B.cod (fst \<mu>\<nu>\<tau>), B.cod (fst (snd \<mu>\<nu>\<tau>)), B.cod (snd (snd \<mu>\<nu>\<tau>)))"
using \<mu>\<nu>\<tau> VVV.cod_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
by simp
also have "... = UP (B.cod (snd (snd \<mu>\<nu>\<tau>)), B.cod (fst (snd \<mu>\<nu>\<tau>)), B.cod (fst \<mu>\<nu>\<tau>))"
using \<mu>\<nu>\<tau> VVV_arr_char B.VV.arr_char B.VVV.arr_char
apply simp
by (metis (no_types, lifting) B.arr_cod_iff_arr src_cod trg_cod VV.arrE VVV.arrE)
also have "... = UP (B.VVV.cod (DN \<mu>\<nu>\<tau>))"
using \<mu>\<nu>\<tau> B.VVV.cod_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
by simp
finally show ?thesis by blast
qed
qed
lemma HoHV_char:
shows "HoHV \<mu>\<nu>\<tau> = B.HoVH (DN \<mu>\<nu>\<tau>)"
using HoHV_def B.HoVH_def VVV_arr_char by simp
lemma HoVH_char:
shows "HoVH \<mu>\<nu>\<tau> = B.HoHV (DN \<mu>\<nu>\<tau>)"
using HoVH_def B.HoHV_def VVV_arr_char by simp
definition \<a> ("\<a>[_, _, _]")
where "\<a>[\<mu>, \<nu>, \<tau>] \<equiv> B.\<alpha>' (DN (\<mu>, \<nu>, \<tau>))"
interpretation natural_isomorphism VVV.comp \<open>(\<cdot>)\<close> HoHV HoVH
\<open>\<lambda>\<mu>\<nu>\<tau>. \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]\<close>
proof
fix \<mu>\<nu>\<tau>
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = B.null"
using VVV.arr_char B.VVV.arr_char \<a>_def B.\<alpha>'.is_extensional by auto
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
show "B.dom \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> \<a>_def HoHV_def B.\<alpha>'.preserves_dom VVV.arr_char VVV.dom_char VV.arr_char
B.HoVH_def B.VVV.arr_char B.VV.arr_char B.VVV.dom_char
by auto
show "B.cod \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)] = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> \<a>_def HoVH_def B.\<alpha>'.preserves_cod VVV.arr_char VVV.cod_char VV.arr_char
B.HoHV_def B.VVV.arr_char B.VV.arr_char B.VVV.cod_char
by auto
show "HoVH \<mu>\<nu>\<tau> \<cdot>
\<a>[fst (VVV.dom \<mu>\<nu>\<tau>), fst (snd (VVV.dom \<mu>\<nu>\<tau>)), snd (snd (VVV.dom \<mu>\<nu>\<tau>))] =
\<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
proof -
have "HoVH \<mu>\<nu>\<tau> \<cdot>
\<a>[fst (VVV.dom \<mu>\<nu>\<tau>), fst (snd (VVV.dom \<mu>\<nu>\<tau>)), snd (snd (VVV.dom \<mu>\<nu>\<tau>))] =
HoVH \<mu>\<nu>\<tau> \<cdot> B.\<alpha>' (B.VVV.dom (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>))"
unfolding \<a>_def
using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char VVV.dom_char B.VVV.arr_char B.VVV.dom_char
by auto
also have "... = B.\<alpha>' (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)"
using B.\<alpha>'.is_natural_1 VVV_arr_char \<mu>\<nu>\<tau> HoVH_char by presburger
also have "... = \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
using \<mu>\<nu>\<tau> \<a>_def by simp
finally show ?thesis by blast
qed
show "\<a>[fst (VVV.cod \<mu>\<nu>\<tau>), fst (snd (VVV.cod \<mu>\<nu>\<tau>)), snd (snd (VVV.cod \<mu>\<nu>\<tau>))] \<cdot>
HoHV \<mu>\<nu>\<tau> =
\<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))"
proof -
have "\<a>[fst (VVV.cod \<mu>\<nu>\<tau>), fst (snd (VVV.cod \<mu>\<nu>\<tau>)), snd (snd (VVV.cod \<mu>\<nu>\<tau>))] \<cdot>
HoHV \<mu>\<nu>\<tau> =
B.\<alpha>' (B.VVV.cod (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)) \<cdot> HoHV \<mu>\<nu>\<tau>"
unfolding \<a>_def
using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char VVV.cod_char B.VVV.arr_char B.VVV.cod_char
by auto
also have "... = B.\<alpha>' (snd (snd \<mu>\<nu>\<tau>), fst (snd \<mu>\<nu>\<tau>), fst \<mu>\<nu>\<tau>)"
using B.\<alpha>'.is_natural_2 VVV_arr_char \<mu>\<nu>\<tau> HoHV_char by presburger
also have "... = \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
using \<mu>\<nu>\<tau> \<a>_def by simp
finally show ?thesis by blast
qed
next
fix \<mu>\<nu>\<tau>
assume \<mu>\<nu>\<tau>: "VVV.ide \<mu>\<nu>\<tau>"
show "B.iso \<a>[fst \<mu>\<nu>\<tau>, fst (snd \<mu>\<nu>\<tau>), snd (snd \<mu>\<nu>\<tau>)]"
proof -
have "B.VVV.ide (DN \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV_ide_char by blast
thus ?thesis
using \<mu>\<nu>\<tau> \<a>_def B.\<alpha>'.components_are_iso by force
qed
qed
sublocale bicategory V H \<a> \<i> src trg
proof
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>\<i> a : H a a \<rightarrow>\<^sub>B a\<guillemotright>"
using obj_def objE B.obj_def B.objE B.unit_in_hom by meson
show "\<And>a. obj a \<Longrightarrow> B.iso (\<i> a)"
using obj_def objE B.obj_def B.objE B.iso_unit by meson
show "\<And>f g h k. \<lbrakk> B.ide f; B.ide g; B.ide h; B.ide k;
src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<a>[g, h, k]) \<cdot> \<a>[f, g \<star> h, k] \<cdot> (\<a>[f, g, h] \<star> k) = \<a>[f, g, h \<star> k] \<cdot> \<a>[f \<star> g, h, k]"
unfolding \<a>_def
using B.\<a>'_def B.comp_assoc B.pentagon' VVV.arr_char VV.arr_char by simp
qed
proposition is_bicategory:
shows "bicategory V H \<a> \<i> src trg"
..
lemma assoc_ide_simp:
assumes "B.ide f" and "B.ide g" and "B.ide h"
and "src f = trg g" and "src g = trg h"
shows "\<a>[f, g, h] = B.\<a>' h g f"
using assms \<a>_def B.\<a>'_def by fastforce
lemma lunit_ide_simp:
assumes "B.ide f"
shows "lunit f = B.runit f"
proof (intro B.runit_eqI)
show "B.ide f" by fact
show "\<guillemotleft>lunit f : H (trg f) f \<rightarrow>\<^sub>B f\<guillemotright>"
using assms by simp
show "trg f \<star> lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>\<^sub>B[f, trg f, trg f]"
proof -
have "trg f \<star> lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
using assms lunit_char(1-2) [of f] by simp
moreover have "\<a>' (trg f) (trg f) f = \<a>\<^sub>B[f, trg f, trg f]"
proof (unfold \<a>'_def)
have 1: "VVV.ide (trg f, trg f, f)"
using assms VVV.ide_char VVV.arr_char VV.arr_char by simp
have "\<alpha>' (trg f, trg f, f) = B.inv \<a>[trg f, trg f, f]"
using 1 B.\<alpha>'.inverts_components by simp
also have "... = B.inv (B.\<alpha>' (f, trg f, trg f))"
unfolding \<a>_def using 1 by simp
also have "... = \<a>\<^sub>B[f, trg f, trg f]"
using 1 B.VVV.ide_char B.VVV.arr_char B.VV.arr_char VVV.ide_char
VVV.arr_char VV.arr_char B.\<alpha>'.inverts_components B.\<alpha>_def
by force
finally show "\<alpha>' (trg f, trg f, f) = \<a>\<^sub>B[f, trg f, trg f]"
by blast
qed
ultimately show ?thesis by simp
qed
qed
lemma runit_ide_simp:
assumes "B.ide f"
shows "runit f = B.lunit f"
using assms runit_char(1-2) [of f] B.\<a>'_def assoc_ide_simp
by (intro B.lunit_eqI) auto
end
context pseudofunctor
begin
interpretation C': op_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C ..
interpretation D': op_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D ..
notation C'.H (infixr "\<star>\<^sub>C\<^sup>o\<^sup>p" 53)
notation D'.H (infixr "\<star>\<^sub>D\<^sup>o\<^sup>p" 53)
interpretation F': weak_arrow_of_homs V\<^sub>C C'.src C'.trg V\<^sub>D D'.src D'.trg F
apply unfold_locales
using weakly_preserves_src weakly_preserves_trg by simp_all
interpretation H\<^sub>D'oFF: composite_functor C'.VV.comp D'.VV.comp V\<^sub>D F'.FF
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D\<^sup>o\<^sup>p snd \<mu>\<nu>\<close> ..
interpretation FoH\<^sub>C': composite_functor C'.VV.comp V\<^sub>C V\<^sub>D
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>C\<^sup>o\<^sup>p snd \<mu>\<nu>\<close> F
..
interpretation \<Phi>': natural_isomorphism C'.VV.comp V\<^sub>D H\<^sub>D'oFF.map FoH\<^sub>C'.map
\<open>\<lambda>f. \<Phi> (snd f, fst f)\<close>
using C.VV.arr_char C'.VV.arr_char C'.VV.ide_char C.VV.ide_char FF_def F'.FF_def
\<Phi>.is_extensional \<Phi>.is_natural_1 \<Phi>.is_natural_2
C.VV.dom_simp C.VV.cod_simp C'.VV.dom_simp C'.VV.cod_simp
apply unfold_locales by auto
interpretation F': pseudofunctor V\<^sub>C C'.H C'.\<a> \<i>\<^sub>C C'.src C'.trg
V\<^sub>D D'.H D'.\<a> \<i>\<^sub>D D'.src D'.trg
F \<open>\<lambda>f. \<Phi> (snd f, fst f)\<close>
proof
fix f g h
assume f: "C.ide f" and g: "C.ide g" and h: "C.ide h"
and fg: "C'.src f = C'.trg g" and gh: "C'.src g = C'.trg h"
show "F (C'.\<a> f g h) \<cdot>\<^sub>D \<Phi> (snd (f \<star>\<^sub>C\<^sup>o\<^sup>p g, h), fst (f \<star>\<^sub>C\<^sup>o\<^sup>p g, h)) \<cdot>\<^sub>D
(\<Phi> (snd (f, g), fst (f, g)) \<star>\<^sub>D\<^sup>o\<^sup>p F h) =
\<Phi> (snd (f, g \<star>\<^sub>C\<^sup>o\<^sup>p h), fst (f, g \<star>\<^sub>C\<^sup>o\<^sup>p h)) \<cdot>\<^sub>D (F f \<star>\<^sub>D\<^sup>o\<^sup>p \<Phi> (snd (g, h), fst (g, h))) \<cdot>\<^sub>D
D'.\<a> (F f) (F g) (F h)"
using f g h fg gh C.VV.in_hom_char FF_def C.VV.arr_char C.VV.dom_simp C.VV.cod_simp
C'.assoc_ide_simp D'.assoc_ide_simp preserves_inv assoc_coherence
D.invert_opposite_sides_of_square
[of "F (\<a>\<^sub>C h g f)" "\<Phi> (g \<star>\<^sub>C\<^sup>o\<^sup>p h, f) \<cdot>\<^sub>D (F f \<star>\<^sub>D\<^sup>o\<^sup>p \<Phi> (h, g))"
"\<Phi> (h, f \<star>\<^sub>C\<^sup>o\<^sup>p g) \<cdot>\<^sub>D (\<Phi> (g, f) \<star>\<^sub>D\<^sup>o\<^sup>p F h)" "\<a>\<^sub>D (F h) (F g) (F f)"]
D.comp_assoc
by auto
qed
lemma induces_pseudofunctor_between_opposites:
shows "pseudofunctor (\<cdot>\<^sub>C) (\<star>\<^sub>C\<^sup>o\<^sup>p) C'.\<a> \<i>\<^sub>C C'.src C'.trg
(\<cdot>\<^sub>D) (\<star>\<^sub>D\<^sup>o\<^sup>p) D'.\<a> \<i>\<^sub>D D'.src D'.trg
F (\<lambda>f. \<Phi> (snd f, fst f))"
..
text \<open>
It is now easy to dualize the coherence condition for \<open>F\<close> with respect to
left unitors to obtain the corresponding condition for right unitors.
\<close>
lemma runit_coherence:
assumes "C.ide f"
shows "\<r>\<^sub>D[F f] = F \<r>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
proof -
have "C'.lunit f = \<r>\<^sub>C[f]"
using assms C'.lunit_ide_simp by simp
moreover have "D'.lunit (F f) = \<r>\<^sub>D[F f]"
using assms D'.lunit_ide_simp by simp
moreover have "F'.unit (src\<^sub>C f) = unit (src\<^sub>C f)"
using assms F'.unit_char F'.preserves_trg
by (intro unit_eqI, auto)
moreover have "D'.lunit (F f) =
F (C'.lunit f) \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D F'.unit (src\<^sub>C f))"
using assms F'.lunit_coherence by simp
ultimately show ?thesis by simp
qed
end
subsection "Preservation Properties"
text \<open>
The objective of this section is to establish explicit formulas for the result
of applying a pseudofunctor to expressions of various forms.
\<close>
context pseudofunctor
begin
lemma preserves_lunit:
assumes "C.ide f"
shows "F \<l>\<^sub>C[f] = \<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
and "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
have "\<l>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) = F \<l>\<^sub>C[f]"
proof -
have "D.arr \<l>\<^sub>D[F f]"
using assms by simp
moreover have "\<l>\<^sub>D[F f] = F \<l>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)"
using assms lunit_coherence by simp
moreover have "D.iso (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f))"
using assms unit_char cmp_components_are_iso
by (intro D.isos_compose D.seqI) auto
ultimately show ?thesis
using assms D.invert_side_of_triangle(2) by metis
qed
moreover have "D.inv (\<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f)) =
(D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
using assms C.VV.arr_char unit_char FF_def D.inv_comp C.VV.dom_simp by simp
ultimately show "F \<l>\<^sub>C[f] =
\<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f))"
by simp
hence "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] =
D.inv (\<l>\<^sub>D[F f] \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<star>\<^sub>D F f) \<cdot>\<^sub>D D.inv (\<Phi> (trg\<^sub>C f, f)))"
using assms preserves_inv by simp
also have "... = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
using assms unit_char D.comp_assoc D.isos_compose
by (auto simp add: D.inv_comp)
finally show "F \<l>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (trg\<^sub>C f, f) \<cdot>\<^sub>D (unit (trg\<^sub>C f) \<star>\<^sub>D F f) \<cdot>\<^sub>D \<l>\<^sub>D\<^sup>-\<^sup>1[F f]"
by simp
qed
lemma preserves_runit:
assumes "C.ide f"
shows "F \<r>\<^sub>C[f] = \<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
and "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f]"
proof -
have "F \<r>\<^sub>C[f] = \<r>\<^sub>D[F f] \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)))"
proof -
have "\<r>\<^sub>D[F f] = F \<r>\<^sub>C[f] \<cdot>\<^sub>D \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))"
using assms runit_coherence by simp
moreover have "D.iso (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)))"
using assms unit_char D.iso_hcomp FF_def
apply (intro D.isos_compose D.seqI) by auto
moreover have "D.arr \<r>\<^sub>D[F f]"
using assms by simp
ultimately show ?thesis
using assms D.invert_side_of_triangle(2) by metis
qed
moreover have "D.inv (\<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f))) =
(F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
using assms C.VV.arr_char unit_char D.iso_hcomp FF_def D.inv_comp C.VV.dom_simp
by simp
ultimately show "F \<r>\<^sub>C[f] =
\<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f))"
by simp
hence "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] =
D.inv (\<r>\<^sub>D[F f] \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv (unit (src\<^sub>C f))) \<cdot>\<^sub>D D.inv (\<Phi> (f, src\<^sub>C f)))"
using assms preserves_inv C.iso_runit by simp
also have "... = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f]"
using assms unit_char D.comp_assoc D.isos_compose
by (auto simp add: D.inv_comp)
finally show "F \<r>\<^sub>C\<^sup>-\<^sup>1[f] = \<Phi> (f, src\<^sub>C f) \<cdot>\<^sub>D (F f \<star>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D \<r>\<^sub>D\<^sup>-\<^sup>1[F f]"
by simp
qed
lemma preserves_assoc:
assumes "C.ide f" and "C.ide g" and "C.ide h"
and "src\<^sub>C f = trg\<^sub>C g" and "src\<^sub>C g = trg\<^sub>C h"
shows "F \<a>\<^sub>C[f, g, h] = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
and "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
have "F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
using assms assoc_coherence [of f g h] by simp
moreover have "D.seq (\<Phi> (f, g \<star>\<^sub>C h)) ((F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h])"
using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto
moreover have 1: "D.iso (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h))"
using assms C.VV.arr_char FF_def C.VV.dom_simp C.VV.cod_simp by auto
moreover have "D.inv (\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)) =
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
using assms 1 C.VV.arr_char FF_def D.inv_comp C.VV.dom_simp C.VV.cod_simp
by simp
ultimately show 1: "F \<a>\<^sub>C[f, g, h] =
\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h))"
using D.invert_side_of_triangle(2)
[of "\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
"F \<a>\<^sub>C[f, g, h]" "\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)"]
D.comp_assoc
by simp
hence "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] =
D.inv (\<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h] \<cdot>\<^sub>D
(D.inv (\<Phi> (f, g)) \<star>\<^sub>D F h) \<cdot>\<^sub>D D.inv (\<Phi> (f \<star>\<^sub>C g, h)))"
using assms preserves_inv by simp
also have "... = \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
proof -
have "\<guillemotleft>\<Phi> (f, g \<star>\<^sub>C h) : F f \<star>\<^sub>D F (g \<star>\<^sub>C h) \<Rightarrow>\<^sub>D F (f \<star>\<^sub>C g \<star>\<^sub>C h)\<guillemotright> \<and> D.iso (\<Phi> (f, g \<star>\<^sub>C h))"
using assms by auto
moreover have "\<guillemotleft>F f \<star>\<^sub>D \<Phi> (g, h) : F f \<star>\<^sub>D F g \<star>\<^sub>D F h \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F (g \<star>\<^sub>C h)\<guillemotright> \<and>
D.iso (F f \<star>\<^sub>D \<Phi> (g, h))"
using assms
by (intro conjI D.hcomp_in_vhom, auto)
ultimately show ?thesis
using assms D.isos_compose D.comp_assoc
by (elim conjE D.in_homE) (auto simp add: D.inv_comp)
qed
finally show "F \<a>\<^sub>C\<^sup>-\<^sup>1[f, g, h] =
\<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h) \<cdot>\<^sub>D \<a>\<^sub>D\<^sup>-\<^sup>1[F f, F g, F h] \<cdot>\<^sub>D
(F f \<star>\<^sub>D D.inv (\<Phi> (g, h))) \<cdot>\<^sub>D D.inv (\<Phi> (f, g \<star>\<^sub>C h))"
by simp
qed
lemma preserves_hcomp:
assumes "C.hseq \<mu> \<nu>"
shows "F (\<mu> \<star>\<^sub>C \<nu>) =
\<Phi> (C.cod \<mu>, C.cod \<nu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D F \<nu>) \<cdot>\<^sub>D D.inv (\<Phi> (C.dom \<mu>, C.dom \<nu>))"
proof -
have "F (\<mu> \<star>\<^sub>C \<nu>) \<cdot>\<^sub>D \<Phi> (C.dom \<mu>, C.dom \<nu>) = \<Phi> (C.cod \<mu>, C.cod \<nu>) \<cdot>\<^sub>D (F \<mu> \<star>\<^sub>D F \<nu>)"
using assms \<Phi>.naturality C.VV.arr_char C.VV.cod_char FF_def C.VV.dom_simp
by auto
thus ?thesis
by (metis (no_types) assms C.hcomp_simps(3) C.hseqE C.ide_dom C.src_dom C.trg_dom
D.comp_arr_inv' D.comp_assoc cmp_components_are_iso cmp_simps(5) is_natural_1)
qed
lemma preserves_adjunction_data:
assumes "adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "adjunction_data_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof
interpret adjunction_data_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms by auto
show "D.ide (F f)"
using preserves_ide by simp
show "D.ide (F g)"
using preserves_ide by simp
show "\<guillemotleft>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) : src\<^sub>D (F f) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F f\<guillemotright>"
using antipar C.VV.ide_char C.VV.arr_char cmp_in_hom(2) unit_in_hom FF_def by auto
show "\<guillemotleft>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D src\<^sub>D (F g)\<guillemotright>"
using antipar C.VV.ide_char C.VV.arr_char FF_def cmp_in_hom(2) unit_in_hom(2)
unit_char(2)
by auto
qed
lemma preserves_equivalence:
assumes "equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
shows "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
proof -
interpret equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using assms by auto
show "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(F f) (F g) (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
using antipar unit_is_iso preserves_iso unit_char(2) C.VV.ide_char C.VV.arr_char
FF_def D.isos_compose
by (unfold_locales) auto
qed
lemma preserves_equivalence_maps:
assumes "C.equivalence_map f"
shows "D.equivalence_map (F f)"
proof -
obtain g \<eta> \<epsilon> where E: "equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms C.equivalence_map_def by auto
interpret E: equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>
using E by auto
show ?thesis
using E preserves_equivalence C.equivalence_map_def D.equivalence_map_def map\<^sub>0_simps
by blast
qed
lemma preserves_equivalent_objects:
assumes "C.equivalent_objects a b"
shows "D.equivalent_objects (map\<^sub>0 a) (map\<^sub>0 b)"
using assms D.equivalent_objects_def preserves_equivalence_maps
unfolding C.equivalent_objects_def by fast
lemma preserves_isomorphic:
assumes "C.isomorphic f g"
shows "D.isomorphic (F f) (F g)"
using assms C.isomorphic_def D.isomorphic_def preserves_iso by auto
lemma preserves_quasi_inverses:
assumes "C.quasi_inverses f g"
shows "D.quasi_inverses (F f) (F g)"
using assms C.quasi_inverses_def D.quasi_inverses_def preserves_equivalence by blast
lemma preserves_quasi_inverse:
assumes "C.equivalence_map f"
shows "D.isomorphic (F (C.some_quasi_inverse f)) (D.some_quasi_inverse (F f))"
using assms preserves_quasi_inverses C.quasi_inverses_some_quasi_inverse
D.quasi_inverse_unique D.quasi_inverses_some_quasi_inverse
preserves_equivalence_maps
by blast
end
subsection "Identity Pseudofunctors"
locale identity_pseudofunctor =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
begin
text\<open>
The underlying vertical functor is just the identity functor on the vertical category,
which is already available as \<open>B.map\<close>.
\<close>
abbreviation map
where "map \<equiv> B.map"
interpretation I: weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B src\<^sub>B trg\<^sub>B map
proof
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> B.isomorphic (map (src\<^sub>B \<mu>)) (src\<^sub>B (map \<mu>))"
by (simp add: B.isomorphic_reflexive)
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> B.isomorphic (map (trg\<^sub>B \<mu>)) (trg\<^sub>B (map \<mu>))"
by (simp add: B.isomorphic_reflexive)
qed
interpretation II: "functor" B.VV.comp B.VV.comp I.FF
using I.functor_FF by simp
interpretation H\<^sub>BoII: composite_functor B.VV.comp B.VV.comp V\<^sub>B I.FF
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
..
interpretation IoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>B \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close> map
..
text\<open>
The horizontal composition provides the compositor.
\<close>
abbreviation cmp
where "cmp \<equiv> \<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>"
interpretation cmp: natural_transformation B.VV.comp V\<^sub>B H\<^sub>BoII.map IoH\<^sub>B.map cmp
using B.VV.arr_char B.VV.dom_simp B.VV.cod_simp B.H.is_natural_1 B.H.is_natural_2
I.FF_def
apply unfold_locales
apply auto
by (meson B.hseqE B.hseq_char')+
interpretation cmp: natural_isomorphism B.VV.comp V\<^sub>B H\<^sub>BoII.map IoH\<^sub>B.map cmp
by unfold_locales simp
sublocale pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B map cmp
apply unfold_locales
by (metis B.assoc_is_natural_2 B.assoc_naturality B.assoc_simps(1) B.comp_ide_self
B.hcomp_simps(1) B.ide_char B.ide_hcomp B.map_simp fst_conv snd_conv)
lemma is_pseudofunctor:
shows "pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B map cmp"
..
lemma unit_char':
assumes "B.obj a"
shows "unit a = a"
proof -
have "src\<^sub>B a = a \<and> B.ide a"
using assms by auto
hence "a = unit a"
using assms B.comp_arr_dom B.comp_cod_arr I.map\<^sub>0_def map_def
B.ide_in_hom(2) B.objE [of a] B.ide_is_iso [of a]
by (intro unit_eqI) auto
thus ?thesis by simp
qed
end
lemma (in identity_pseudofunctor) map\<^sub>0_simp [simp]:
assumes "B.obj a"
shows "map\<^sub>0 a = a"
using assms map\<^sub>0_def by auto
(* TODO: Does not recognize map\<^sub>0_def unless the context is closed, then re-opened. *)
subsection "Embedding Pseudofunctors"
text \<open>
In this section, we construct the embedding pseudofunctor of a sub-bicategory
into the ambient bicategory.
\<close>
locale embedding_pseudofunctor =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
S: subbicategory
begin
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
notation B.in_hhom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
definition map
where "map \<mu> = (if S.arr \<mu> then \<mu> else B.null)"
lemma map_in_hom [intro]:
assumes "S.arr \<mu>"
shows "\<guillemotleft>map \<mu> : src\<^sub>B (map (S.src \<mu>)) \<rightarrow>\<^sub>B src\<^sub>B (map (S.trg \<mu>))\<guillemotright>"
and "\<guillemotleft>map \<mu> : map (S.dom \<mu>) \<Rightarrow>\<^sub>B map (S.cod \<mu>)\<guillemotright>"
proof -
show 1: "\<guillemotleft>map \<mu> : map (S.dom \<mu>) \<Rightarrow>\<^sub>B map (S.cod \<mu>)\<guillemotright>"
using assms map_def S.in_hom_char by fastforce
show "\<guillemotleft>map \<mu> : src\<^sub>B (map (S.src \<mu>)) \<rightarrow>\<^sub>B src\<^sub>B (map (S.trg \<mu>))\<guillemotright>"
using assms 1 map_def S.arr_char S.src_def S.trg_def S.obj_char S.obj_src S.obj_trg
by auto
qed
lemma map_simps [simp]:
assumes "S.arr \<mu>"
shows "B.arr (map \<mu>)"
and "src\<^sub>B (map \<mu>) = src\<^sub>B (map (S.src \<mu>))" and "trg\<^sub>B (map \<mu>) = src\<^sub>B (map (S.trg \<mu>))"
and "B.dom (map \<mu>) = map (S.dom \<mu>)" and "B.cod (map \<mu>) = map (S.cod \<mu>)"
using assms map_in_hom by blast+
interpretation "functor" S.comp V map
apply unfold_locales
apply auto
using map_def S.comp_char S.seq_char S.arr_char
apply auto[1]
using map_def S.comp_simp by auto
interpretation weak_arrow_of_homs S.comp S.src S.trg V src\<^sub>B trg\<^sub>B map
using S.arr_char map_def S.src_def S.trg_def S.src_closed S.trg_closed B.isomorphic_reflexive
preserves_ide S.ide_src S.ide_trg
apply unfold_locales
by presburger+
interpretation HoFF: composite_functor S.VV.comp B.VV.comp V FF
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
..
interpretation FoH: composite_functor S.VV.comp S.comp V \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> map
..
no_notation B.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>B _\<guillemotright>")
definition cmp
where "cmp \<mu>\<nu> = (if S.VV.arr \<mu>\<nu> then fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu> else B.null)"
lemma cmp_in_hom [intro]:
assumes "S.VV.arr \<mu>\<nu>"
shows "\<guillemotleft>cmp \<mu>\<nu> : src\<^sub>B (snd \<mu>\<nu>) \<rightarrow>\<^sub>B trg\<^sub>B (fst \<mu>\<nu>)\<guillemotright>"
and "\<guillemotleft>cmp \<mu>\<nu> : map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))\<guillemotright>"
proof -
show "\<guillemotleft>cmp \<mu>\<nu> : map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))
\<Rightarrow>\<^sub>B map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))\<guillemotright>"
proof
show 1: "B.arr (cmp \<mu>\<nu>)"
unfolding cmp_def
using assms S.arr_char S.VV.arr_char S.inclusion S.src_def S.trg_def by auto
show "B.dom (cmp \<mu>\<nu>) = map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))"
unfolding cmp_def map_def
using assms 1 cmp_def S.dom_simp S.cod_simp S.VV.arr_char S.arr_char S.hcomp_def
S.inclusion S.dom_closed
by auto
show "B.cod (cmp \<mu>\<nu>) = map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))"
unfolding cmp_def map_def
using assms 1 cmp_def S.dom_simp S.cod_simp S.VV.arr_char S.arr_char S.hcomp_def
S.inclusion S.cod_closed
apply auto
using S.hcomp_closed apply auto[1]
by (metis (no_types, lifting) B.VV.arrE)
qed
thus "\<guillemotleft>cmp \<mu>\<nu> : src\<^sub>B (snd \<mu>\<nu>) \<rightarrow>\<^sub>B trg\<^sub>B (fst \<mu>\<nu>)\<guillemotright>"
using cmp_def by auto
qed
lemma cmp_simps [simp]:
assumes "S.VV.arr \<mu>\<nu>"
shows "B.arr (cmp \<mu>\<nu>)"
and "src\<^sub>B (cmp \<mu>\<nu>) = S.src (snd \<mu>\<nu>)" and "trg\<^sub>B (cmp \<mu>\<nu>) = S.trg (fst \<mu>\<nu>)"
and "B.dom (cmp \<mu>\<nu>) = map (S.dom (fst \<mu>\<nu>)) \<star>\<^sub>B map (S.dom (snd \<mu>\<nu>))"
and "B.cod (cmp \<mu>\<nu>) = map (S.cod (fst \<mu>\<nu>) \<star> S.cod (snd \<mu>\<nu>))"
using assms cmp_in_hom S.src_def S.trg_def S.VV.arr_char
by simp_all blast+
lemma iso_cmp:
assumes "S.VV.ide \<mu>\<nu>"
shows "B.iso (cmp \<mu>\<nu>)"
using assms S.VV.ide_char S.VV.arr_char S.arr_char cmp_def S.ide_char S.src_def S.trg_def
by auto
interpretation \<Phi>\<^sub>E: natural_isomorphism S.VV.comp V HoFF.map FoH.map cmp
proof
show "\<And>\<mu>\<nu>. \<not> S.VV.arr \<mu>\<nu> \<Longrightarrow> cmp \<mu>\<nu> = B.null"
using cmp_def by simp
fix \<mu>\<nu>
assume \<mu>\<nu>: "S.VV.arr \<mu>\<nu>"
have 1: "S.arr (fst \<mu>\<nu>) \<and> S.arr (snd \<mu>\<nu>) \<and> S.src (fst \<mu>\<nu>) = S.trg (snd \<mu>\<nu>)"
using \<mu>\<nu> S.VV.arr_char by simp
show "B.dom (cmp \<mu>\<nu>) = HoFF.map (S.VV.dom \<mu>\<nu>)"
using \<mu>\<nu> FF_def S.VV.arr_char S.VV.dom_char S.arr_dom S.src_def S.trg_def
S.dom_char S.src.preserves_dom S.trg.preserves_dom
apply simp
by (metis (no_types, lifting))
show "B.cod (cmp \<mu>\<nu>) = FoH.map (S.VV.cod \<mu>\<nu>)"
using \<mu>\<nu> 1 map_def S.hseq_char S.hcomp_def S.cod_char S.arr_cod S.VV.cod_simp
by simp
show "cmp (S.VV.cod \<mu>\<nu>) \<cdot>\<^sub>B HoFF.map \<mu>\<nu> = cmp \<mu>\<nu>"
using \<mu>\<nu> 1 cmp_def S.VV.arr_char S.VV.cod_char FF_def S.arr_cod S.cod_simp
S.src_def S.trg_def map_def
apply simp
by (metis (no_types, lifting) B.comp_cod_arr B.hcomp_simps(4) B.hseqE B.src_cod
B.trg_cod cmp_simps(1) \<mu>\<nu>)
show "FoH.map \<mu>\<nu> \<cdot>\<^sub>B cmp (S.VV.dom \<mu>\<nu>) = cmp \<mu>\<nu>"
unfolding cmp_def map_def
using \<mu>\<nu> S.VV.arr_char B.VV.arr_char S.VV.dom_char S.VV.cod_char B.comp_arr_dom
S.hcomp_def
apply simp
by (metis (no_types, lifting) B.hcomp_simps(3) cmp_def cmp_simps(1) S.arr_char
S.dom_char S.hcomp_closed S.src_def S.trg_def)
next
show "\<And>fg. S.VV.ide fg \<Longrightarrow> B.iso (cmp fg)"
using iso_cmp by simp
qed
sublocale pseudofunctor S.comp S.hcomp S.\<a> \<i> S.src S.trg V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B map cmp
proof
fix f g h
assume f: "S.ide f" and g: "S.ide g" and h: "S.ide h"
and fg: "S.src f = S.trg g" and gh: "S.src g = S.trg h"
have 1: "B.ide f \<and> B.ide g \<and> B.ide h \<and> src\<^sub>B f = trg\<^sub>B g \<and> src\<^sub>B g = trg\<^sub>B h"
using f g h fg gh S.ide_char S.arr_char S.src_def S.trg_def by simp
show "map (S.\<a> f g h) \<cdot>\<^sub>B cmp (f \<star> g, h) \<cdot>\<^sub>B cmp (f, g) \<star>\<^sub>B map h =
cmp (f, g \<star> h) \<cdot>\<^sub>B (map f \<star>\<^sub>B cmp (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[map f, map g, map h]"
proof -
have "map (S.\<a> f g h) \<cdot>\<^sub>B cmp (f \<star> g, h) \<cdot>\<^sub>B cmp (f, g) \<star>\<^sub>B map h =
\<a>\<^sub>B[f, g, h] \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h) \<cdot>\<^sub>B ((f \<star>\<^sub>B g) \<star>\<^sub>B h)"
unfolding map_def cmp_def
using 1 f g h fg gh S.VVV.arr_char S.VV.arr_char B.VVV.arr_char B.VV.arr_char
B.comp_arr_dom S.arr_char S.comp_char S.hcomp_closed S.hcomp_def
S.ideD(1) S.src_def
by (simp add: S.assoc_closed)
also have "... = cmp (f, g \<star> h) \<cdot>\<^sub>B (map f \<star>\<^sub>B cmp (g, h)) \<cdot>\<^sub>B \<a>\<^sub>B[map f, map g, map h]"
unfolding cmp_def map_def
using 1 f g h fg gh S.VV.arr_char B.VVV.arr_char B.VV.arr_char
B.comp_arr_dom B.comp_cod_arr S.hcomp_def S.comp_char
S.arr_char S.assoc_closed S.hcomp_closed S.ideD(1) S.trg_def
by auto
finally show ?thesis by blast
qed
qed
lemma is_pseudofunctor:
shows "pseudofunctor S.comp S.hcomp S.\<a> \<i> S.src S.trg V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B map cmp"
..
lemma map\<^sub>0_simp [simp]:
assumes "S.obj a"
shows "map\<^sub>0 a = a"
using assms map\<^sub>0_def map_def S.obj_char by auto
lemma unit_char':
assumes "S.obj a"
shows "unit a = a"
proof -
have a: "S.arr a"
using assms by auto
have 1: "B.ide a"
using assms S.obj_char by auto
have 2: "src\<^sub>B a = a"
using assms S.obj_char by auto
have "a = unit a"
proof (intro unit_eqI)
show "S.obj a" by fact
show "\<guillemotleft>a : map\<^sub>0 a \<Rightarrow>\<^sub>B map a\<guillemotright>"
using assms a 2 map\<^sub>0_def map_def S.src_def S.dom_char S.cod_char S.obj_char
by auto
show "B.iso a"
using assms 1 by simp
show "a \<cdot>\<^sub>B \<i>[map\<^sub>0 a] = (map \<i>[a] \<cdot>\<^sub>B cmp (a, a)) \<cdot>\<^sub>B (a \<star>\<^sub>B a)"
proof -
have "a \<cdot>\<^sub>B \<i>[a] = \<i>[a] \<cdot>\<^sub>B cmp (a, a) \<cdot>\<^sub>B (a \<star>\<^sub>B a)"
proof -
have "a \<cdot>\<^sub>B \<i>[a] = \<i>[a]"
using assms 1 2 S.comp_cod_arr S.unitor_coincidence S.lunit_in_hom
S.obj_char S.comp_simp
by auto
moreover have "(a \<star>\<^sub>B a) \<cdot>\<^sub>B (a \<star>\<^sub>B a) = a \<star>\<^sub>B a"
using assms S.obj_char S.comp_ide_self by auto
moreover have "B.dom \<i>[a] = a \<star>\<^sub>B a"
using assms S.obj_char by simp
moreover have "\<i>[a] \<cdot>\<^sub>B (a \<star>\<^sub>B a) = \<i>[a]"
using assms S.obj_char B.comp_arr_dom by simp
ultimately show ?thesis
using assms cmp_def S.VV.arr_char by auto
qed
thus ?thesis
using assms a 2 map\<^sub>0_def map_def S.src_def B.comp_assoc by simp
qed
qed
thus ?thesis by simp
qed
end
subsection "Composition of Pseudofunctors"
text \<open>
In this section, we show how pseudofunctors may be composed. The main work is to
establish the coherence condition for associativity.
\<close>
locale composite_pseudofunctor =
B: bicategory V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B +
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
G: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
for V\<^sub>B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and H\<^sub>B :: "'b comp" (infixr "\<star>\<^sub>B" 53)
and \<a>\<^sub>B :: "'b \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b" ("\<a>\<^sub>B[_, _, _]")
and \<i>\<^sub>B :: "'b \<Rightarrow> 'b" ("\<i>\<^sub>B[_]")
and src\<^sub>B :: "'b \<Rightarrow> 'b"
and trg\<^sub>B :: "'b \<Rightarrow> 'b"
and V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'b \<Rightarrow> 'c"
and \<Phi>\<^sub>F :: "'b * 'b \<Rightarrow> 'c"
and G :: "'c \<Rightarrow> 'd"
and \<Phi>\<^sub>G :: "'c * 'c \<Rightarrow> 'd"
begin
sublocale composite_functor V\<^sub>B V\<^sub>C V\<^sub>D F G ..
sublocale weak_arrow_of_homs V\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D src\<^sub>D trg\<^sub>D \<open>G o F\<close>
proof
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> D.isomorphic ((G o F) (src\<^sub>B \<mu>)) (src\<^sub>D ((G o F) \<mu>))"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
show "D.isomorphic ((G o F) (src\<^sub>B \<mu>)) (src\<^sub>D ((G o F) \<mu>))"
proof -
have "(G o F) (src\<^sub>B \<mu>) = G (F (src\<^sub>B \<mu>))"
using \<mu> by simp
also have "D.isomorphic ... (G (src\<^sub>C (F \<mu>)))"
using \<mu> F.weakly_preserves_src G.preserves_iso
by (meson C.isomorphicE D.isomorphic_def G.preserves_hom)
also have "D.isomorphic ... (src\<^sub>D (G (F \<mu>)))"
using \<mu> G.weakly_preserves_src by blast
also have "... = src\<^sub>D ((G o F) \<mu>)"
by simp
finally show ?thesis by blast
qed
qed
show "\<And>\<mu>. B.arr \<mu> \<Longrightarrow> D.isomorphic ((G o F) (trg\<^sub>B \<mu>)) (trg\<^sub>D ((G o F) \<mu>))"
proof -
fix \<mu>
assume \<mu>: "B.arr \<mu>"
show "D.isomorphic ((G o F) (trg\<^sub>B \<mu>)) (trg\<^sub>D ((G o F) \<mu>))"
proof -
have "(G o F) (trg\<^sub>B \<mu>) = G (F (trg\<^sub>B \<mu>))"
using \<mu> by simp
also have "D.isomorphic ... (G (trg\<^sub>C (F \<mu>)))"
using \<mu> F.weakly_preserves_trg G.preserves_iso
by (meson C.isomorphicE D.isomorphic_def G.preserves_hom)
also have "D.isomorphic ... (trg\<^sub>D (G (F \<mu>)))"
using \<mu> G.weakly_preserves_trg by blast
also have "... = trg\<^sub>D ((G o F) \<mu>)"
by simp
finally show ?thesis by blast
qed
qed
qed
interpretation H\<^sub>DoGF_GF: composite_functor B.VV.comp D.VV.comp V\<^sub>D FF
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D snd \<mu>\<nu>\<close>
..
interpretation GFoH\<^sub>B: composite_functor B.VV.comp V\<^sub>B V\<^sub>D \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>\<close>
\<open>G o F\<close>
..
definition cmp
where "cmp \<mu>\<nu> = (if B.VV.arr \<mu>\<nu> then
G (F (H\<^sub>B (fst \<mu>\<nu>) (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>)))
else D.null)"
lemma cmp_in_hom [intro]:
assumes "B.VV.arr \<mu>\<nu>"
shows "\<guillemotleft>cmp \<mu>\<nu> : H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>) \<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof -
have "cmp \<mu>\<nu> = G (F (H\<^sub>B (fst \<mu>\<nu>) (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>)))"
using assms cmp_def by simp
moreover have "\<guillemotleft> ... : H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)
\<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>\<Phi>\<^sub>G (F (B.dom (fst \<mu>\<nu>)), F (B.dom (snd \<mu>\<nu>))) :
H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)
\<Rightarrow>\<^sub>D G (F (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C F (B.dom (snd \<mu>\<nu>)))\<guillemotright>"
using assms F.FF_def FF_def B.VV.arr_char B.VV.dom_simp by auto
show "\<guillemotleft>G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) :
G (F (B.dom (fst \<mu>\<nu>)) \<star>\<^sub>C F (B.dom (snd \<mu>\<nu>)))
\<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.dom \<mu>\<nu>)\<guillemotright>"
using assms B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
show "\<guillemotleft>G (F (fst \<mu>\<nu> \<star>\<^sub>B snd \<mu>\<nu>)) :
GFoH\<^sub>B.map (B.VV.dom \<mu>\<nu>) \<Rightarrow>\<^sub>D GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)\<guillemotright>"
proof -
have "B.VV.in_hom \<mu>\<nu> (B.VV.dom \<mu>\<nu>) (B.VV.cod \<mu>\<nu>)"
using assms by auto
thus ?thesis by auto
qed
qed
ultimately show ?thesis by argo
qed
lemma cmp_simps [simp]:
assumes "B.VV.arr \<mu>\<nu>"
shows "D.arr (cmp \<mu>\<nu>)"
and "D.dom (cmp \<mu>\<nu>) = H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)"
and "D.cod (cmp \<mu>\<nu>) = GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)"
using assms cmp_in_hom by blast+
interpretation \<Phi>: natural_transformation
B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp
proof
show "\<And>\<mu>\<nu>. \<not> B.VV.arr \<mu>\<nu> \<Longrightarrow> cmp \<mu>\<nu> = D.null"
unfolding cmp_def by simp
fix \<mu>\<nu>
assume \<mu>\<nu>: "B.VV.arr \<mu>\<nu>"
show "D.dom (cmp \<mu>\<nu>) = H\<^sub>DoGF_GF.map (B.VV.dom \<mu>\<nu>)"
using \<mu>\<nu> cmp_in_hom by blast
show "D.cod (cmp \<mu>\<nu>) = GFoH\<^sub>B.map (B.VV.cod \<mu>\<nu>)"
using \<mu>\<nu> cmp_in_hom by blast
show "GFoH\<^sub>B.map \<mu>\<nu> \<cdot>\<^sub>D cmp (B.VV.dom \<mu>\<nu>) = cmp \<mu>\<nu>"
unfolding cmp_def
using \<mu>\<nu> B.VV.ide_char B.VV.arr_char D.comp_ide_arr B.VV.dom_char D.comp_assoc
is_natural_1
apply simp
by (metis (no_types, lifting) B.H.preserves_arr B.hcomp_simps(3))
show "cmp (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>D H\<^sub>DoGF_GF.map \<mu>\<nu> = cmp \<mu>\<nu>"
proof -
have "cmp (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>D H\<^sub>DoGF_GF.map \<mu>\<nu> =
(G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>))"
unfolding cmp_def
using \<mu>\<nu> B.VV.arr_char B.VV.dom_simp B.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>))"
proof -
have "D.ide (G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))))"
using \<mu>\<nu> B.VV.arr_char by simp
moreover have "D.seq (G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))))
(G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
using \<mu>\<nu> B.VV.arr_char B.VV.dom_char B.VV.cod_char F.FF_def
apply (intro D.seqI)
apply auto
proof -
show "G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>))) =
D.cod (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
proof -
have "D.seq (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))))
(\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))))"
using \<mu>\<nu> B.VV.arr_char F.FF_def B.VV.arr_char B.VV.dom_char B.VV.cod_char
by (intro D.seqI) auto
thus ?thesis
using \<mu>\<nu> B.VV.arr_char B.VV.cod_simp by simp
qed
qed
ultimately show ?thesis
using \<mu>\<nu> D.comp_ide_arr [of "G (F (B.cod (fst \<mu>\<nu>) \<star>\<^sub>B B.cod (snd \<mu>\<nu>)))"
"G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>)))"]
by simp
qed
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F (B.cod (fst \<mu>\<nu>)), F (B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
(fst (FF \<mu>\<nu>) \<star>\<^sub>D snd (FF \<mu>\<nu>)))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.cod (F.FF \<mu>\<nu>)) \<cdot>\<^sub>D G.H\<^sub>DoFF.map (F.FF \<mu>\<nu>)"
using \<mu>\<nu> B.VV.arr_char F.FF_def G.FF_def FF_def C.VV.cod_simp by auto
also have "... = G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G.FoH\<^sub>C.map (F.FF \<mu>\<nu>) \<cdot>\<^sub>D \<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
using \<mu>\<nu> B.VV.arr_char G.\<Phi>.naturality C.VV.dom_simp C.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D
G.FoH\<^sub>C.map (F.FF \<mu>\<nu>)) \<cdot>\<^sub>D \<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
using D.comp_assoc by simp
also have "... = (G (\<Phi>\<^sub>F (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>C F.H\<^sub>DoFF.map \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
proof -
have "(B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>)) = B.VV.cod \<mu>\<nu>"
using \<mu>\<nu> B.VV.arr_char B.VV.cod_simp by simp
moreover have "G.FoH\<^sub>C.map (F.FF \<mu>\<nu>) = G (F.H\<^sub>DoFF.map \<mu>\<nu>)"
using \<mu>\<nu> F.FF_def by simp
moreover have "G (\<Phi>\<^sub>F (B.cod (fst \<mu>\<nu>), B.cod (snd \<mu>\<nu>))) \<cdot>\<^sub>D G (F.H\<^sub>DoFF.map \<mu>\<nu>) =
G (\<Phi>\<^sub>F (B.VV.cod \<mu>\<nu>) \<cdot>\<^sub>C F.H\<^sub>DoFF.map \<mu>\<nu>)"
using \<mu>\<nu> B.VV.arr_char
by (metis (no_types, lifting) F.\<Phi>.is_natural_2 F.\<Phi>.preserves_reflects_arr
G.preserves_comp calculation(1))
ultimately show ?thesis by argo
qed
also have "... = G (F.FoH\<^sub>C.map \<mu>\<nu> \<cdot>\<^sub>C \<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
using \<mu>\<nu> F.\<Phi>.naturality [of \<mu>\<nu>] F.FF_def by simp
also have "... = G (F.FoH\<^sub>C.map \<mu>\<nu>) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (C.VV.dom (F.FF \<mu>\<nu>))"
proof -
have "G (F.FoH\<^sub>C.map \<mu>\<nu> \<cdot>\<^sub>C \<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>)) =
G (F.FoH\<^sub>C.map \<mu>\<nu>) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom \<mu>\<nu>))"
using \<mu>\<nu>
by (metis (mono_tags, lifting) F.\<Phi>.is_natural_1 F.\<Phi>.preserves_reflects_arr
G.preserves_comp)
thus ?thesis
using \<mu>\<nu> D.comp_assoc by simp
qed
also have "... = cmp \<mu>\<nu>"
using \<mu>\<nu> B.VV.arr_char cmp_def F.FF_def F.FF.preserves_dom B.VV.dom_simp
by auto
finally show ?thesis by simp
qed
qed
interpretation \<Phi>: natural_isomorphism B.VV.comp V\<^sub>D H\<^sub>DoGF_GF.map GFoH\<^sub>B.map cmp
proof
show "\<And>hk. B.VV.ide hk \<Longrightarrow> D.iso (cmp hk)"
proof -
fix hk
assume hk: "B.VV.ide hk"
have "D.iso (G (F (fst hk \<star>\<^sub>B snd hk)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (B.VV.dom hk)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
proof (intro D.isos_compose)
show "D.iso (\<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
using hk G.\<Phi>.components_are_iso [of "(F (B.dom (fst hk)), F (B.dom (snd hk)))"]
C.VV.ide_char B.VV.arr_char B.VV.dom_char
by (metis (no_types, lifting) B.VV.ideD(1) B.VV.ideD(2) B.VxV.dom_char
F.FF_def F.FF.components_are_iso G.\<Phi>.preserves_iso fst_conv snd_conv)
show "D.iso (G (\<Phi>\<^sub>F (B.VV.dom hk)))"
using hk F.\<Phi>.components_are_iso B.VV.arr_char B.VV.dom_char B.VV.ideD(2)
by auto
show "D.seq (G (\<Phi>\<^sub>F (B.VV.dom hk))) (\<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
using hk B.VV.arr_char B.VV.ide_char B.VV.dom_char C.VV.arr_char F.FF_def
C.VV.dom_simp C.VV.cod_simp
by auto
show "D.iso (G (F (fst hk \<star>\<^sub>B snd hk)))"
using hk F.\<Phi>.components_are_iso B.VV.arr_char by simp
show "D.seq (G (F (fst hk \<star>\<^sub>B snd hk)))
(G (\<Phi>\<^sub>F (B.VV.dom hk)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (B.dom (fst hk)), F (B.dom (snd hk))))"
using hk B.VV.arr_char B.VV.dom_char
by (metis (no_types, lifting) B.VV.ideD(1) cmp_def cmp_simps(1))
qed
thus "D.iso (cmp hk)"
unfolding cmp_def using hk by simp
qed
qed
sublocale pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>G o F\<close> cmp
proof
fix f g h
assume f: "B.ide f" and g: "B.ide g" and h: "B.ide h"
assume fg: "src\<^sub>B f = trg\<^sub>B g" and gh: "src\<^sub>B g = trg\<^sub>B h"
show "map \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>D cmp (f \<star>\<^sub>B g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h) =
cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[map f, map g, map h]"
proof -
have "map \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>D cmp (f \<star>\<^sub>B g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h) =
G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (F ((f \<star>\<^sub>B g) \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (F (f \<star>\<^sub>B g)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
unfolding cmp_def
using f g h fg gh B.VV.arr_char B.VV.dom_simp by simp
also have "... = G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (F (f \<star>\<^sub>B g)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh D.comp_ide_arr D.comp_assoc
by (metis B.ideD(1) B.ide_hcomp B.src_hcomp F.cmp_simps(1) F.cmp_simps(5)
G.is_natural_2)
also have "... = G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h)) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g fg
by (metis (no_types) D.comp_assoc F.cmp_simps(1) F.cmp_simps(5) G.is_natural_2)
also have "... = (G (F \<a>\<^sub>B[f, g, h]) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f \<star>\<^sub>B g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh B.VV.arr_char B.VV.cod_simp by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh B.VV.arr_char C.VV.arr_char F.FF_def D.whisker_right
B.VV.dom_simp C.VV.cod_simp
by auto
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h))) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D
(G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
proof -
have "\<Phi>\<^sub>G (F (f \<star>\<^sub>B g), F h) = \<Phi>\<^sub>G (C.VV.cod (\<Phi>\<^sub>F (f, g), F h))"
using f g h fg gh B.VV.arr_char C.VV.arr_char B.VV.cod_simp C.VV.cod_simp
by simp
moreover have "G (\<Phi>\<^sub>F (f, g)) \<star>\<^sub>D G (F h) = G.H\<^sub>DoFF.map (\<Phi>\<^sub>F (f, g), F h)"
using f g h fg gh B.VV.arr_char C.VV.arr_char G.FF_def by simp
moreover have "G.FoH\<^sub>C.map (\<Phi>\<^sub>F (f, g), F h) = G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)"
using f g h fg gh B.VV.arr_char by simp
moreover have "\<Phi>\<^sub>G (C.VV.dom (\<Phi>\<^sub>F (f, g), F h)) = \<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h)"
using f g h fg gh C.VV.arr_char F.cmp_in_hom [of f g] C.VV.dom_simp by auto
ultimately show ?thesis
using f g h fg gh B.VV.arr_char G.\<Phi>.naturality
by (metis (mono_tags, lifting) C.VV.arr_cod_iff_arr C.VV.arr_dom_iff_arr
G.FoH\<^sub>C.is_extensional G.H\<^sub>DoFF.is_extensional G.\<Phi>.is_extensional)
qed
also have "... = (G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>D (G (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G ((F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h)) \<cdot>\<^sub>C (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
also have "... = G (F \<a>\<^sub>B[f, g, h] \<cdot>\<^sub>C \<Phi>\<^sub>F (f \<star>\<^sub>B g, h) \<cdot>\<^sub>C (\<Phi>\<^sub>F (f, g) \<star>\<^sub>C F h)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using C.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>C \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh F.assoc_coherence [of f g h] by simp
also have "... = G ((\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>C \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using C.comp_assoc by simp
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D G \<a>\<^sub>C[F f, F g, F h]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D (\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using f g h fg gh B.VV.arr_char F.FF_def B.VV.dom_simp B.VV.cod_simp by auto
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
G \<a>\<^sub>C[F f, F g, F h] \<cdot>\<^sub>D \<Phi>\<^sub>G (F f \<star>\<^sub>C F g, F h) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F g) \<star>\<^sub>D G (F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
using f g h fg gh G.assoc_coherence [of "F f" "F g" "F h"] by simp
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
using D.comp_assoc by simp
also have "... = (cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h))) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
proof -
have "G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h)) =
cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h))"
proof -
have "cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h)) =
(G (F (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
unfolding cmp_def
using f g h fg gh B.VV.arr_char B.VV.dom_simp B.VV.cod_simp by simp
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "G (F (f \<star>\<^sub>B g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h))"
using f g h fg gh B.VV.arr_char D.comp_ide_arr B.VV.dom_simp B.VV.cod_simp
by simp
thus ?thesis
using D.comp_assoc by metis
qed
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "G (F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (g, h)) = G (\<Phi>\<^sub>F (g, h))"
using f g h fg gh B.VV.arr_char D.comp_ide_arr B.VV.dom_simp B.VV.cod_simp
by simp
thus ?thesis
using D.comp_assoc by metis
qed
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using f g h fg gh
D.whisker_left [of "G (F f)" "G (\<Phi>\<^sub>F (g, h))" "\<Phi>\<^sub>G (F g, F h)"]
by fastforce
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h)) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h)))) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D
(G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h)) \<cdot>\<^sub>D
(G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
proof -
have "\<Phi>\<^sub>G (C.VV.cod (F f, \<Phi>\<^sub>F (g, h))) = \<Phi>\<^sub>G (F f, F (g \<star>\<^sub>B h))"
using f g h fg gh B.VV.arr_char C.VV.cod_char B.VV.dom_simp B.VV.cod_simp
by auto
moreover have "G.H\<^sub>DoFF.map (F f, \<Phi>\<^sub>F (g, h)) = G (F f) \<star>\<^sub>D G (\<Phi>\<^sub>F (g, h))"
using f g h fg gh B.VV.arr_char G.FF_def by auto
moreover have "G.FoH\<^sub>C.map (F f, \<Phi>\<^sub>F (g, h)) = G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))"
using f g h fg gh B.VV.arr_char C.VV.arr_char by simp
moreover have "C.VV.dom (F f, \<Phi>\<^sub>F (g, h)) = (F f, F g \<star>\<^sub>C F h)"
using f g h fg gh B.VV.arr_char C.VV.arr_char C.VV.dom_char
F.cmp_in_hom [of g h]
by auto
ultimately show ?thesis
using f g h fg gh B.VV.arr_char C.VV.arr_char
G.\<Phi>.naturality [of "(F f, \<Phi>\<^sub>F (g, h))"]
by simp
qed
also have "... = (G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h)) \<cdot>\<^sub>D (G (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h)))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using D.comp_assoc by simp
also have "... = G (\<Phi>\<^sub>F (f, g \<star>\<^sub>B h) \<cdot>\<^sub>C (F f \<star>\<^sub>C \<Phi>\<^sub>F (g, h))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F f, F g \<star>\<^sub>C F h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D \<Phi>\<^sub>G (F g, F h))"
using f g h fg gh B.VV.arr_char
by (metis (no_types, lifting) B.assoc_simps(1) C.comp_assoc C.seqE
F.preserves_assoc(1) F.preserves_reflects_arr G.preserves_comp)
finally show ?thesis by simp
qed
thus ?thesis by simp
qed
also have "... = cmp (f, g \<star>\<^sub>B h) \<cdot>\<^sub>D (G (F f) \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D
\<a>\<^sub>D[G (F f), G (F g), G (F h)]"
using D.comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma is_pseudofunctor:
shows "pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (G o F) cmp"
..
lemma map\<^sub>0_simp [simp]:
assumes "B.obj a"
shows "map\<^sub>0 a = G.map\<^sub>0 (F.map\<^sub>0 a)"
using assms map\<^sub>0_def by auto
lemma unit_char':
assumes "B.obj a"
shows "unit a = G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)"
proof -
have "G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) = unit a"
proof (intro unit_eqI [of a "G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)"])
show "B.obj a" by fact
show "\<guillemotleft>G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) : map\<^sub>0 a \<Rightarrow>\<^sub>D map a\<guillemotright>"
using assms by auto
show "D.iso (G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a))"
by (simp add: D.isos_compose F.unit_char(2) G.unit_char(2) assms)
show "(G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] =
(map \<i>\<^sub>B[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D
(G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a))"
proof -
have 1: "cmp (a, a) = G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F a, F a)"
proof -
have "cmp (a, a) = (G (F (a \<star>\<^sub>B a)) \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a))) \<cdot>\<^sub>D \<Phi>\<^sub>G (F a, F a)"
using assms cmp_def B.VV.ide_char B.VV.arr_char B.VV.dom_char B.VV.cod_char
B.VxV.dom_char B.objE D.comp_assoc B.obj_simps
by simp
also have "... = G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D \<Phi>\<^sub>G (F a, F a)"
using assms D.comp_cod_arr B.VV.arr_char B.VV.ide_char by auto
finally show ?thesis by blast
qed
have "(map \<i>\<^sub>B[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D
(G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) =
map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D
(\<Phi>\<^sub>G (F a, F a) \<cdot>\<^sub>D (G (F.unit a) \<star>\<^sub>D G (F.unit a))) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms 1 D.comp_assoc D.interchange by simp
also have "... = (map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D G (F.unit a \<star>\<^sub>C F.unit a)) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms G.\<Phi>.naturality [of "(F.unit a, F.unit a)"]
C.VV.arr_char C.VV.dom_char C.VV.cod_char G.FF_def D.comp_assoc
by simp
also have "... = (G (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
proof -
have "C.arr (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))"
using assms B.VV.ide_char B.VV.arr_char F.cmp_in_hom(2)
by (intro C.seqI' C.comp_in_homI) auto
hence "map \<i>\<^sub>B[a] \<cdot>\<^sub>D G (\<Phi>\<^sub>F (a, a)) \<cdot>\<^sub>D G (F.unit a \<star>\<^sub>C F.unit a) =
G (F \<i>\<^sub>B[a] \<cdot>\<^sub>C \<Phi>\<^sub>F (a, a) \<cdot>\<^sub>C (F.unit a \<star>\<^sub>C F.unit a))"
by auto
thus ?thesis by argo
qed
also have "... = G (F.unit a \<cdot>\<^sub>C \<i>\<^sub>C[F.map\<^sub>0 a]) \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms F.unit_char C.comp_assoc by simp
also have "... = G (F.unit a) \<cdot>\<^sub>D (G \<i>\<^sub>C[F.map\<^sub>0 a] \<cdot>\<^sub>D
\<Phi>\<^sub>G (F.map\<^sub>0 a, F.map\<^sub>0 a)) \<cdot>\<^sub>D
(G.unit (F.map\<^sub>0 a) \<star>\<^sub>D G.unit (F.map\<^sub>0 a))"
using assms D.comp_assoc by simp
also have "... = (G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^sub>D[G.map\<^sub>0 (F.map\<^sub>0 a)]"
using assms G.unit_char D.comp_assoc by simp
also have "... = (G (F.unit a) \<cdot>\<^sub>D G.unit (F.map\<^sub>0 a)) \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a]"
using assms map\<^sub>0_def by auto
finally show ?thesis by simp
qed
qed
thus ?thesis by simp
qed
end
subsection "Restriction of Pseudofunctors"
text \<open>
In this section, we construct the restriction and corestriction of a pseudofunctor to
a subbicategory of its domain and codomain, respectively.
\<close>
locale restricted_pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi> +
C': subbicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C Arr
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd"
and Arr :: "'c \<Rightarrow> bool"
begin
abbreviation map
where "map \<equiv> \<lambda>\<mu>. if C'.arr \<mu> then F \<mu> else D.null"
abbreviation cmp
where "cmp \<equiv> \<lambda>\<mu>\<nu>. if C'.VV.arr \<mu>\<nu> then \<Phi> \<mu>\<nu> else D.null"
interpretation "functor" C'.comp V\<^sub>D map
using C'.inclusion C'.arr_char C'.dom_char C'.cod_char C'.seq_char C'.comp_char
C'.arr_dom C'.arr_cod
apply (unfold_locales)
apply auto
by presburger
interpretation weak_arrow_of_homs C'.comp C'.src C'.trg V\<^sub>D src\<^sub>D trg\<^sub>D map
using C'.arrE C'.ide_src C'.ide_trg C'.inclusion C'.src_def C'.trg_def
F.weakly_preserves_src F.weakly_preserves_trg
by unfold_locales auto
interpretation H\<^sub>D\<^sub>'oFF: composite_functor C'.VV.comp D.VV.comp V\<^sub>D FF
\<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>D snd \<mu>\<nu>\<close>
..
interpretation FoH\<^sub>C\<^sub>': composite_functor C'.VV.comp C'.comp V\<^sub>D
\<open>\<lambda>\<mu>\<nu>. C'.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> map
..
interpretation \<Phi>: natural_transformation C'.VV.comp V\<^sub>D H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C\<^sub>'.map cmp
using C'.arr_char C'.dom_char C'.cod_char C'.VV.arr_char C'.VV.dom_char C'.VV.cod_char
FF_def C'.inclusion C'.dom_closed C'.cod_closed C'.src_def C'.trg_def
F.cmp_simps'(4) C'.hcomp_def C'.hcomp_closed F.\<Phi>.is_natural_1 F.\<Phi>.is_natural_2
C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def F.cmp_simps'(5)
by unfold_locales auto
interpretation \<Phi>: natural_isomorphism C'.VV.comp V\<^sub>D H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C\<^sub>'.map cmp
using C.VV.ide_char C.VV.arr_char C'.VV.ide_char C'.VV.arr_char C'.arr_char
C'.src_def C'.trg_def C'.ide_char F.\<Phi>.components_are_iso
by unfold_locales auto
sublocale pseudofunctor C'.comp C'.hcomp C'.\<a> \<i>\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
map cmp
using F.assoc_coherence C'.VVV.arr_char C'.VV.arr_char C'.arr_char C'.hcomp_def
C'.src_def C'.trg_def C'.assoc_closed C'.hcomp_closed C'.ide_char
by unfold_locales (simp add: C'.ide_char C'.src_def C'.trg_def)
lemma is_pseudofunctor:
shows "pseudofunctor C'.comp C'.hcomp C'.\<a> \<i>\<^sub>C C'.src C'.trg V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D map cmp"
..
lemma map\<^sub>0_simp [simp]:
assumes "C'.obj a"
shows "map\<^sub>0 a = F.map\<^sub>0 a"
using assms map\<^sub>0_def C'.obj_char by auto
lemma unit_char':
assumes "C'.obj a"
shows "F.unit a = unit a"
proof (intro unit_eqI)
show "C'.obj a" by fact
show "\<guillemotleft>F.unit a : map\<^sub>0 a \<Rightarrow>\<^sub>D map a\<guillemotright>"
using assms map\<^sub>0_def F.unit_in_hom(2) [of a] C'.obj_char by auto
show "D.iso (F.unit a)"
using assms by (simp add: C'.obj_char F.unit_char(2))
show "F.unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.unit a)"
proof -
have "F.unit a \<cdot>\<^sub>D \<i>\<^sub>D[map\<^sub>0 a] = F.unit a \<cdot>\<^sub>D \<i>\<^sub>D[F.map\<^sub>0 a]"
using assms map\<^sub>0_def F.map\<^sub>0_def by auto
also have "... = (F \<i>\<^sub>C[a] \<cdot>\<^sub>D \<Phi> (a, a)) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.unit a)"
using assms C'.obj_char F.unit_char(3) [of a] by simp
also have "... = (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.unit a)"
using assms C.VV.arr_char D.iso_is_arr iso_\<i> by auto
finally show ?thesis by simp
qed
qed
end
text \<open>
We define the corestriction construction only for the case of sub-bicategories
determined by a set of objects of the ambient bicategory.
There are undoubtedly more general constructions, but this one is adequate for our
present needs.
\<close>
locale corestricted_pseudofunctor =
C: bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C +
D: bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D +
F: pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D F \<Phi> +
D': subbicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>\<lambda>\<mu>. D.arr \<mu> \<and> Obj (src\<^sub>D \<mu>) \<and> Obj (trg\<^sub>D \<mu>)\<close>
for V\<^sub>C :: "'c comp" (infixr "\<cdot>\<^sub>C" 55)
and H\<^sub>C :: "'c comp" (infixr "\<star>\<^sub>C" 53)
and \<a>\<^sub>C :: "'c \<Rightarrow> 'c \<Rightarrow> 'c \<Rightarrow> 'c" ("\<a>\<^sub>C[_, _, _]")
and \<i>\<^sub>C :: "'c \<Rightarrow> 'c" ("\<i>\<^sub>C[_]")
and src\<^sub>C :: "'c \<Rightarrow> 'c"
and trg\<^sub>C :: "'c \<Rightarrow> 'c"
and V\<^sub>D :: "'d comp" (infixr "\<cdot>\<^sub>D" 55)
and H\<^sub>D :: "'d comp" (infixr "\<star>\<^sub>D" 53)
and \<a>\<^sub>D :: "'d \<Rightarrow> 'd \<Rightarrow> 'd \<Rightarrow> 'd" ("\<a>\<^sub>D[_, _, _]")
and \<i>\<^sub>D :: "'d \<Rightarrow> 'd" ("\<i>\<^sub>D[_]")
and src\<^sub>D :: "'d \<Rightarrow> 'd"
and trg\<^sub>D :: "'d \<Rightarrow> 'd"
and F :: "'c \<Rightarrow> 'd"
and \<Phi> :: "'c * 'c \<Rightarrow> 'd"
and Obj :: "'d \<Rightarrow> bool" +
assumes preserves_arr: "\<And>\<mu>. C.arr \<mu> \<Longrightarrow> D'.arr (F \<mu>)"
begin
abbreviation map
where "map \<equiv> F"
abbreviation cmp
where "cmp \<equiv> \<Phi>"
interpretation "functor" V\<^sub>C D'.comp F
using preserves_arr F.is_extensional D'.arr_char D'.dom_char D'.cod_char D'.comp_char
by (unfold_locales) auto
interpretation weak_arrow_of_homs V\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.src D'.trg F
proof
fix \<mu>
assume \<mu>: "C.arr \<mu>"
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : F (src\<^sub>C \<mu>) \<Rightarrow>\<^sub>D src\<^sub>D (F \<mu>)\<guillemotright> \<and> D.iso \<phi>"
using \<mu> F.weakly_preserves_src by auto
have 2: "D'.in_hom \<phi> (F (src\<^sub>C \<mu>)) (D'.src (F \<mu>))"
using \<mu> \<phi> D'.arr_char D'.dom_char D'.cod_char D'.src_def D.vconn_implies_hpar(1-2)
preserves_arr
by (intro D'.in_homI) auto
moreover have "D'.iso \<phi>"
using 2 \<phi> D'.iso_char D'.arr_char by fastforce
ultimately show "D'.isomorphic (F (src\<^sub>C \<mu>)) (D'.src (F \<mu>))"
using D'.isomorphic_def by auto
obtain \<psi> where \<psi>: "\<guillemotleft>\<psi> : F (trg\<^sub>C \<mu>) \<Rightarrow>\<^sub>D trg\<^sub>D (F \<mu>)\<guillemotright> \<and> D.iso \<psi>"
using \<mu> F.weakly_preserves_trg by auto
have 2: "D'.in_hom \<psi> (F (trg\<^sub>C \<mu>)) (D'.trg (F \<mu>))"
using \<mu> \<psi> D'.arr_char D'.dom_char D'.cod_char D'.trg_def D.vconn_implies_hpar(1-2)
preserves_arr
by (intro D'.in_homI) auto
moreover have "D'.iso \<psi>"
using 2 \<psi> D'.iso_char D'.arr_char by fastforce
ultimately show "D'.isomorphic (F (trg\<^sub>C \<mu>)) (D'.trg (F \<mu>))"
using D'.isomorphic_def by auto
qed
interpretation H\<^sub>D\<^sub>'oFF: composite_functor C.VV.comp D'.VV.comp D'.comp FF
\<open>\<lambda>\<mu>\<nu>. D'.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close>
..
interpretation FoH\<^sub>C: composite_functor C.VV.comp V\<^sub>C D'.comp \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star>\<^sub>C snd \<mu>\<nu>\<close>
F
..
interpretation natural_transformation C.VV.comp D'.comp H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C.map \<Phi>
proof
show "\<And>\<mu>\<nu>. \<not> C.VV.arr \<mu>\<nu> \<Longrightarrow> \<Phi> \<mu>\<nu> = D'.null"
by (simp add: F.\<Phi>.is_extensional)
fix \<mu>\<nu>
assume \<mu>\<nu>: "C.VV.arr \<mu>\<nu>"
have 1: "D'.arr (\<Phi> \<mu>\<nu>)"
using \<mu>\<nu> D'.arr_char F.\<Phi>.is_natural_1 F.\<Phi>.components_are_iso
by (metis (no_types, lifting) D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr
F.\<Phi>.preserves_reflects_arr)
show "D'.dom (\<Phi> \<mu>\<nu>) = H\<^sub>D\<^sub>'oFF.map (C.VV.dom \<mu>\<nu>)"
using 1 \<mu>\<nu> D'.dom_char C.VV.arr_char C.VV.dom_char F.FF_def FF_def D'.hcomp_def
by simp
show "D'.cod (\<Phi> \<mu>\<nu>) = FoH\<^sub>C.map (C.VV.cod \<mu>\<nu>)"
using 1 \<mu>\<nu> D'.cod_char C.VV.arr_char F.FF_def FF_def D'.hcomp_def by simp
show "D'.comp (FoH\<^sub>C.map \<mu>\<nu>) (\<Phi> (C.VV.dom \<mu>\<nu>)) = \<Phi> \<mu>\<nu>"
using 1 \<mu>\<nu> D'.arr_char D'.comp_char C.VV.dom_char F.\<Phi>.is_natural_1
C.VV.arr_dom D.src_vcomp D.trg_vcomp FoH\<^sub>C.preserves_arr F.\<Phi>.preserves_reflects_arr
by (metis (mono_tags, lifting))
show "D'.comp (\<Phi> (C.VV.cod \<mu>\<nu>)) (H\<^sub>D\<^sub>'oFF.map \<mu>\<nu>) = \<Phi> \<mu>\<nu>"
proof -
have "Obj (F.map\<^sub>0 (trg\<^sub>C (fst \<mu>\<nu>))) \<and> Obj (F.map\<^sub>0 (trg\<^sub>C (snd \<mu>\<nu>)))"
using \<mu>\<nu> C.VV.arr_char
by (metis (no_types, lifting) C.src_trg C.trg.preserves_reflects_arr D'.arr_char
F.map\<^sub>0_def preserves_hseq)
moreover have "Obj (F.map\<^sub>0 (src\<^sub>C (snd \<mu>\<nu>)))"
using \<mu>\<nu> C.VV.arr_char
by (metis (no_types, lifting) C.src.preserves_reflects_arr C.trg_src D'.arr_char
F.map\<^sub>0_def preserves_hseq)
ultimately show ?thesis
using \<mu>\<nu> 1 D'.arr_char D'.comp_char D'.hseq_char C.VV.arr_char C.VV.cod_char
C.VxV.cod_char FF_def F.FF_def D'.hcomp_char preserves_hseq
apply simp
using F.\<Phi>.is_natural_2 by force
qed
qed
interpretation natural_isomorphism C.VV.comp D'.comp H\<^sub>D\<^sub>'oFF.map FoH\<^sub>C.map \<Phi>
apply unfold_locales
using D'.iso_char D'.arr_char by fastforce
sublocale pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\<a> \<i>\<^sub>D D'.src D'.trg
F \<Phi>
proof
fix f g h
assume f: "C.ide f" and g: "C.ide g" and h: "C.ide h"
and fg: "src\<^sub>C f = trg\<^sub>C g" and gh: "src\<^sub>C g = trg\<^sub>C h"
have "D'.comp (F \<a>\<^sub>C[f, g, h]) (D'.comp (\<Phi> (f \<star>\<^sub>C g, h)) (D'.hcomp (\<Phi> (f, g)) (F h))) =
F \<a>\<^sub>C[f, g, h] \<cdot>\<^sub>D \<Phi> (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (\<Phi> (f, g) \<star>\<^sub>D F h)"
proof -
have 1: "D'.arr (cmp (f, g) \<star>\<^sub>D map h)"
using f g h fg gh D'.arr_char C.VV.arr_char
by (metis (no_types, lifting) C.ideD(1) D'.hcomp_closed F.\<Phi>.preserves_reflects_arr
F.cmp_simps(1) F.cmp_simps(2) F.preserves_hseq preserves_reflects_arr)
moreover have 2: "D.seq (cmp (f \<star>\<^sub>C g, h)) (cmp (f, g) \<star>\<^sub>D map h)"
using 1 f g h fg gh D'.arr_char C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def
by (intro D.seqI) auto
moreover have "D'.arr (cmp (f \<star>\<^sub>C g, h) \<cdot>\<^sub>D (cmp (f, g) \<star>\<^sub>D map h))"
using 1 2 D'.arr_char
by (metis (no_types, lifting) D'.comp_char D'.seq_char D.seqE F.\<Phi>.preserves_reflects_arr
preserves_reflects_arr)
ultimately show ?thesis
using f g h fg gh D'.dom_char D'.cod_char D'.comp_char D'.hcomp_def C.VV.arr_char
apply simp
by force
qed
also have "... = \<Phi> (f, g \<star>\<^sub>C h) \<cdot>\<^sub>D (F f \<star>\<^sub>D \<Phi> (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[F f, F g, F h]"
using f g h fg gh F.assoc_coherence [of f g h] by blast
also have "... = D'.comp (\<Phi> (f, g \<star>\<^sub>C h))
(D'.comp (D'.hcomp (F f) (\<Phi> (g, h))) (D'.\<a> (F f) (F g) (F h)))"
proof -
have "D.seq (map f \<star>\<^sub>D cmp (g, h)) \<a>\<^sub>D[map f, map g, map h]"
using f g h fg gh C.VV.arr_char C.VV.dom_char C.VV.cod_char F.FF_def
by (intro D.seqI) auto
moreover have "D'.arr \<a>\<^sub>D[map f, map g, map h]"
using f g h fg gh D'.arr_char preserves_arr by auto
moreover have "D'.arr (map f \<star>\<^sub>D cmp (g, h))"
using f g h fg gh
by (metis (no_types, lifting) D'.arr_char D.seqE D.vseq_implies_hpar(1)
D.vseq_implies_hpar(2) calculation(1-2))
moreover have "D'.arr ((map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[map f, map g, map h])"
using f g h fg gh
by (metis (no_types, lifting) D'.arr_char D'.comp_closed D.seqE
calculation(1-3))
moreover have "D.seq (cmp (f, g \<star>\<^sub>C h))
((map f \<star>\<^sub>D cmp (g, h)) \<cdot>\<^sub>D \<a>\<^sub>D[map f, map g, map h])"
using f g h fg gh F.cmp_simps'(1) F.cmp_simps(4) F.cmp_simps(5) by auto
ultimately show ?thesis
using f g h fg gh C.VV.arr_char D'.VVV.arr_char D'.VV.arr_char D'.comp_char
D'.hcomp_def
by simp
qed
finally show "D'.comp (F \<a>\<^sub>C[f, g, h])
(D'.comp (\<Phi> (f \<star>\<^sub>C g, h)) (D'.hcomp (\<Phi> (f, g)) (F h))) =
D'.comp (\<Phi> (f, g \<star>\<^sub>C h))
(D'.comp (D'.hcomp (F f) (\<Phi> (g, h))) (D'.\<a> (F f) (F g) (F h)))"
by blast
qed
lemma is_pseudofunctor:
shows "pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C D'.comp D'.hcomp D'.\<a> \<i>\<^sub>D D'.src D'.trg F \<Phi>"
..
lemma map\<^sub>0_simp [simp]:
assumes "C.obj a"
shows "map\<^sub>0 a = F.map\<^sub>0 a"
using assms map\<^sub>0_def D'.src_def by auto
lemma unit_char':
assumes "C.obj a"
shows "F.unit a = unit a"
proof (intro unit_eqI)
show "C.obj a" by fact
show 1: "D'.in_hom (F.unit a) (map\<^sub>0 a) (map a)"
proof -
have "D.in_hom (F.unit a) (F.map\<^sub>0 a) (map a)"
using assms by auto
moreover have "map\<^sub>0 a = map\<^sub>0 a"
using assms map\<^sub>0_def D'.src_def by auto
ultimately show ?thesis
using assms D'.in_hom_char preserves_arr D'.arr_char
by (metis (no_types, lifting) C.obj_def' D'.obj_def D'.src_def D.arrI
D.vconn_implies_hpar(1,3) F.unit_simps(2-3) map\<^sub>0_def map\<^sub>0_simps(1))
qed
show "D'.iso (F.unit a)"
using assms D'.iso_char D'.arr_char F.unit_char(2)
\<open>D'.in_hom (F.unit a) (map\<^sub>0 a) (map a)\<close>
by auto
show "D'.comp (F.unit a) \<i>\<^sub>D[map\<^sub>0 a] =
D'.comp (D'.comp (map \<i>\<^sub>C[a]) (cmp (a, a)))
(D'.hcomp (F.unit a) (F.unit a))"
proof -
have "D'.comp (F.unit a) \<i>\<^sub>D[map\<^sub>0 a] = F.unit a \<cdot>\<^sub>D \<i>\<^sub>D[src\<^sub>D (map a)]"
using assms D'.comp_char D'.arr_char
apply simp
by (metis (no_types, lifting) C.obj_simps(1-2) F.preserves_src preserves_arr)
also have "... = (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) \<cdot>\<^sub>D (F.unit a \<star>\<^sub>D F.unit a)"
using assms F.unit_char(3) [of a] by auto
also have "... = D'.comp (D'.comp (map \<i>\<^sub>C[a]) (cmp (a, a)))
(D'.hcomp (F.unit a) (F.unit a))"
proof -
have "D'.arr (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a))"
using assms D'.comp_simp by auto
moreover have "D.seq (map \<i>\<^sub>C[a] \<cdot>\<^sub>D cmp (a, a)) (F.unit a \<star>\<^sub>D F.unit a)"
using assms C.VV.arr_char F.cmp_simps(4-5)
by (intro D.seqI) auto
ultimately show ?thesis
using assms 1 D'.comp_char D'.hcomp_def C.VV.arr_char D'.hseq_char
D'.arr_char F.\<i>_simps(2)
by auto
qed
finally show ?thesis by blast
qed
qed
end
subsection "Equivalence Pseudofunctors"
text \<open>
In this section, we define ``equivalence pseudofunctors'', which are pseudofunctors
that are locally fully faithful, locally essentially surjective, and biessentially
surjective on objects. In a later section, we will show that a pseudofunctor is
an equivalence pseudofunctor if and only if it can be extended to an equivalence
of bicategories.
The definition below requires that an equivalence pseudofunctor be (globally) faithful
with respect to vertical composition. Traditional formulations do not consider a
pseudofunctor as a single global functor, so we have to consider whether this condition
is too strong. In fact, a pseudofunctor (as defined here) is locally faithful if and
only if it is globally faithful.
\<close>
context pseudofunctor
begin
definition locally_faithful
where "locally_faithful \<equiv>
\<forall>f g \<mu> \<mu>'. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> F \<mu> = F \<mu>' \<longrightarrow> \<mu> = \<mu>'"
lemma locally_faithful_iff_faithful:
shows "locally_faithful \<longleftrightarrow> faithful_functor V\<^sub>C V\<^sub>D F"
proof
show "faithful_functor V\<^sub>C V\<^sub>D F \<Longrightarrow> locally_faithful"
proof -
assume 1: "faithful_functor V\<^sub>C V\<^sub>D F"
interpret faithful_functor V\<^sub>C V\<^sub>D F
using 1 by blast
show "locally_faithful"
unfolding locally_faithful_def using is_faithful by blast
qed
show "locally_faithful \<Longrightarrow> faithful_functor V\<^sub>C V\<^sub>D F"
proof -
assume 1: "locally_faithful"
show "faithful_functor V\<^sub>C V\<^sub>D F"
proof
fix \<mu> \<mu>'
assume par: "C.par \<mu> \<mu>'" and eq: "F \<mu> = F \<mu>'"
obtain f g where fg: "\<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C g\<guillemotright> \<and> \<guillemotleft>\<mu>' : f \<Rightarrow>\<^sub>C g\<guillemotright>"
using par by auto
show "\<mu> = \<mu>'"
using 1 fg locally_faithful_def eq by simp
qed
qed
qed
end
text \<open>
In contrast, it is not true that a pseudofunctor that is locally full is also
globally full, because we can have \<open>\<guillemotleft>\<nu> : F h \<Rightarrow>\<^sub>D F k\<guillemotright>\<close> even if \<open>h\<close> and \<open>k\<close>
are not in the same hom-category. So it would be a mistake to require that an
equivalence functor be globally full.
\<close>
locale equivalence_pseudofunctor =
pseudofunctor +
faithful_functor V\<^sub>C V\<^sub>D F +
assumes biessentially_surjective_on_objects:
"D.obj a' \<Longrightarrow> \<exists>a. C.obj a \<and> D.equivalent_objects (map\<^sub>0 a) a'"
and locally_essentially_surjective:
"\<lbrakk> C.obj a; C.obj b; \<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 b\<guillemotright>; D.ide g \<rbrakk> \<Longrightarrow>
\<exists>f. \<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) g"
and locally_full:
"\<lbrakk> C.ide f; C.ide f'; src\<^sub>C f = src\<^sub>C f'; trg\<^sub>C f = trg\<^sub>C f'; \<guillemotleft>\<nu> : F f \<Rightarrow>\<^sub>D F f'\<guillemotright> \<rbrakk> \<Longrightarrow>
\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow>\<^sub>C f'\<guillemotright> \<and> F \<mu> = \<nu>"
begin
lemma reflects_ide:
assumes "C.endo \<mu>" and "D.ide (F \<mu>)"
shows "C.ide \<mu>"
using assms is_faithful [of "C.dom \<mu>" \<mu>] C.ide_char'
by (metis C.arr_dom C.cod_dom C.dom_dom C.seqE D.ide_char preserves_dom)
lemma reflects_iso:
assumes "C.arr \<mu>" and "D.iso (F \<mu>)"
shows "C.iso \<mu>"
proof -
obtain \<mu>' where \<mu>': "\<guillemotleft>\<mu>' : C.cod \<mu> \<Rightarrow>\<^sub>C C.dom \<mu>\<guillemotright> \<and> F \<mu>' = D.inv (F \<mu>)"
using assms locally_full [of "C.cod \<mu>" "C.dom \<mu>" "D.inv (F \<mu>)"]
D.inv_in_hom C.in_homE preserves_hom C.in_homI
by auto
have "C.inverse_arrows \<mu> \<mu>'"
using assms \<mu>' reflects_ide
apply (intro C.inverse_arrowsI)
apply (metis C.cod_comp C.dom_comp C.ide_dom C.in_homE C.seqI D.comp_inv_arr'
faithful_functor_axioms faithful_functor_def functor.preserves_ide
preserves_comp_2 preserves_dom)
by (metis C.cod_comp C.dom_comp C.ide_cod C.in_homE C.seqI D.comp_arr_inv'
faithful_functor_axioms faithful_functor_def functor.preserves_ide
preserves_cod preserves_comp_2)
thus ?thesis by auto
qed
lemma reflects_isomorphic:
assumes "C.ide f" and "C.ide f'" and "src\<^sub>C f = src\<^sub>C f'" and "trg\<^sub>C f = trg\<^sub>C f'"
and "D.isomorphic (F f) (F f')"
shows "C.isomorphic f f'"
using assms C.isomorphic_def D.isomorphic_def locally_full reflects_iso C.arrI
by metis
lemma reflects_equivalence:
assumes "C.ide f" and "C.ide g"
and "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>" and "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
and "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
(D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f))
(D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g))"
shows "equivalence_in_bicategory V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
proof -
interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close>
\<open>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)\<close>
\<open>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)\<close>
using assms by auto
show ?thesis
proof
show "C.ide f"
using assms(1) by simp
show "C.ide g"
using assms(2) by simp
show "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>"
using assms(3) by simp
show "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>"
using assms(4) by simp
have 0: "src\<^sub>C f = trg\<^sub>C g \<and> src\<^sub>C g = trg\<^sub>C f"
- using `\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>`
+ using \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close>
by (metis C.hseqE C.ideD(1) C.ide_cod C.ide_dom C.in_homE assms(4))
show "C.iso \<eta>"
proof -
have "D.iso (F \<eta>)"
proof -
have 1: "\<guillemotleft>D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f) : map\<^sub>0 (src\<^sub>C f) \<Rightarrow>\<^sub>D F g \<star>\<^sub>D F f\<guillemotright>"
- using `C.ide f` `C.ide g` `\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>`
+ using \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close>
unit_char cmp_in_hom cmp_components_are_iso
by (metis (mono_tags, lifting) C.ideD(1) E'.unit_in_vhom preserves_src)
have 2: "D.iso (\<Phi> (g, f)) \<and> \<guillemotleft>\<Phi> (g, f) : F g \<star>\<^sub>D F f \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
- using 0 `C.ide f` `C.ide g` cmp_in_hom by simp
+ using 0 \<open>C.ide f\<close> \<open>C.ide g\<close> cmp_in_hom by simp
have 3: "D.iso (D.inv (unit (src\<^sub>C f))) \<and>
\<guillemotleft>D.inv (unit (src\<^sub>C f)) : F (src\<^sub>C f) \<Rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C f)\<guillemotright>"
- using `C.ide f` unit_char by simp
+ using \<open>C.ide f\<close> unit_char by simp
have "D.iso (\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)))"
using 1 2 3 E'.unit_is_iso D.isos_compose by blast
moreover have "\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)) =
F \<eta>"
proof -
have "\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f)) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f))
= (\<Phi> (g, f) \<cdot>\<^sub>D (D.inv (\<Phi> (g, f))) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D (unit (src\<^sub>C f)) \<cdot>\<^sub>D
D.inv (unit (src\<^sub>C f)))"
using D.comp_assoc by simp
also have "... = F (g \<star>\<^sub>C f) \<cdot>\<^sub>D F \<eta> \<cdot>\<^sub>D F (src\<^sub>C f)"
using 2 3 D.comp_arr_inv D.comp_inv_arr D.inv_is_inverse
by (metis C.ideD(1) C.obj_src D.comp_assoc D.dom_inv D.in_homE unit_char(2)
assms(1))
also have "... = F \<eta>"
- using `\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>` D.comp_arr_dom D.comp_cod_arr by auto
+ using \<open>\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright>\<close> D.comp_arr_dom D.comp_cod_arr by auto
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
thus ?thesis
using assms reflects_iso by auto
qed
show "C.iso \<epsilon>"
proof -
have "D.iso (F \<epsilon>)"
proof -
have 1: "\<guillemotleft>D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g) : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D map\<^sub>0 (src\<^sub>C g)\<guillemotright>"
- using `C.ide f` `C.ide g` `\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>`
+ using \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close>
unit_char cmp_in_hom cmp_components_are_iso
by (metis (mono_tags, lifting) C.ideD(1) E'.counit_in_vhom preserves_src)
have 2: "D.iso (D.inv (\<Phi> (f, g))) \<and>
\<guillemotleft>D.inv (\<Phi> (f, g)) : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F g\<guillemotright>"
- using 0 `C.ide f` `C.ide g` `\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>` cmp_in_hom(2) D.inv_in_hom
+ using 0 \<open>C.ide f\<close> \<open>C.ide g\<close> \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close> cmp_in_hom(2) D.inv_in_hom
by simp
have 3: "D.iso (unit (trg\<^sub>C f)) \<and> \<guillemotleft>unit (trg\<^sub>C f) : map\<^sub>0 (trg\<^sub>C f) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
- using `C.ide f` unit_char by simp
+ using \<open>C.ide f\<close> unit_char by simp
have "D.iso (unit (trg\<^sub>C f) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g)))"
using 0 1 2 3 E'.counit_is_iso D.isos_compose
by (metis D.arrI D.cod_comp D.cod_inv D.seqI D.seqI')
moreover have "unit (trg\<^sub>C f) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g)) =
F \<epsilon>"
proof -
have "unit (trg\<^sub>C f) \<cdot>\<^sub>D (D.inv (unit (trg\<^sub>C f)) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D \<Phi> (f, g)) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g)) =
(unit (trg\<^sub>C f) \<cdot>\<^sub>D D.inv (unit (trg\<^sub>C f))) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D (\<Phi> (f, g) \<cdot>\<^sub>D D.inv (\<Phi> (f, g)))"
using D.comp_assoc by simp
also have "... = F (trg\<^sub>C f) \<cdot>\<^sub>D F \<epsilon> \<cdot>\<^sub>D F (f \<star>\<^sub>C g)"
using 0 3 D.comp_arr_inv' D.comp_inv_arr'
by (simp add: C.VV.arr_char C.VV.ide_char assms(1-2))
also have "... = F \<epsilon>"
- using 0 `\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>` D.comp_arr_dom D.comp_cod_arr by auto
+ using 0 \<open>\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright>\<close> D.comp_arr_dom D.comp_cod_arr by auto
finally show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
thus ?thesis
using assms reflects_iso by auto
qed
qed
qed
lemma reflects_equivalence_map:
assumes "C.ide f" and "D.equivalence_map (F f)"
shows "C.equivalence_map f"
proof -
obtain g' \<phi> \<psi> where E': "equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) g' \<phi> \<psi>"
using assms D.equivalence_map_def by auto
interpret E': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> g' \<phi> \<psi>
using E' by auto
obtain g where g: "\<guillemotleft>g : trg\<^sub>C f \<rightarrow>\<^sub>C src\<^sub>C f\<guillemotright> \<and> C.ide g \<and> D.isomorphic (F g) g'"
using assms E'.antipar locally_essentially_surjective [of "trg\<^sub>C f" "src\<^sub>C f" g']
by auto
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : g' \<Rightarrow>\<^sub>D F g\<guillemotright> \<and> D.iso \<mu>"
using g D.isomorphic_def D.isomorphic_symmetric by blast
have E'': "equivalence_in_bicategory (\<cdot>\<^sub>D) (\<star>\<^sub>D) \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D (F f) (F g)
((F g \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D F f) \<cdot>\<^sub>D \<phi>)
(\<psi> \<cdot>\<^sub>D (D.inv (F f) \<star>\<^sub>D g') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<mu>))"
using assms \<mu> E'.equivalence_in_bicategory_axioms D.ide_is_iso
D.equivalence_respects_iso [of "F f" g' \<phi> \<psi> "F f" "F f" \<mu> "F g"]
by auto
interpret E'': equivalence_in_bicategory V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D \<open>F f\<close> \<open>F g\<close>
\<open>(F g \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D F f) \<cdot>\<^sub>D \<phi>\<close>
\<open>\<psi> \<cdot>\<^sub>D (D.inv (F f) \<star>\<^sub>D g') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<mu>)\<close>
using E'' by auto
let ?\<eta>' = "\<Phi> (g, f) \<cdot>\<^sub>D (F g \<star>\<^sub>D F f) \<cdot>\<^sub>D (\<mu> \<star>\<^sub>D F f) \<cdot>\<^sub>D \<phi> \<cdot>\<^sub>D D.inv (unit (src\<^sub>C f))"
have \<eta>': "\<guillemotleft>?\<eta>' : F (src\<^sub>C f) \<Rightarrow>\<^sub>D F (g \<star>\<^sub>C f)\<guillemotright>"
using assms \<mu> g unit_char E'.unit_in_hom(2) E'.antipar E''.antipar cmp_in_hom
apply (intro D.comp_in_homI)
apply auto
using E'.antipar(2) by blast
have iso_\<eta>': "D.iso ?\<eta>'"
using assms g \<mu> \<eta>' E'.antipar E''.antipar cmp_in_hom unit_char
E'.unit_in_hom E'.unit_is_iso
by (intro D.isos_compose) auto
let ?\<epsilon>' = "unit (src\<^sub>C g) \<cdot>\<^sub>D \<psi> \<cdot>\<^sub>D (D.inv (F f) \<star>\<^sub>D g') \<cdot>\<^sub>D (F f \<star>\<^sub>D D.inv \<mu>) \<cdot>\<^sub>D
D.inv (\<Phi> (f, g))"
have \<epsilon>': "\<guillemotleft>?\<epsilon>' : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
proof (intro D.comp_in_homI)
show "\<guillemotleft>D.inv (\<Phi> (f, g)) : F (f \<star>\<^sub>C g) \<Rightarrow>\<^sub>D F f \<star>\<^sub>D F g\<guillemotright>"
using assms g cmp_in_hom C.VV.ide_char C.VV.arr_char by auto
show "\<guillemotleft>F f \<star>\<^sub>D D.inv \<mu> : F f \<star>\<^sub>D F g \<Rightarrow>\<^sub>D F f \<star>\<^sub>D g'\<guillemotright>"
using assms g \<mu> E'.antipar E''.antipar D.ide_in_hom(2) by auto
show "\<guillemotleft>D.inv (F f) \<star>\<^sub>D g' : F f \<star>\<^sub>D g' \<Rightarrow>\<^sub>D F f \<star>\<^sub>D g'\<guillemotright>"
using assms E'.antipar E''.antipar D.ide_is_iso
by (simp add: D.ide_in_hom(2))
show "\<guillemotleft>\<psi> : F f \<star>\<^sub>D g' \<Rightarrow>\<^sub>D trg\<^sub>D (F f)\<guillemotright>"
using E'.counit_in_hom by simp
show "\<guillemotleft>unit (src\<^sub>C g) : trg\<^sub>D (F f) \<Rightarrow>\<^sub>D F (trg\<^sub>C f)\<guillemotright>"
using assms g unit_char E'.antipar by auto
qed
have iso_\<epsilon>': "D.iso ?\<epsilon>'"
proof -
have "D.iso (\<Phi> (f, g))"
using assms g C.VV.ide_char C.VV.arr_char by auto
thus ?thesis
using assms g \<mu> E'.antipar E''.antipar cmp_in_hom unit_char
E'.counit_in_hom D.iso_inv_iso E'.counit_is_iso \<epsilon>'
by (metis C.ideD(1) C.obj_src D.arrI D.iso_hcomp D.hseq_char D.ide_is_iso
D.isos_compose D.seqE E'.ide_right components_are_iso)
qed
obtain \<eta> where \<eta>: "\<guillemotleft>\<eta> : src\<^sub>C f \<Rightarrow>\<^sub>C g \<star>\<^sub>C f\<guillemotright> \<and> F \<eta> = ?\<eta>'"
using assms g E'.antipar \<eta>' locally_full [of "src\<^sub>C f" "g \<star>\<^sub>C f" ?\<eta>']
by (metis C.ide_hcomp C.ideD(1) C.in_hhomE C.src.preserves_ide C.hcomp_simps(1-2)
C.src_trg C.trg_trg)
have iso_\<eta>: "C.iso \<eta>"
using \<eta> \<eta>' iso_\<eta>' reflects_iso by auto
have 1: "\<exists>\<epsilon>. \<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> F \<epsilon> = ?\<epsilon>'"
using assms g \<epsilon>' locally_full [of "f \<star>\<^sub>C g" "src\<^sub>C g" ?\<epsilon>'] by force
obtain \<epsilon> where \<epsilon>: "\<guillemotleft>\<epsilon> : f \<star>\<^sub>C g \<Rightarrow>\<^sub>C src\<^sub>C g\<guillemotright> \<and> F \<epsilon> = ?\<epsilon>'"
using 1 by blast
have iso_\<epsilon>: "C.iso \<epsilon>"
using \<epsilon> \<epsilon>' iso_\<epsilon>' reflects_iso by auto
have "equivalence_in_bicategory (\<cdot>\<^sub>C) (\<star>\<^sub>C) \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C f g \<eta> \<epsilon>"
using assms g \<eta> \<epsilon> iso_\<eta> iso_\<epsilon> by (unfold_locales, auto)
thus ?thesis
using C.equivalence_map_def by auto
qed
lemma reflects_equivalent_objects:
assumes "C.obj a" and "C.obj b" and "D.equivalent_objects (map\<^sub>0 a) (map\<^sub>0 b)"
shows "C.equivalent_objects a b"
proof -
obtain f' where f': "\<guillemotleft>f' : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 b\<guillemotright> \<and> D.equivalence_map f'"
using assms D.equivalent_objects_def D.equivalence_map_def by auto
obtain f where f: "\<guillemotleft>f : a \<rightarrow>\<^sub>C b\<guillemotright> \<and> C.ide f \<and> D.isomorphic (F f) f'"
using assms f' locally_essentially_surjective [of a b f'] D.equivalence_map_is_ide
by auto
have "D.equivalence_map (F f)"
using f f' D.equivalence_map_preserved_by_iso [of f' "F f"] D.isomorphic_symmetric
by simp
hence "C.equivalence_map f"
using f f' reflects_equivalence_map [of f] by simp
thus ?thesis
using f C.equivalent_objects_def by auto
qed
end
text\<open>
For each pair of objects \<open>a\<close>, \<open>b\<close> of \<open>C\<close>, an equivalence pseudofunctor restricts
to an equivalence of categories between \<open>C.hhom a b\<close> and \<open>D.hhom (map\<^sub>0 a) (map\<^sub>0 b)\<close>.
\<close>
(* TODO: Change the "perspective" of this locale to be the defined functor. *)
locale equivalence_pseudofunctor_at_hom =
equivalence_pseudofunctor +
fixes a :: 'a and a' :: 'a
assumes obj_a: "C.obj a"
and obj_a': "C.obj a'"
begin
sublocale hhom\<^sub>C: subcategory V\<^sub>C \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : a \<rightarrow>\<^sub>C a'\<guillemotright>\<close>
using C.hhom_is_subcategory by simp
sublocale hhom\<^sub>D: subcategory V\<^sub>D \<open>\<lambda>\<mu>. \<guillemotleft>\<mu> : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>\<close>
using D.hhom_is_subcategory by simp
definition F\<^sub>1
where "F\<^sub>1 = (\<lambda>\<mu>. if hhom\<^sub>C.arr \<mu> then F \<mu> else D.null)"
interpretation F\<^sub>1: "functor" hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1
unfolding F\<^sub>1_def
using hhom\<^sub>C.arr_char hhom\<^sub>D.arr_char hhom\<^sub>C.dom_char hhom\<^sub>D.dom_char
hhom\<^sub>C.cod_char hhom\<^sub>D.cod_char hhom\<^sub>C.seq_char hhom\<^sub>C.comp_char hhom\<^sub>D.comp_char
by unfold_locales auto
interpretation F\<^sub>1: fully_faithful_and_essentially_surjective_functor
hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1
proof
show "\<And>\<mu> \<mu>'. \<lbrakk>hhom\<^sub>C.par \<mu> \<mu>'; F\<^sub>1 \<mu> = F\<^sub>1 \<mu>'\<rbrakk> \<Longrightarrow> \<mu> = \<mu>'"
unfolding F\<^sub>1_def
using is_faithful hhom\<^sub>C.dom_char hhom\<^sub>D.dom_char hhom\<^sub>C.cod_char hhom\<^sub>D.cod_char
by (metis C.in_hhom_def hhom\<^sub>C.arrE)
show "\<And>f f' \<nu>. \<lbrakk>hhom\<^sub>C.ide f; hhom\<^sub>C.ide f'; hhom\<^sub>D.in_hom \<nu> (F\<^sub>1 f') (F\<^sub>1 f)\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> F\<^sub>1 \<mu> = \<nu>"
proof (unfold F\<^sub>1_def)
fix f f' \<nu>
assume f: "hhom\<^sub>C.ide f" and f': "hhom\<^sub>C.ide f'"
assume "hhom\<^sub>D.in_hom \<nu> (if hhom\<^sub>C.arr f' then F f' else D.null)
(if hhom\<^sub>C.arr f then F f else D.null)"
hence \<nu>: "hhom\<^sub>D.in_hom \<nu> (F f') (F f)"
using f f' by simp
have "\<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
proof -
have 1: "src\<^sub>C f' = src\<^sub>C f \<and> trg\<^sub>C f' = trg\<^sub>C f"
using f f' hhom\<^sub>C.ide_char by (metis C.in_hhomE hhom\<^sub>C.arrE)
hence ex: "\<exists>\<mu>. C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
using f f' \<nu> hhom\<^sub>C.in_hom_char hhom\<^sub>D.in_hom_char hhom\<^sub>C.ide_char locally_full
by simp
obtain \<mu> where \<mu>: "C.in_hom \<mu> f' f \<and> F \<mu> = \<nu>"
using ex by blast
have "hhom\<^sub>C.in_hom \<mu> f' f"
proof -
have 2: "hhom\<^sub>C.arr f' \<and> hhom\<^sub>C.arr f"
using f f' hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char by simp
moreover have "hhom\<^sub>C.arr \<mu>"
using \<mu> 1 2 hhom\<^sub>C.arr_char C.arrI C.vconn_implies_hpar(1-2) by auto
ultimately show ?thesis
using f f' \<mu> hhom\<^sub>C.arr_char hhom\<^sub>C.ide_char hhom\<^sub>C.in_hom_char by simp
qed
thus ?thesis
using \<mu> by auto
qed
thus "\<exists>\<mu>. hhom\<^sub>C.in_hom \<mu> f' f \<and> (if hhom\<^sub>C.arr \<mu> then F \<mu> else D.null) = \<nu>"
by auto
qed
show "\<And>g. hhom\<^sub>D.ide g \<Longrightarrow> \<exists>f. hhom\<^sub>C.ide f \<and> hhom\<^sub>D.isomorphic (F\<^sub>1 f) g"
proof (unfold F\<^sub>1_def)
fix g
assume g: "hhom\<^sub>D.ide g"
show "\<exists>f. hhom\<^sub>C.ide f \<and> hhom\<^sub>D.isomorphic (if hhom\<^sub>C.arr f then F f else D.null) g"
proof -
have "C.obj a \<and> C.obj a'"
using obj_a obj_a' by simp
moreover have 1: "D.ide g \<and> \<guillemotleft>g : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>"
using g obj_a obj_a' hhom\<^sub>D.ide_char by auto
ultimately have 2: "\<exists>f. C.in_hhom f a a' \<and> C.ide f \<and> D.isomorphic (F f) g"
using locally_essentially_surjective [of a a' g] by simp
obtain f \<phi> where f: "C.in_hhom f a a' \<and> C.ide f \<and> D.in_hom \<phi> (F f) g \<and> D.iso \<phi>"
using 2 by auto
have "hhom\<^sub>C.ide f"
using f hhom\<^sub>C.ide_char hhom\<^sub>C.arr_char by simp
moreover have "hhom\<^sub>D.isomorphic (F f) g"
proof -
have "hhom\<^sub>D.arr \<phi> \<and> hhom\<^sub>D.arr (D.inv \<phi>)"
using f 1 hhom\<^sub>D.arr_char D.in_hhom_def by auto
hence "hhom\<^sub>D.in_hom \<phi> (F f) g \<and> hhom\<^sub>D.iso \<phi>"
using f g hhom\<^sub>D.in_hom_char hhom\<^sub>D.iso_char hhom\<^sub>C.arr_char hhom\<^sub>D.arr_char
hhom\<^sub>D.ideD(1) preserves_arr
by (simp add: C.in_hhom_def)
thus ?thesis
unfolding hhom\<^sub>D.isomorphic_def by blast
qed
ultimately show "\<exists>f. hhom\<^sub>C.ide f \<and>
hhom\<^sub>D.isomorphic (if hhom\<^sub>C.arr f then F f else D.null) g"
by force
qed
qed
qed
lemma equivalence_functor_F\<^sub>1:
shows "fully_faithful_and_essentially_surjective_functor hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1"
and "equivalence_functor hhom\<^sub>C.comp hhom\<^sub>D.comp F\<^sub>1"
..
definition G\<^sub>1
where "G\<^sub>1 = (SOME G. \<exists>\<eta>\<epsilon>.
adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>))"
lemma G\<^sub>1_props:
assumes "C.obj a" and "C.obj a'"
shows "\<exists>\<eta> \<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"
proof -
have "\<exists>G. \<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"
using F\<^sub>1.extends_to_adjoint_equivalence by simp
hence "\<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"
unfolding G\<^sub>1_def
using someI_ex
[of "\<lambda>G. \<exists>\<eta>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G F\<^sub>1 (fst \<eta>\<epsilon>) (snd \<eta>\<epsilon>)"]
by blast
thus ?thesis by simp
qed
definition \<eta>
where "\<eta> = (SOME \<eta>. \<exists>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>)"
definition \<epsilon>
where "\<epsilon> = (SOME \<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>)"
lemma \<eta>\<epsilon>_props:
shows "adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"
using obj_a obj_a' \<eta>_def \<epsilon>_def G\<^sub>1_props
someI_ex [of "\<lambda>\<eta>. \<exists>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"]
someI_ex [of "\<lambda>\<epsilon>. adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>"]
by simp
sublocale \<eta>\<epsilon>: adjoint_equivalence hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta> \<epsilon>
using \<eta>\<epsilon>_props by simp
sublocale \<eta>\<epsilon>: meta_adjunction hhom\<^sub>C.comp hhom\<^sub>D.comp G\<^sub>1 F\<^sub>1 \<eta>\<epsilon>.\<phi> \<eta>\<epsilon>.\<psi>
using \<eta>\<epsilon>.induces_meta_adjunction by simp
end
context identity_pseudofunctor
begin
sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
map cmp
proof
show "\<And>f f'. \<lbrakk>B.par f f'; map f = map f'\<rbrakk> \<Longrightarrow> f = f'"
by simp
show "\<And>a'. B.obj a' \<Longrightarrow> \<exists>a. B.obj a \<and> B.equivalent_objects (map\<^sub>0 a) a'"
by (auto simp add: B.equivalent_objects_reflexive map\<^sub>0_def B.obj_simps)
show "\<And>a b g. \<lbrakk>B.obj a; B.obj b; B.in_hhom g (map\<^sub>0 a) (map\<^sub>0 b); B.ide g\<rbrakk>
\<Longrightarrow> \<exists>f. B.in_hhom f a b \<and> B.ide f \<and> B.isomorphic (map f) g"
using B.isomorphic_reflexive map\<^sub>0_def B.obj_simps by auto
show "\<And>f f' \<nu>. \<lbrakk>B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f';
\<guillemotleft>\<nu> : map f \<rightarrow>\<^sub>B map f'\<guillemotright>\<rbrakk>
\<Longrightarrow> \<exists>\<mu>. \<guillemotleft>\<mu> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> map \<mu> = \<nu>"
using B.arrI by auto
qed
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B
map cmp"
..
end
locale composite_equivalence_pseudofunctor =
composite_pseudofunctor +
F: equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C F \<Phi>\<^sub>F +
G: equivalence_pseudofunctor V\<^sub>C H\<^sub>C \<a>\<^sub>C \<i>\<^sub>C src\<^sub>C trg\<^sub>C V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D G \<Phi>\<^sub>G
begin
interpretation faithful_functor V\<^sub>B V\<^sub>D \<open>G o F\<close>
using F.faithful_functor_axioms G.faithful_functor_axioms faithful_functors_compose
by blast
interpretation equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>G o F\<close> cmp
proof
show "\<And>c. D.obj c \<Longrightarrow> \<exists>a. B.obj a \<and> D.equivalent_objects (map\<^sub>0 a) c"
proof -
fix c
assume c: "D.obj c"
obtain b where b: "C.obj b \<and> D.equivalent_objects (G.map\<^sub>0 b) c"
using c G.biessentially_surjective_on_objects by auto
obtain a where a: "B.obj a \<and> C.equivalent_objects (F.map\<^sub>0 a) b"
using b F.biessentially_surjective_on_objects by auto
have "D.equivalent_objects (map\<^sub>0 a) c"
using a b map\<^sub>0_def G.preserves_equivalent_objects D.equivalent_objects_transitive
by fastforce
thus "\<exists>a. B.obj a \<and> D.equivalent_objects (map\<^sub>0 a) c"
using a by auto
qed
show "\<And>a a' h. \<lbrakk>B.obj a; B.obj a'; \<guillemotleft>h : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>; D.ide h\<rbrakk>
\<Longrightarrow> \<exists>f. B.in_hhom f a a' \<and> B.ide f \<and> D.isomorphic ((G \<circ> F) f) h"
proof -
fix a a' h
assume a: "B.obj a" and a': "B.obj a'"
and h_in_hom: "\<guillemotleft>h : map\<^sub>0 a \<rightarrow>\<^sub>D map\<^sub>0 a'\<guillemotright>" and ide_h: "D.ide h"
obtain g
where g: "C.in_hhom g (F.map\<^sub>0 a) (F.map\<^sub>0 a') \<and> C.ide g \<and> D.isomorphic (G g) h"
using a a' h_in_hom ide_h map\<^sub>0_def B.obj_simps
G.locally_essentially_surjective [of "F.map\<^sub>0 a" "F.map\<^sub>0 a'" h]
by auto
obtain f where f: "B.in_hhom f a a' \<and> B.ide f \<and> C.isomorphic (F f) g"
using a a' g F.locally_essentially_surjective by blast
have "D.isomorphic ((G o F) f) h"
proof -
have "(G o F) f = G (F f)"
by simp
also have "D.isomorphic ... (G g)"
using f G.preserves_iso D.isomorphic_def by blast
also have "D.isomorphic ... h"
using g by simp
finally show "D.isomorphic ((G \<circ> F) f) h" by simp
qed
thus "\<exists>f. B.in_hhom f a a' \<and> B.ide f \<and> D.isomorphic ((G \<circ> F) f) h"
using f by auto
qed
show "\<And>f f' \<nu>. \<lbrakk>B.ide f; B.ide f'; src\<^sub>B f = src\<^sub>B f'; trg\<^sub>B f = trg\<^sub>B f';
\<guillemotleft>\<nu> : (G \<circ> F) f \<Rightarrow>\<^sub>D (G \<circ> F) f'\<guillemotright>\<rbrakk>
\<Longrightarrow> \<exists>\<tau>. \<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> (G \<circ> F) \<tau> = \<nu>"
proof -
fix f f' \<nu>
assume f: "B.ide f" and f': "B.ide f'"
and src: "src\<^sub>B f = src\<^sub>B f'" and trg: "trg\<^sub>B f = trg\<^sub>B f'"
and \<nu>: "\<guillemotleft>\<nu> : (G \<circ> F) f \<Rightarrow>\<^sub>D (G \<circ> F) f'\<guillemotright>"
have \<nu>: "\<guillemotleft>\<nu> : G (F f) \<Rightarrow>\<^sub>D G (F f')\<guillemotright>"
using \<nu> by simp
have 1: "src\<^sub>C (F f) = src\<^sub>C (F f') \<and> trg\<^sub>C (F f) = trg\<^sub>C (F f')"
using f f' src trg by simp
have 2: "\<exists>\<mu>. \<guillemotleft>\<mu> : F f \<Rightarrow>\<^sub>C F f'\<guillemotright> \<and> G \<mu> = \<nu>"
using f f' 1 \<nu> G.locally_full F.preserves_ide by simp
obtain \<mu> where \<mu>: "\<guillemotleft>\<mu> : F f \<Rightarrow>\<^sub>C F f'\<guillemotright> \<and> G \<mu> = \<nu>"
using 2 by auto
obtain \<tau> where \<tau>: "\<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> F \<tau> = \<mu>"
using f f' src trg 2 \<mu> F.locally_full by blast
show "\<exists>\<tau>. \<guillemotleft>\<tau> : f \<rightarrow>\<^sub>B f'\<guillemotright> \<and> (G \<circ> F) \<tau> = \<nu>"
using \<mu> \<tau> by auto
qed
qed
sublocale equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
\<open>G o F\<close> cmp ..
lemma is_equivalence_pseudofunctor:
shows "equivalence_pseudofunctor V\<^sub>B H\<^sub>B \<a>\<^sub>B \<i>\<^sub>B src\<^sub>B trg\<^sub>B V\<^sub>D H\<^sub>D \<a>\<^sub>D \<i>\<^sub>D src\<^sub>D trg\<^sub>D
(G o F) cmp"
..
end
end
diff --git a/thys/Bicategory/Subbicategory.thy b/thys/Bicategory/Subbicategory.thy
--- a/thys/Bicategory/Subbicategory.thy
+++ b/thys/Bicategory/Subbicategory.thy
@@ -1,1482 +1,1482 @@
(* Title: Subbicategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2019
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
section "Sub-Bicategories"
text \<open>
In this section we give a construction of a sub-bicategory in terms of a predicate
on the arrows of an ambient bicategory that has certain closure properties with respect
to that bicategory. While the construction given here is likely to be of general use,
it is not the most general sub-bicategory construction that one could imagine,
because it requires that the sub-bicategory actually contain the unit and associativity
isomorphisms of the ambient bicategory. Our main motivation for including this construction
here is to apply it to exploit the fact that the sub-bicategory of endo-arrows of a fixed
object is a monoidal category, which will enable us to transfer to bicategories a result
about unit isomorphisms in monoidal categories.
\<close>
theory Subbicategory
imports Bicategory
begin
subsection "Construction"
locale subbicategory =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
subcategory V Arr
for V :: "'a comp" (infixr "\<cdot>\<^sub>B" 55)
and H :: "'a comp" (infixr "\<star>\<^sub>B" 55)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
and Arr :: "'a \<Rightarrow> bool" +
assumes src_closed: "Arr f \<Longrightarrow> Arr (src\<^sub>B f)"
and trg_closed: "Arr f \<Longrightarrow> Arr (trg\<^sub>B f)"
and hcomp_closed: "\<lbrakk> Arr f; Arr g; trg\<^sub>B f = src\<^sub>B g \<rbrakk> \<Longrightarrow> Arr (g \<star>\<^sub>B f)"
and assoc_closed: "\<lbrakk> Arr f \<and> B.ide f; Arr g \<and> B.ide g; Arr h \<and> B.ide h;
src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \<rbrakk> \<Longrightarrow> Arr (\<a>\<^sub>B f g h)"
and assoc'_closed: "\<lbrakk> Arr f \<and> B.ide f; Arr g \<and> B.ide g; Arr h \<and> B.ide h;
src\<^sub>B f = trg\<^sub>B g; src\<^sub>B g = trg\<^sub>B h \<rbrakk> \<Longrightarrow> Arr (B.inv (\<a>\<^sub>B f g h))"
and lunit_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.\<ll> f)"
and lunit'_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.inv (B.\<ll> f))"
and runit_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.\<rr> f)"
and runit'_closed: "\<lbrakk> Arr f; B.ide f \<rbrakk> \<Longrightarrow> Arr (B.inv (B.\<rr> f))"
begin
notation B.in_hom ("\<guillemotleft>_ : _ \<Rightarrow>\<^sub>B _\<guillemotright>")
notation comp (infixr "\<cdot>" 55)
definition hcomp (infixr "\<star>" 53)
where "g \<star> f = (if arr f \<and> arr g \<and> trg\<^sub>B f = src\<^sub>B g then g \<star>\<^sub>B f else null)"
definition src
where "src \<mu> = (if arr \<mu> then src\<^sub>B \<mu> else null)"
definition trg
where "trg \<mu> = (if arr \<mu> then trg\<^sub>B \<mu> else null)"
interpretation src: endofunctor \<open>(\<cdot>)\<close> src
using src_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed
dom_simp cod_simp
apply unfold_locales
apply auto[4]
by (metis B.src.preserves_comp_2 comp_char seq_char)
interpretation trg: endofunctor \<open>(\<cdot>)\<close> trg
using trg_def null_char inclusion arr_char src_closed trg_closed dom_closed cod_closed
dom_simp cod_simp
apply unfold_locales
apply auto[4]
by (metis B.trg.preserves_comp_2 comp_char seq_char)
interpretation horizontal_homs \<open>(\<cdot>)\<close> src trg
using src_def trg_def src.preserves_arr trg.preserves_arr null_char ide_char arr_char
inclusion
by (unfold_locales, simp_all)
interpretation "functor" VV.comp \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
using hcomp_def VV.arr_char src_def trg_def arr_char hcomp_closed dom_char cod_char
VV.dom_char VV.cod_char
apply unfold_locales
apply auto[2]
proof -
fix f
assume f: "VV.arr f"
show "dom (fst f \<star> snd f) = fst (VV.dom f) \<star> snd (VV.dom f)"
proof -
have "dom (fst f \<star> snd f) = B.dom (fst f) \<star>\<^sub>B B.dom (snd f)"
proof -
have "dom (fst f \<star> snd f) = B.dom (fst f \<star> snd f)"
using f dom_char
by (simp add: arr_char hcomp_closed hcomp_def)
also have "... = B.dom (fst f) \<star>\<^sub>B B.dom (snd f)"
using f
by (metis (no_types, lifting) B.hcomp_simps(3) B.hseqI' VV.arrE arrE hcomp_def
inclusion src_def trg_def)
finally show ?thesis by blast
qed
also have "... = fst (VV.dom f) \<star> snd (VV.dom f)"
using f VV.arr_char VV.dom_char arr_char hcomp_def B.seq_if_composable dom_closed
dom_simp
by (simp, metis)
finally show ?thesis by simp
qed
show "cod (fst f \<star> snd f) = fst (VV.cod f) \<star> snd (VV.cod f)"
proof -
have "cod (fst f \<star> snd f) = B.cod (fst f) \<star>\<^sub>B B.cod (snd f)"
using f VV.arr_char arr_char cod_char hcomp_def src_def trg_def
src_closed trg_closed hcomp_closed inclusion B.hseq_char arrE
by auto
also have "... = fst (VV.cod f) \<star> snd (VV.cod f)"
using f VV.arr_char VV.cod_char arr_char hcomp_def B.seq_if_composable cod_closed
cod_simp
by (simp, metis)
finally show ?thesis by simp
qed
next
fix f g
assume fg: "VV.seq g f"
show "fst (VV.comp g f) \<star> snd (VV.comp g f) = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
proof -
have "fst (VV.comp g f) \<star> snd (VV.comp g f) = fst g \<cdot> fst f \<star> snd g \<cdot> snd f"
using fg VV.seq_char VV.comp_char VxV.comp_char VxV.not_Arr_Null
by (metis (no_types, lifting) VxV.seqE prod.sel(1) prod.sel(2))
also have "... = (fst g \<cdot>\<^sub>B fst f) \<star>\<^sub>B (snd g \<cdot>\<^sub>B snd f)"
using fg comp_char hcomp_def VV.seq_char inclusion arr_char seq_char B.hseq_char
by (metis (no_types, lifting) B.hseq_char' VxV.seq_char null_char)
also have 1: "... = (fst g \<star>\<^sub>B snd g) \<cdot>\<^sub>B (fst f \<star>\<^sub>B snd f)"
proof -
have "src\<^sub>B (fst g) = trg\<^sub>B (snd g)"
by (metis (no_types, lifting) VV.arrE VV.seq_char fg src_def trg_def)
thus ?thesis
using fg VV.seq_char VV.arr_char arr_char seq_char inclusion B.interchange
by (meson VxV.seqE)
qed
also have "... = (fst g \<star> snd g) \<cdot> (fst f \<star> snd f)"
using fg comp_char hcomp_def VV.seq_char VV.arr_char arr_char seq_char inclusion
B.hseq_char' hcomp_closed src_def trg_def
by (metis (no_types, lifting) 1)
finally show ?thesis by auto
qed
qed
interpretation horizontal_composition \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> src trg
using arr_char src_def trg_def src_closed trg_closed
apply (unfold_locales)
using hcomp_def inclusion not_arr_null by auto
abbreviation \<a>
where "\<a> \<mu> \<nu> \<tau> \<equiv> if VVV.arr (\<mu>, \<nu>, \<tau>) then \<a>\<^sub>B \<mu> \<nu> \<tau> else null"
abbreviation (input) \<alpha>\<^sub>S\<^sub>B
where "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> \<equiv> \<a> (fst \<mu>\<nu>\<tau>) (fst (snd \<mu>\<nu>\<tau>)) (snd (snd \<mu>\<nu>\<tau>))"
lemma assoc_closed':
assumes "VVV.arr \<mu>\<nu>\<tau>"
shows "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>)"
proof -
have 1: "B.VVV.arr \<mu>\<nu>\<tau>"
using assms VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char
src_def trg_def inclusion
by auto
show "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>)"
proof -
have "Arr (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) =
Arr ((fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)) \<cdot>\<^sub>B \<alpha>\<^sub>S\<^sub>B (B.VVV.dom \<mu>\<nu>\<tau>))"
using assms B.\<alpha>_def 1 B.VVV.arr_char B.VV.arr_char B.VVV.dom_char B.VV.dom_char
B.assoc_is_natural_1 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
VV.arr_char VVV.arr_char arr_dom src_dom trg_dom dom_simp
by auto
also have "..."
proof (intro comp_closed)
show "Arr (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>))"
using assms 1 B.VVV.arr_char B.VV.arr_char hcomp_closed
by (metis (no_types, lifting) B.H.preserves_reflects_arr B.trg_hcomp
VV.arr_char VVV.arrE arr_char)
show "B.cod (\<a> (fst (B.VVV.dom \<mu>\<nu>\<tau>)) (fst (snd (B.VVV.dom \<mu>\<nu>\<tau>)))
(snd (snd (B.VVV.dom \<mu>\<nu>\<tau>)))) =
B.dom (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>))"
using assms 1 VVV.arr_char VV.arr_char B.VxVxV.dom_char
B.VVV.dom_simp B.VVV.cod_simp
apply simp
by (metis (no_types, lifting) B.VV.arr_char B.VVV.arrE B.\<alpha>.preserves_reflects_arr
B.assoc_is_natural_1 B.seqE arr_dom dom_char src_dom trg_dom)
show "Arr (\<a> (fst (B.VVV.dom \<mu>\<nu>\<tau>)) (fst (snd (B.VVV.dom \<mu>\<nu>\<tau>)))
(snd (snd (B.VVV.dom \<mu>\<nu>\<tau>))))"
proof -
have "VVV.arr (B.VVV.dom \<mu>\<nu>\<tau>)"
using 1 B.VVV.dom_char B.VVV.arr_char B.VV.arr_char VVV.arr_char VV.arr_char
apply simp
by (metis (no_types, lifting) VVV.arrE arr_dom assms dom_simp src_dom trg_dom)
moreover have "Arr (\<a>\<^sub>B (B.dom (fst \<mu>\<nu>\<tau>)) (B.dom (fst (snd \<mu>\<nu>\<tau>)))
(B.dom (snd (snd \<mu>\<nu>\<tau>))))"
proof -
have "B.VVV.ide (B.VVV.dom \<mu>\<nu>\<tau>)"
using 1 B.VVV.ide_dom by blast
thus ?thesis
using assms B.\<alpha>_def B.VVV.arr_char B.VV.arr_char B.VVV.ide_char B.VV.ide_char
dom_closed assoc_closed
by (metis (no_types, lifting) "1" B.ide_dom B.src_dom B.trg_dom VV.arr_char
VVV.arrE arr_char)
qed
ultimately show ?thesis
using 1 B.VVV.ide_dom assoc_closed B.VVV.dom_char
apply simp
by (metis (no_types, lifting) B.VV.arr_char B.VVV.arrE B.VVV.inclusion
B.VxV.dom_char B.VxVxV.arrE B.VxVxV.dom_char prod.sel(1) prod.sel(2))
qed
qed
finally show ?thesis by blast
qed
qed
lemma lunit_closed':
assumes "Arr f"
shows "Arr (B.\<ll> f)"
proof -
have 1: "arr f \<and> arr (B.\<ll> (B.dom f))"
using assms arr_char lunit_closed dom_closed B.ide_dom inclusion by simp
moreover have "B.dom f = B.cod (B.\<ll> (B.dom f))"
using 1 arr_char B.\<ll>.preserves_cod inclusion by simp
moreover have "B.\<ll> f = f \<cdot> B.\<ll> (B.dom f)"
using assms 1 B.\<ll>.is_natural_1 inclusion comp_char arr_char by simp
ultimately show ?thesis
using arr_char comp_closed cod_char seqI dom_simp by auto
qed
lemma runit_closed':
assumes "Arr f"
shows "Arr (B.\<rr> f)"
proof -
have 1: "arr f \<and> arr (B.\<rr> (B.dom f))"
using assms arr_char runit_closed dom_closed B.ide_dom inclusion
by simp
moreover have "B.dom f = B.cod (B.\<rr> (B.dom f))"
using 1 arr_char B.\<ll>.preserves_cod inclusion by simp
moreover have "B.\<rr> f = f \<cdot> B.\<rr> (B.dom f)"
using assms 1 B.\<rr>.is_natural_1 inclusion comp_char arr_char by simp
ultimately show ?thesis
using arr_char comp_closed cod_char seqI dom_simp by auto
qed
interpretation natural_isomorphism VVV.comp \<open>(\<cdot>)\<close> HoHV HoVH \<alpha>\<^sub>S\<^sub>B
proof
fix \<mu>\<nu>\<tau>
show "\<not> VVV.arr \<mu>\<nu>\<tau> \<Longrightarrow> \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> = null"
by simp
assume \<mu>\<nu>\<tau>: "VVV.arr \<mu>\<nu>\<tau>"
have 1: "B.VVV.arr \<mu>\<nu>\<tau>"
using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char arr_char
src_def trg_def inclusion
by auto
show "dom (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
proof -
have "dom (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = B.HoHV (B.VVV.dom \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
B.\<alpha>_def assoc_closed' dom_simp
by simp
also have "... = HoHV (VVV.dom \<mu>\<nu>\<tau>)"
proof -
have "HoHV (VVV.dom \<mu>\<nu>\<tau>) = HoHV (VxVxV.dom \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV.dom_char VV.arr_char src_def trg_def VVV.arr_char by auto
also have "... = B.HoHV (B.VVV.dom \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV.dom_char VVV.arr_char VV.arr_char src_def trg_def
HoHV_def B.HoHV_def arr_char B.VVV.arr_char B.VVV.dom_char B.VV.arr_char
dom_closed hcomp_closed hcomp_def inclusion dom_simp
by auto
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
show "cod (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
proof -
have "cod (\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>) = B.HoVH (B.VVV.cod \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> 1 arr_char VVV.arr_char VV.arr_char B.VVV.arr_char B.VV.arr_char
B.\<alpha>_def assoc_closed' cod_simp
by simp
also have "... = HoVH (VVV.cod \<mu>\<nu>\<tau>)"
proof -
have "HoVH (VVV.cod \<mu>\<nu>\<tau>) = HoVH (VxVxV.cod \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV.cod_char VV.arr_char src_def trg_def VVV.arr_char by auto
also have "... = B.HoVH (B.VVV.cod \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> VVV.cod_char VV.arr_char src_def trg_def VVV.arr_char
HoVH_def B.HoVH_def arr_char B.VVV.arr_char B.VVV.cod_char B.VV.arr_char
cod_closed hcomp_closed hcomp_def inclusion cod_simp
by simp
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
have 3: "Arr (fst \<mu>\<nu>\<tau>) \<and> Arr (fst (snd \<mu>\<nu>\<tau>)) \<and> Arr (snd (snd \<mu>\<nu>\<tau>)) \<and>
src\<^sub>B (fst \<mu>\<nu>\<tau>) = trg\<^sub>B (fst (snd \<mu>\<nu>\<tau>)) \<and>
src\<^sub>B (fst (snd \<mu>\<nu>\<tau>)) = trg\<^sub>B (snd (snd \<mu>\<nu>\<tau>))"
using \<mu>\<nu>\<tau> VVV.arr_char VV.arr_char src_def trg_def arr_char by auto
show "HoVH \<mu>\<nu>\<tau> \<cdot> \<alpha>\<^sub>S\<^sub>B (VVV.dom \<mu>\<nu>\<tau>) = \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>"
proof -
have "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> = (fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)) \<cdot>\<^sub>B
\<a>\<^sub>B (B.dom (fst \<mu>\<nu>\<tau>)) (B.dom (fst (snd \<mu>\<nu>\<tau>))) (B.dom (snd (snd \<mu>\<nu>\<tau>)))"
using 3 inclusion B.assoc_is_natural_1 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
by (simp add: \<mu>\<nu>\<tau>)
also have "... = (fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>) \<star> snd (snd \<mu>\<nu>\<tau>)) \<cdot>
\<a>\<^sub>B (dom (fst \<mu>\<nu>\<tau>)) (dom (fst (snd \<mu>\<nu>\<tau>))) (dom (snd (snd \<mu>\<nu>\<tau>)))"
using 1 3 \<mu>\<nu>\<tau> hcomp_closed assoc_closed dom_closed hcomp_def comp_def inclusion
comp_char dom_char VVV.arr_char VV.arr_char
apply simp
using B.hcomp_simps(2-3) arr_char by presburger
also have "... = HoVH \<mu>\<nu>\<tau> \<cdot> \<alpha>\<^sub>S\<^sub>B (VVV.dom \<mu>\<nu>\<tau>)"
using \<mu>\<nu>\<tau> B.\<alpha>_def HoVH_def VVV.dom_char VV.dom_char VxVxV.dom_char
apply simp
by (metis (no_types, lifting) VV.arr_char VVV.arrE VVV.arr_dom VxV.dom_char
dom_simp)
finally show ?thesis by argo
qed
show "\<alpha>\<^sub>S\<^sub>B (VVV.cod \<mu>\<nu>\<tau>) \<cdot> HoHV \<mu>\<nu>\<tau> = \<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau>"
proof -
have "\<alpha>\<^sub>S\<^sub>B \<mu>\<nu>\<tau> =
\<a>\<^sub>B (B.cod (fst \<mu>\<nu>\<tau>)) (B.cod (fst (snd \<mu>\<nu>\<tau>))) (B.cod (snd (snd \<mu>\<nu>\<tau>))) \<cdot>\<^sub>B
(fst \<mu>\<nu>\<tau> \<star>\<^sub>B fst (snd \<mu>\<nu>\<tau>)) \<star>\<^sub>B snd (snd \<mu>\<nu>\<tau>)"
using 3 inclusion B.assoc_is_natural_2 [of "fst \<mu>\<nu>\<tau>" "fst (snd \<mu>\<nu>\<tau>)" "snd (snd \<mu>\<nu>\<tau>)"]
by (simp add: \<mu>\<nu>\<tau>)
also have "... = \<a>\<^sub>B (cod (fst \<mu>\<nu>\<tau>)) (cod (fst (snd \<mu>\<nu>\<tau>))) (cod (snd (snd \<mu>\<nu>\<tau>))) \<cdot>
((fst \<mu>\<nu>\<tau> \<star> fst (snd \<mu>\<nu>\<tau>)) \<star> snd (snd \<mu>\<nu>\<tau>)) "
using 1 3 \<mu>\<nu>\<tau> hcomp_closed assoc_closed cod_closed hcomp_def comp_def inclusion
comp_char cod_char VVV.arr_char VV.arr_char
by auto
also have "... = \<alpha>\<^sub>S\<^sub>B (VVV.cod \<mu>\<nu>\<tau>) \<cdot> HoHV \<mu>\<nu>\<tau>"
using \<mu>\<nu>\<tau> B.\<alpha>_def HoHV_def VVV.cod_char VV.cod_char VxVxV.cod_char
VVV.arr_char VV.arr_char arr_cod src_cod trg_cod
by simp
finally show ?thesis by argo
qed
next
fix fgh
assume fgh: "VVV.ide fgh"
show "iso (\<alpha>\<^sub>S\<^sub>B fgh)"
proof -
have 1: "B.arr (fst (snd fgh)) \<and> B.arr (snd (snd fgh)) \<and>
src\<^sub>B (fst (snd fgh)) = trg\<^sub>B (snd (snd fgh)) \<and>
src\<^sub>B (fst fgh) = trg\<^sub>B (fst (snd fgh))"
using fgh VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def
arr_char inclusion
by auto
have 2: "B.ide (fst fgh) \<and> B.ide (fst (snd fgh)) \<and> B.ide (snd (snd fgh))"
using fgh VVV.ide_char ide_char by blast
have "\<alpha>\<^sub>S\<^sub>B fgh = \<a>\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))"
using fgh B.\<alpha>_def by simp
moreover have "B.VVV.ide fgh"
using fgh 1 2 VVV.ide_char B.VVV.ide_char VVV.arr_char B.VVV.arr_char
src_def trg_def inclusion arr_char B.VV.arr_char
by simp
moreover have "Arr (\<a>\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh)))"
using fgh 1 VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def
arr_char assoc_closed' B.\<alpha>_def
by simp
moreover have "Arr (B.inv (\<a>\<^sub>B (fst fgh) (fst (snd fgh)) (snd (snd fgh))))"
using fgh 1 VVV.ide_char VVV.arr_char VV.arr_char src_def trg_def
arr_char assoc'_closed
by (simp add: VVV.arr_char "2" B.VVV.ide_char calculation(2))
ultimately show ?thesis
using fgh iso_char B.\<alpha>.components_are_iso by auto
qed
qed
interpretation L: endofunctor \<open>(\<cdot>)\<close> L
using endofunctor_L by auto
interpretation R: endofunctor \<open>(\<cdot>)\<close> R
using endofunctor_R by auto
interpretation L: faithful_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> L
proof
fix f f'
assume par: "par f f'"
assume eq: "L f = L f'"
have "B.par f f'"
using par inclusion arr_char dom_simp cod_simp by fastforce
moreover have "B.L f = B.L f'"
proof -
have "\<forall>a. Arr a \<longrightarrow> B.arr a"
by (simp add: inclusion)
moreover have 1: "\<forall>a. arr a \<longrightarrow> (if arr a then hseq (trg a) a else arr null)"
using L.preserves_arr by presburger
moreover have "Arr f \<and> Arr (trg f) \<and> trg\<^sub>B f = src\<^sub>B (trg f)"
by (simp add: \<open>B.par f f'\<close> arrE par trg_closed trg_def)
ultimately show ?thesis
by (metis \<open>B.par f f'\<close> eq hcomp_def hseq_char' par trg_def)
qed
ultimately show "f = f'"
using B.L.is_faithful by blast
qed
interpretation L: full_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> L
proof
fix f f' \<nu>
assume f: "ide f" and f': "ide f'" and \<nu>: "\<guillemotleft>\<nu> : L f \<Rightarrow> L f'\<guillemotright>"
have 1: "L f = trg\<^sub>B f \<star>\<^sub>B f \<and> L f' = trg\<^sub>B f' \<star>\<^sub>B f'"
using f f' hcomp_def trg_def arr_char ide_char trg_closed by simp
have 2: "\<guillemotleft>\<nu> : trg\<^sub>B f \<star>\<^sub>B f \<Rightarrow>\<^sub>B trg\<^sub>B f' \<star>\<^sub>B f'\<guillemotright>"
using 1 f f' \<nu> hcomp_def trg_def src_def inclusion
dom_char cod_char hseq_char' arr_char ide_char trg_closed null_char
by (simp add: arr_char in_hom_char)
show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> L \<mu> = \<nu>"
proof -
let ?\<mu> = "B.\<ll> f' \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B B.inv (B.\<ll> f)"
have \<mu>: "\<guillemotleft>?\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
proof -
have "\<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
using f f' \<nu> 2 B.\<ll>_ide_simp lunit'_closed lunit_closed' ide_char by auto
thus ?thesis
using f f' \<nu> in_hom_char arr_char comp_closed ide_char
lunit'_closed lunit_closed
by (metis (no_types, lifting) B.arrI B.seqE in_homE)
qed
have \<mu>_eq: "?\<mu> = B.\<ll> f' \<cdot> \<nu> \<cdot> B.inv (B.\<ll> f)"
proof -
have "?\<mu> = B.\<ll> f' \<cdot> \<nu> \<cdot>\<^sub>B B.inv (B.\<ll> f)"
using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
lunit'_closed lunit_closed
by (metis (no_types, lifting) B.seqE in_homE)
also have "... = B.\<ll> f' \<cdot> \<nu> \<cdot> B.inv (B.\<ll> f)"
using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
lunit'_closed lunit_closed
by (metis (no_types, lifting) B.seqE in_homE)
finally show ?thesis by simp
qed
have "L ?\<mu> = \<nu>"
proof -
have "L ?\<mu> = trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>"
using \<mu> \<mu>_eq hcomp_def trg_def inclusion arr_char trg_closed by auto
also have "... = (trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) \<cdot>\<^sub>B (B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f)"
proof -
have "B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f = trg\<^sub>B f \<star>\<^sub>B f"
using f ide_char B.comp_inv_arr B.inv_is_inverse by auto
moreover have "B.dom (trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) = trg\<^sub>B f \<star>\<^sub>B f"
using f \<mu> \<mu>_eq ide_char arr_char B.trg_dom [of ?\<mu>] by fastforce
ultimately show ?thesis
using \<mu> \<mu>_eq B.comp_arr_dom in_hom_char by auto
qed
also have "... = ((trg\<^sub>B ?\<mu> \<star>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.inv (B.\<ll> f)) \<cdot>\<^sub>B B.\<ll> f"
using B.comp_assoc by simp
also have "... = (B.inv (B.\<ll> f') \<cdot>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.\<ll> f"
using \<mu> \<mu>_eq B.\<ll>'.naturality [of ?\<mu>] by auto
also have "... = (B.inv (B.\<ll> f') \<cdot>\<^sub>B B.\<ll> f') \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B (B.inv (B.\<ll> f) \<cdot>\<^sub>B B.\<ll> f)"
using \<mu> \<mu>_eq arr_char arrI comp_simp B.comp_assoc by metis
also have "... = \<nu>"
using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char
B.\<ll>.components_are_iso B.\<ll>_ide_simp B.comp_inv_arr'
by auto
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> by auto
qed
qed
interpretation R: faithful_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> R
proof
fix f f'
assume par: "par f f'"
assume eq: "R f = R f'"
have "B.par f f'"
using par inclusion arr_char dom_simp cod_simp by fastforce
moreover have "B.R f = B.R f'"
proof -
have "\<forall>a. Arr a \<longrightarrow> B.arr a"
by (simp add: inclusion)
moreover have 1: "\<forall>a. arr a \<longrightarrow> (if arr a then hseq a (src a) else arr null)"
using R.preserves_arr by presburger
moreover have "arr f \<and> arr (src f) \<and> trg\<^sub>B (src f) = src\<^sub>B f"
by (meson 1 hcomp_def hseq_char' par)
ultimately show ?thesis
by (metis \<open>B.par f f'\<close> eq hcomp_def hseq_char' src_def)
qed
ultimately show "f = f'"
using B.R.is_faithful by blast
qed
interpretation R: full_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> R
proof
fix f f' \<nu>
assume f: "ide f" and f': "ide f'" and \<nu>: "\<guillemotleft>\<nu> : R f \<Rightarrow> R f'\<guillemotright>"
have 1: "R f = f \<star>\<^sub>B src\<^sub>B f \<and> R f' = f' \<star>\<^sub>B src\<^sub>B f'"
using f f' hcomp_def src_def arr_char ide_char src_closed by simp
have 2: "\<guillemotleft>\<nu> : f \<star>\<^sub>B src\<^sub>B f \<Rightarrow>\<^sub>B f' \<star>\<^sub>B src\<^sub>B f'\<guillemotright>"
using 1 f f' \<nu> hcomp_def trg_def src_def inclusion
dom_char cod_char hseq_char' arr_char ide_char trg_closed null_char
by (simp add: arr_char in_hom_char)
show "\<exists>\<mu>. \<guillemotleft>\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> R \<mu> = \<nu>"
proof -
let ?\<mu> = "B.\<rr> f' \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B B.inv (B.\<rr> f)"
have \<mu>: "\<guillemotleft>?\<mu> : f \<Rightarrow> f'\<guillemotright> \<and> \<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
proof -
have "\<guillemotleft>?\<mu> : f \<Rightarrow>\<^sub>B f'\<guillemotright>"
using f f' \<nu> 2 B.\<rr>_ide_simp runit'_closed runit_closed' ide_char by auto
thus ?thesis
using f f' \<nu> in_hom_char [of ?\<mu> f f'] arr_char comp_closed ide_char
runit'_closed runit_closed
apply auto
by fastforce
qed
have \<mu>_eq: "?\<mu> = B.\<rr> f' \<cdot> \<nu> \<cdot> B.inv (B.\<rr> f)"
proof -
have "?\<mu> = B.\<rr> f' \<cdot> \<nu> \<cdot>\<^sub>B B.inv (B.\<rr> f)"
using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
runit'_closed runit_closed
by (metis (no_types, lifting) B.seqE in_homE)
also have "... = B.\<rr> f' \<cdot> \<nu> \<cdot> B.inv (B.\<rr> f)"
using f f' \<nu> \<mu> arr_char inclusion comp_char comp_closed ide_char
runit'_closed runit_closed
by (metis (no_types, lifting) B.arrI B.comp_in_homE in_hom_char)
finally show ?thesis by simp
qed
have "R ?\<mu> = \<nu>"
proof -
have "R ?\<mu> = ?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>"
using \<mu> \<mu>_eq hcomp_def src_def inclusion arr_char src_closed by auto
also have "... = (?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) \<cdot>\<^sub>B (B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f)"
proof -
have "B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f = f \<star>\<^sub>B src\<^sub>B f"
using f ide_char B.comp_inv_arr B.inv_is_inverse by auto
moreover have "B.dom (?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) = f \<star>\<^sub>B src\<^sub>B f"
using f \<mu> \<mu>_eq ide_char arr_char B.src_dom [of ?\<mu>] by fastforce
ultimately show ?thesis
using \<mu> \<mu>_eq B.comp_arr_dom in_hom_char by auto
qed
also have "... = ((?\<mu> \<star>\<^sub>B src\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.inv (B.\<rr> f)) \<cdot>\<^sub>B B.\<rr> f"
using B.comp_assoc by simp
also have "... = (B.inv (B.\<rr> f') \<cdot>\<^sub>B ?\<mu>) \<cdot>\<^sub>B B.\<rr> f"
using \<mu> \<mu>_eq B.\<rr>'.naturality [of ?\<mu>] by auto
also have "... = (B.inv (B.\<rr> f') \<cdot>\<^sub>B B.\<rr> f') \<cdot>\<^sub>B \<nu> \<cdot>\<^sub>B (B.inv (B.\<rr> f) \<cdot>\<^sub>B B.\<rr> f)"
using \<mu> \<mu>_eq arr_char arrI comp_simp B.comp_assoc by metis
also have "... = \<nu>"
using f f' \<nu> 2 B.comp_arr_dom B.comp_cod_arr ide_char
B.\<ll>.components_are_iso B.\<ll>_ide_simp B.comp_inv_arr'
by auto
finally show ?thesis by blast
qed
thus ?thesis
using \<mu> by blast
qed
qed
interpretation bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
proof
show "\<And>a. obj a \<Longrightarrow> \<guillemotleft>\<i>[a] : a \<star> a \<Rightarrow> a\<guillemotright>"
proof (intro in_homI)
fix a
assume a: "obj a"
have 1: "Arr (\<i> a)"
using a obj_def src_def trg_def in_hom_char B.unit_in_hom
arr_char hcomp_def B.obj_def ide_char objE hcomp_closed
by (metis (no_types, lifting) B.\<ll>_ide_simp B.unitor_coincidence(1) inclusion lunit_closed)
show 2: "arr \<i>[a]"
using 1 arr_char by simp
show "dom \<i>[a] = a \<star> a"
using a 2 dom_char
by (metis (full_types) B.objI_trg B.unit_simps(4) R.preserves_reflects_arr
hcomp_def hseq_char' inclusion objE obj_simps(1)
subcategory.arrE subcategory_axioms trg_def)
show "cod \<i>[a] = a"
using a 2 cod_char
by (metis B.obj_def' B.unit_simps(5) inclusion objE obj_simps(1)
subcategory.arrE subcategory_axioms trg_def)
qed
show "\<And>a. obj a \<Longrightarrow> iso (\<i> a)"
proof -
fix a
assume a: "obj a"
have 1: "trg\<^sub>B a = src\<^sub>B a"
using a obj_def src_def trg_def B.obj_def arr_char
by (metis horizontal_homs.objE horizontal_homs_axioms)
have 2: "Arr (\<i> a)"
using a 1 obj_def src_def trg_def in_hom_char B.unit_in_hom
arr_char hcomp_def B.obj_def ide_char objE hcomp_closed
by (metis (no_types, lifting) B.\<ll>_ide_simp B.unitor_coincidence(1) inclusion lunit_closed)
have "iso (B.\<ll> a)"
using a 2 obj_def B.iso_unit iso_char arr_char lunit_closed lunit'_closed B.iso_lunit
apply simp
by (metis (no_types, lifting) B.\<ll>.components_are_iso B.ide_src inclusion src_def)
thus "iso (\<i> a)"
using a 2 obj_def B.iso_unit iso_char arr_char B.unitor_coincidence
apply simp
by (metis (no_types, lifting) B.\<ll>_ide_simp B.ide_src B.obj_src inclusion src_def)
qed
show "\<And>f g h k. \<lbrakk> ide f; ide g; ide h; ide k;
src f = trg g; src g = trg h; src h = trg k \<rbrakk> \<Longrightarrow>
(f \<star> \<a> g h k) \<cdot> \<a> f (g \<star> h) k \<cdot> (\<a> f g h \<star> k) =
\<a> f g (h \<star> k) \<cdot> \<a> (f \<star> g) h k"
using B.pentagon VVV.arr_char VV.arr_char hcomp_def assoc_closed arr_char comp_char
hcomp_closed comp_closed ide_char inclusion src_def trg_def
by simp
qed
proposition is_bicategory:
shows "bicategory (\<cdot>) (\<star>) \<a> \<i> src trg"
..
lemma obj_char:
shows "obj a \<longleftrightarrow> arr a \<and> B.obj a"
proof
assume a: "obj a"
show "arr a \<and> B.obj a"
using a obj_def B.obj_def src_def arr_char inclusion by metis
next
assume a: "arr a \<and> B.obj a"
have "src a = a"
using a src_def by auto
thus "obj a"
using a obj_def by simp
qed
lemma hcomp_char:
shows "hcomp = (\<lambda>f g. if arr f \<and> arr g \<and> src f = trg g then f \<star>\<^sub>B g else null)"
using hcomp_def src_def trg_def by metis
lemma assoc_simp:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "\<a> f g h = \<a>\<^sub>B f g h"
using assms VVV.arr_char VV.arr_char by auto
lemma assoc'_simp:
assumes "ide f" and "ide g" and "ide h" and "src f = trg g" and "src g = trg h"
shows "\<a>' f g h = B.\<a>' f g h"
proof -
have "\<a>' f g h = B.inv (\<a>\<^sub>B f g h)"
using assms inv_char by fastforce
also have "... = B.\<a>' f g h"
using assms ide_char src_def trg_def
B.VVV.ide_char B.VVV.arr_char B.VV.arr_char
by force
finally show ?thesis by blast
qed
lemma lunit_simp:
assumes "ide f"
shows "lunit f = B.lunit f"
proof -
have "B.lunit f = lunit f"
proof (intro lunit_eqI)
show "ide f" by fact
show 1: "\<guillemotleft>B.lunit f : trg f \<star> f \<Rightarrow> f\<guillemotright>"
proof
show 2: "arr (B.lunit f)"
using assms arr_char lunit_closed
by (simp add: arr_char B.\<ll>_ide_simp ide_char)
show "dom (B.lunit f) = trg f \<star> f"
using assms 2 dom_char hcomp_char ide_char src_trg trg.preserves_arr trg_def
by auto
show "cod (B.lunit f) = f"
using assms 2 in_hom_char
by (simp add: cod_simp ide_char)
qed
show "trg f \<star> B.lunit f = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
proof -
have "trg f \<star> B.lunit f = trg\<^sub>B f \<star>\<^sub>B B.lunit f"
using assms 1 arr_char hcomp_char
by (metis (no_types, lifting) ideD(1) src_trg trg.preserves_reflects_arr
trg_def vconn_implies_hpar(2,4))
also have "... = (\<i>[trg f] \<star>\<^sub>B f) \<cdot>\<^sub>B B.\<a>' (trg f) (trg f) f"
using assms ide_char B.lunit_char(2) trg_def by simp
also have "... = (\<i>[trg f] \<star>\<^sub>B f) \<cdot>\<^sub>B \<a>' (trg f) (trg f) f"
using assms assoc'_simp [of "trg f" "trg f" f] by simp
also have "... = (\<i>[trg f] \<star> f) \<cdot>\<^sub>B \<a>' (trg f) (trg f) f"
using assms hcomp_char by simp
also have "... = (\<i>[trg f] \<star> f) \<cdot> \<a>' (trg f) (trg f) f"
using assms seq_char [of "\<i>[trg f] \<star> f" "\<a>' (trg f) (trg f) f"]
comp_char [of "\<i>[trg f] \<star> f" "\<a>' (trg f) (trg f) f"]
by simp
finally show ?thesis by blast
qed
qed
thus ?thesis by simp
qed
lemma lunit'_simp:
assumes "ide f"
shows "lunit' f = B.lunit' f"
using assms inv_char [of "lunit f"] lunit_simp by fastforce
lemma runit_simp:
assumes "ide f"
shows "runit f = B.runit f"
proof -
have "B.runit f = runit f"
proof (intro runit_eqI)
show "ide f" by fact
show 1: "\<guillemotleft>B.runit f : f \<star> src f \<Rightarrow> f\<guillemotright>"
proof
show 2: "arr (B.runit f)"
using assms arr_char runit_closed
by (simp add: arr_char B.\<rr>_ide_simp ide_char)
show "dom (B.runit f) = f \<star> src f"
using assms 2 dom_char hcomp_char
by (metis (no_types, lifting) B.runit_simps(4) ide_char src.preserves_reflects_arr
src_def trg_src)
show "cod (B.runit f) = f"
using assms 2 in_hom_char
by (simp add: cod_simp ide_char)
qed
show "B.runit f \<star> src f = (f \<star> \<i>[src f]) \<cdot> \<a> f (src f) (src f)"
proof -
have "B.runit f \<star> src f = B.runit f \<star>\<^sub>B src\<^sub>B f"
using assms 1 arr_char hcomp_char
by (metis (no_types, lifting) ideD(1) src.preserves_reflects_arr src_def
trg_src vconn_implies_hpar(1,3))
also have "... = (f \<star>\<^sub>B \<i>[src f]) \<cdot>\<^sub>B \<a>\<^sub>B f (src f) (src f)"
using assms ide_char B.runit_char(2) src_def by simp
also have "... = (f \<star>\<^sub>B \<i>[src f]) \<cdot>\<^sub>B \<a> f (src f) (src f)"
using assms assoc_simp by simp
also have "... = (f \<star> \<i>[src f]) \<cdot>\<^sub>B \<a> f (src f) (src f)"
using assms 1 hcomp_char by simp
also have "... = (f \<star> \<i>[src f]) \<cdot> \<a> f (src f) (src f)"
proof -
have "B.seq (f \<star> \<i>[src f]) (\<a> f (src f) (src f))"
using assms seq_char [of "f \<star> \<i>[src f]" "\<a> f (src f) (src f)"] by simp
thus ?thesis
using assms comp_char [of "f \<star> \<i>[src f]" "\<a> f (src f) (src f)"] by simp
qed
finally show ?thesis by blast
qed
qed
thus ?thesis by simp
qed
lemma runit'_simp:
assumes "ide f"
shows "runit' f = B.runit' f"
using assms inv_char [of "runit f"] runit_simp by fastforce
lemma comp_eqI [intro]:
assumes "seq f g" and "f = f'" and "g = g'"
shows "f \<cdot> g = f' \<cdot>\<^sub>B g'"
using assms comp_char ext ext not_arr_null by auto
lemma comp_eqI' [intro]:
assumes "seq f g" and "f = f'" and "g = g'"
shows "f \<cdot>\<^sub>B g = f' \<cdot> g'"
using assms comp_char ext ext not_arr_null by auto
lemma hcomp_eqI [intro]:
assumes "hseq f g" and "f = f'" and "g = g'"
shows "f \<star> g = f' \<star>\<^sub>B g'"
using assms hcomp_char not_arr_null by auto
lemma hcomp_eqI' [intro]:
assumes "hseq f g" and "f = f'" and "g = g'"
shows "f \<star>\<^sub>B g = f' \<star> g'"
using assms hcomp_char not_arr_null by auto
lemma arr_compI:
assumes "seq f g"
shows "arr (f \<cdot>\<^sub>B g)"
using assms seq_char dom_char cod_char
by (metis (no_types, lifting) comp_simp)
lemma arr_hcompI:
assumes "hseq f g"
shows "arr (f \<star>\<^sub>B g)"
using assms hseq_char src_def trg_def hcomp_eqI' by auto
end
sublocale subbicategory \<subseteq> bicategory \<open>(\<cdot>)\<close> \<open>(\<star>)\<close> \<a> \<i> src trg
using is_bicategory by auto
subsection "The Sub-bicategory of Endo-arrows of an Object"
text \<open>
We now consider the sub-bicategory consisting of all arrows having the same
object \<open>a\<close> both as their source and their target and we show that the resulting structure
is a monoidal category. We actually prove a slightly more general result,
in which the unit of the monoidal category is taken to be an arbitrary isomorphism
\<open>\<guillemotleft>\<omega> : w \<star>\<^sub>B w \<Rightarrow> w\<guillemotright>\<close> with \<open>w\<close> isomorphic to \<open>a\<close>, rather than the particular choice
\<open>\<guillemotleft>\<i>[a] : a \<star>\<^sub>B a \<Rightarrow> a\<guillemotright>\<close> made by the ambient bicategory.
\<close>
locale subbicategory_at_object =
B: bicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B +
subbicategory V H \<a>\<^sub>B \<i> src\<^sub>B trg\<^sub>B \<open>\<lambda>\<mu>. B.arr \<mu> \<and> src\<^sub>B \<mu> = a \<and> trg\<^sub>B \<mu> = a\<close>
for V :: "'a comp" (infixr "\<cdot>\<^sub>B" 55)
and H :: "'a comp" (infixr "\<star>\<^sub>B" 55)
and \<a>\<^sub>B :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<a>\<^sub>B[_, _, _]")
and \<i> :: "'a \<Rightarrow> 'a" ("\<i>[_]")
and src\<^sub>B :: "'a \<Rightarrow> 'a"
and trg\<^sub>B :: "'a \<Rightarrow> 'a"
and a :: "'a"
and w :: "'a"
and \<omega> :: "'a" +
assumes obj_a: "B.obj a"
and isomorphic_a_w: "B.isomorphic a w"
and \<omega>_in_vhom: "\<guillemotleft>\<omega> : w \<star>\<^sub>B w \<Rightarrow> w\<guillemotright>"
and \<omega>_is_iso: "B.iso \<omega>"
begin
notation hcomp (infixr "\<star>" 53)
lemma arr_simps:
assumes "arr \<mu>"
shows "src \<mu> = a" and "trg \<mu> = a"
apply (metis (no_types, lifting) arrE assms src_def)
by (metis (no_types, lifting) arrE assms trg_def)
lemma \<omega>_simps [simp]:
shows "arr \<omega>"
and "src \<omega> = a" and "trg \<omega> = a"
and "dom \<omega> = w \<star>\<^sub>B w" and "cod \<omega> = w"
using isomorphic_a_w \<omega>_in_vhom in_hom_char arr_simps by auto
lemma ide_w:
shows "B.ide w"
using isomorphic_a_w B.isomorphic_def by auto
lemma w_simps [simp]:
shows "ide w" and "B.ide w"
and "src w = a" and "trg w = a" and "src\<^sub>B w = a" and "trg\<^sub>B w = a"
and "dom w = w" and "cod w = w"
proof -
show w: "ide w"
using \<omega>_in_vhom ide_cod by blast
show "B.ide w"
using w ide_char by simp
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : a \<Rightarrow>\<^sub>B w\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_def by auto
show "src\<^sub>B w = a"
using obj_a w \<phi> B.src_cod by force
show "trg\<^sub>B w = a"
using obj_a w \<phi> B.src_cod by force
show "src w = a"
- using `src\<^sub>B w = a` w ide_w src_def by simp
+ using \<open>src\<^sub>B w = a\<close> w ide_w src_def by simp
show "trg w = a"
- using `src\<^sub>B w = a` w ide_w trg_def
+ using \<open>src\<^sub>B w = a\<close> w ide_w trg_def
by (simp add: \<open>trg\<^sub>B w = a\<close>)
show "dom w = w"
using w by simp
show "cod w = w"
using w by simp
qed
lemma VxV_arr_eq_VV_arr:
shows "VxV.arr f \<longleftrightarrow> VV.arr f"
using inclusion VxV.arr_char VV.arr_char arr_char src_def trg_def
by auto
lemma VxV_comp_eq_VV_comp:
shows "VxV.comp = VV.comp"
proof -
have "\<And>f g. VxV.comp f g = VV.comp f g"
proof -
fix f g
show "VxV.comp f g = VV.comp f g"
unfolding VV.comp_def
using VxV.comp_char arr_simps(1) arr_simps(2)
apply (cases "seq (fst f) (fst g)", cases "seq (snd f) (snd g)")
by (elim seqE) auto
qed
thus ?thesis by blast
qed
lemma VxVxV_arr_eq_VVV_arr:
shows "VxVxV.arr f \<longleftrightarrow> VVV.arr f"
using VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char
by auto
lemma VxVxV_comp_eq_VVV_comp:
shows "VxVxV.comp = VVV.comp"
proof -
have "\<And>f g. VxVxV.comp f g = VVV.comp f g"
proof -
fix f g
show "VxVxV.comp f g = VVV.comp f g"
proof (cases "VxVxV.seq f g")
assume 1: "\<not> VxVxV.seq f g"
have "VxVxV.comp f g = VxVxV.null"
using 1 VxVxV.ext by blast
also have "... = (null, null, null)"
using VxVxV.null_char VxV.null_char by simp
also have "... = VVV.null"
using VVV.null_char VV.null_char by simp
also have "... = VVV.comp f g"
proof -
have "\<not> VVV.seq f g"
using 1 VVV.seq_char by blast
thus ?thesis
by (metis (no_types, lifting) VVV.ext)
qed
finally show ?thesis by simp
next
assume 1: "VxVxV.seq f g"
have 2: "B.arr (fst f) \<and> B.arr (fst (snd f)) \<and> B.arr (snd (snd f)) \<and>
src\<^sub>B (fst f) = a \<and> src\<^sub>B (fst (snd f)) = a \<and> src\<^sub>B (snd (snd f)) = a \<and>
trg\<^sub>B (fst f) = a \<and> trg\<^sub>B (fst (snd f)) = a \<and> trg\<^sub>B (snd (snd f)) = a"
using 1 VxVxV.seq_char VxV.seq_char arr_char by blast
have 3: "B.arr (fst g) \<and> B.arr (fst (snd g)) \<and> B.arr (snd (snd g)) \<and>
src\<^sub>B (fst g) = a \<and> src\<^sub>B (fst (snd g)) = a \<and> src\<^sub>B (snd (snd g)) = a \<and>
trg\<^sub>B (fst g) = a \<and> trg\<^sub>B (fst (snd g)) = a \<and> trg\<^sub>B (snd (snd g)) = a"
using 1 VxVxV.seq_char VxV.seq_char arr_char by blast
have 4: "B.seq (fst f) (fst g) \<and> B.seq (fst (snd f)) (fst (snd g)) \<and>
B.seq (snd (snd f)) (snd (snd g))"
using 1 VxVxV.seq_char VxV.seq_char seq_char by blast
have 5: "VxVxV.comp f g =
(fst f \<cdot> fst g, fst (snd f) \<cdot> fst (snd g), snd (snd f) \<cdot> snd (snd g))"
using 1 2 3 4 VxVxV.seqE VxVxV.comp_char VxV.comp_char seq_char arr_char
by (metis (no_types, lifting))
also have "... = VVV.comp f g"
using 1 VVV.comp_char VVV.arr_char VV.arr_char
apply simp
using 2 3 5 arrI arr_simps(1) arr_simps(2) by presburger
finally show ?thesis by blast
qed
qed
thus ?thesis by blast
qed
interpretation H: "functor" VxV.comp \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close>
using H.functor_axioms hcomp_def VxV_comp_eq_VV_comp by simp
interpretation H: binary_endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> ..
lemma HoHV_eq_ToTC:
shows "HoHV = H.ToTC"
using HoHV_def H.ToTC_def VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char
by auto
lemma HoVH_eq_ToCT:
shows "HoVH = H.ToCT"
using HoVH_def H.ToCT_def VVV.arr_char VV.arr_char src_def trg_def inclusion arr_char
by auto
interpretation ToTC: "functor" VxVxV.comp \<open>(\<cdot>)\<close> H.ToTC
using HoHV_eq_ToTC VxVxV_comp_eq_VVV_comp HoHV.functor_axioms by simp
interpretation ToCT: "functor" VxVxV.comp \<open>(\<cdot>)\<close> H.ToCT
using HoVH_eq_ToCT VxVxV_comp_eq_VVV_comp HoVH.functor_axioms by simp
interpretation \<alpha>: natural_isomorphism VxVxV.comp \<open>(\<cdot>)\<close> H.ToTC H.ToCT \<alpha>
unfolding \<alpha>_def
using \<alpha>.natural_isomorphism_axioms HoHV_eq_ToTC HoVH_eq_ToCT \<alpha>_def
VxVxV_comp_eq_VVV_comp
by simp
interpretation L: endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close>
proof
fix f
show "\<not> arr f \<Longrightarrow> fst (w, f) \<star> snd (w, f) = null"
using arr_char hcomp_def by auto
assume f: "arr f"
show "hseq (fst (w, f)) (snd (w, f))"
using f hseq_char arr_char src_def trg_def \<omega>_in_vhom cod_char by simp
show "dom (fst (w, f) \<star> snd (w, f)) = fst (w, dom f) \<star> snd (w, dom f)"
using f arr_char hcomp_def dom_simp by simp
show "cod (fst (w, f) \<star> snd (w, f)) = fst (w, cod f) \<star> snd (w, cod f)"
using f arr_char hcomp_def cod_simp by simp
next
fix f g
assume fg: "seq g f"
show "fst (w, g \<cdot> f) \<star> snd (w, g \<cdot> f) = (fst (w, g) \<star> snd (w, g)) \<cdot> (fst (w, f) \<star> snd (w, f))"
by (simp add: fg whisker_left)
qed
interpretation L': equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close>
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_symmetric by force
have "\<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright>"
using \<phi> in_hom_char
by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg
\<omega>_in_vhom arr_char arr_cod cod_simp)
hence \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi> \<and> \<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright> \<and> iso \<phi>"
using \<phi> iso_char arr_char by auto
interpret \<l>: natural_isomorphism \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close>
\<open>\<lambda>f. fst (w, f) \<star> snd (w, f)\<close> map \<open>\<lambda>f. \<ll> f \<cdot> (\<phi> \<star> dom f)\<close>
proof
fix \<mu>
show "\<not> arr \<mu> \<Longrightarrow> \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>) = null"
using \<phi> arr_char dom_char ext
apply simp
using comp_null(2) hcomp_def by fastforce
assume \<mu>: "arr \<mu>"
have 0: "in_hhom (dom \<mu>) a a"
using \<mu> arr_char src_dom trg_dom src_def trg_def dom_simp by simp
have 1: "in_hhom \<phi> a a"
using \<phi> arr_char src_dom trg_dom src_def trg_def by auto
have 2: "hseq \<phi> (B.dom \<mu>)"
using \<mu> 0 1 dom_simp by (intro hseqI) auto
have 3: "seq (\<ll> \<mu>) (\<phi> \<star> dom \<mu>)"
proof (intro seqI')
show "\<guillemotleft>\<phi> \<star> dom \<mu> : w \<star> dom \<mu> \<Rightarrow> a \<star> dom \<mu>\<guillemotright>"
by (metis (no_types, lifting) 0 \<mu> \<phi> hcomp_in_vhom ide_dom ide_in_hom(2)
in_hhom_def w_simps(3))
show "\<guillemotleft>\<ll> \<mu> : a \<star> dom \<mu> \<Rightarrow> cod \<mu>\<guillemotright>"
using \<mu> 2 \<ll>.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] arr_simps(2) arr_cod by fastforce
qed
show "dom (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = fst (w, dom \<mu>) \<star> snd (w, dom \<mu>)"
proof -
have "dom (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = dom \<phi> \<star> dom \<mu>"
using \<mu> 3 hcomp_simps(3) dom_comp dom_dom
apply (elim seqE) by auto
also have "... = fst (w, dom \<mu>) \<star> snd (w, dom \<mu>)"
using \<omega>_in_vhom \<phi>
by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2))
finally show ?thesis by simp
qed
show "cod (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = map (cod \<mu>)"
proof -
have "seq (\<ll> \<mu>) (\<phi> \<star> dom \<mu>)"
using 3 by simp
hence "cod (\<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)) = cod (\<ll> \<mu>)"
using cod_comp by blast
also have "... = map (cod \<mu>)"
using \<mu> by blast
finally show ?thesis by blast
qed
show "map \<mu> \<cdot> \<ll> (dom \<mu>) \<cdot> (\<phi> \<star> dom (dom \<mu>)) = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "map \<mu> \<cdot> \<ll> (dom \<mu>) \<cdot> (\<phi> \<star> dom (dom \<mu>)) = (map \<mu> \<cdot> \<ll> (dom \<mu>)) \<cdot> (\<phi> \<star> dom \<mu>)"
using \<mu> comp_assoc by simp
also have "... = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
using \<mu> \<phi> \<ll>.is_natural_1 by auto
finally show ?thesis by blast
qed
show "(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> dom (cod \<mu>))) \<cdot> (fst (w, \<mu>) \<star> snd (w, \<mu>)) = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> dom (cod \<mu>))) \<cdot> (fst (w, \<mu>) \<star> snd (w, \<mu>)) =
(\<ll> (cod \<mu>) \<cdot> (\<phi> \<star> B.cod \<mu>)) \<cdot> (w \<star> \<mu>)"
using \<mu> \<phi> dom_char arr_char \<omega>_in_vhom cod_simp by simp
also have "... = \<ll> (cod \<mu>) \<cdot> (\<phi> \<cdot> w \<star> B.cod \<mu> \<cdot> \<mu>)"
proof -
have "seq \<phi> w"
using \<phi> \<omega>_in_vhom w_simps(1) by blast
moreover have 2: "seq (B.cod \<mu>) \<mu>"
using \<mu> seq_char cod_simp by (simp add: comp_cod_arr)
moreover have "src \<phi> = trg (B.cod \<mu>)"
using \<mu> \<phi> 2
by (metis (no_types, lifting) arr_simps(2) seqE vconn_implies_hpar(1) w_simps(3))
ultimately show ?thesis
using interchange comp_assoc by simp
qed
also have "... = \<ll> (cod \<mu>) \<cdot> (\<phi> \<star> \<mu>)"
using \<mu> \<phi> \<omega>_in_vhom comp_arr_dom comp_cod_arr cod_simp
apply (elim conjE in_homE) by auto
also have "... = (\<ll> (cod \<mu>) \<cdot> (cod \<phi> \<star> \<mu>)) \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have 1: "seq (cod \<phi>) \<phi>"
using \<phi> arr_cod_iff_arr dom_cod iso_is_arr seqI by presburger
moreover have 2: "seq \<mu> (dom \<mu>)"
using \<mu> by (simp add: comp_arr_dom)
moreover have "src (cod \<phi>) = trg \<mu>"
using \<mu> \<phi> arr_cod arr_simps(1-2) iso_is_arr by auto
ultimately show ?thesis
using 1 2 interchange [of "cod \<phi>" \<phi> \<mu> "dom \<mu>"] comp_arr_dom comp_cod_arr
comp_assoc by fastforce
qed
also have "... = \<ll> \<mu> \<cdot> (\<phi> \<star> dom \<mu>)"
proof -
have "L \<mu> = cod \<phi> \<star> \<mu>"
using \<mu> \<phi> arr_simps(2) in_homE by auto
hence "\<ll> (cod \<mu>) \<cdot> (cod \<phi> \<star> \<mu>) = \<ll> \<mu>"
using \<mu> \<ll>.is_natural_2 [of \<mu>] by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
next
show "\<And>f. ide f \<Longrightarrow> iso (\<ll> f \<cdot> (\<phi> \<star> dom f))"
proof -
fix f
assume f: "ide f"
have "iso (\<ll> f)"
using f iso_lunit by simp
moreover have "iso (\<phi> \<star> dom f)"
using \<phi> f src_def trg_def ide_char arr_char
apply (intro iso_hcomp, simp_all)
by (metis (no_types, lifting) in_homE)
moreover have "seq (\<ll> f) (\<phi> \<star> dom f)"
proof (intro seqI')
show " \<guillemotleft>\<ll> f : a \<star> f \<Rightarrow> f\<guillemotright>"
using f lunit_in_hom(2) \<ll>_ide_simp ide_char arr_char trg_def by simp
show "\<guillemotleft>\<phi> \<star> dom f : w \<star> f \<Rightarrow> a \<star> f\<guillemotright>"
using \<phi> f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom
in_hom_char
by (intro hcomp_in_vhom, auto)
qed
ultimately show "iso (\<ll> f \<cdot> (\<phi> \<star> dom f))"
using isos_compose by simp
qed
qed
show "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (w, f) \<star> snd (w, f))"
using \<l>.natural_isomorphism_axioms L.isomorphic_to_identity_is_equivalence by simp
qed
interpretation L: equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f)\<close>
proof -
have "(\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f)) = (\<lambda>f. fst (w, f) \<star> snd (w, f))"
using \<omega>_in_vhom by simp
thus "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (cod \<omega>, f) \<star> snd (cod \<omega>, f))"
using L'.equivalence_functor_axioms by simp
qed
interpretation R: endofunctor \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close>
proof
fix f
show "\<not> arr f \<Longrightarrow> fst (f, w) \<star> snd (f, w) = null"
using arr_char hcomp_def by auto
assume f: "arr f"
show "hseq (fst (f, w)) (snd (f, w))"
using f hseq_char arr_char src_def trg_def \<omega>_in_vhom cod_char isomorphic_a_w
B.isomorphic_def in_hom_char
by simp
show "dom (fst (f, w) \<star> snd (f, w)) = fst (dom f, w) \<star> snd (dom f, w)"
using f arr_char dom_char cod_char hcomp_def \<omega>_in_vhom by simp
show "cod (fst (f, w) \<star> snd (f, w)) = fst (cod f, w) \<star> snd (cod f, w)"
using f arr_char dom_char cod_char hcomp_def \<omega>_in_vhom by simp
next
fix f g
assume fg: "seq g f"
have 1: "a \<cdot>\<^sub>B a = a"
using obj_a by auto
show "fst (g \<cdot> f, w) \<star> snd (g \<cdot> f, w) = (fst (g, w) \<star> snd (g, w)) \<cdot> (fst (f, w) \<star> snd (f, w))"
by (simp add: fg whisker_right)
qed
interpretation R': equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close>
proof -
obtain \<phi> where \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi>"
using isomorphic_a_w B.isomorphic_symmetric by force
have "\<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright>"
using \<phi> in_hom_char
by (metis (no_types, lifting) B.in_homE B.src_cod B.src_src B.trg_cod B.trg_trg
\<omega>_in_vhom arr_char arr_cod cod_simp)
hence \<phi>: "\<guillemotleft>\<phi> : w \<Rightarrow>\<^sub>B a\<guillemotright> \<and> B.iso \<phi> \<and> \<guillemotleft>\<phi> : w \<Rightarrow> a\<guillemotright> \<and> iso \<phi>"
using \<phi> iso_char arr_char by auto
interpret \<r>: natural_isomorphism comp comp
\<open>\<lambda>f. fst (f, w) \<star> snd (f, w)\<close> map \<open>\<lambda>f. \<rr> f \<cdot> (dom f \<star> \<phi>)\<close>
proof
fix \<mu>
show "\<not> arr \<mu> \<Longrightarrow> \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>) = null"
using \<phi> arr_char dom_char ext
apply simp
using comp_null(2) hcomp_def by fastforce
assume \<mu>: "arr \<mu>"
have 0: "in_hhom (dom \<mu>) a a"
using \<mu> arr_char src_dom trg_dom src_def trg_def dom_simp by simp
have 1: "in_hhom \<phi> a a"
using \<phi> arr_char src_dom trg_dom src_def trg_def by auto
have 2: "hseq (B.dom \<mu>) \<phi>"
using \<mu> 0 1 dom_simp hseqI by auto
have 3: "seq (\<rr> \<mu>) (dom \<mu> \<star> \<phi>)"
proof (intro seqI')
show "\<guillemotleft>dom \<mu> \<star> \<phi> : dom \<mu> \<star> w \<Rightarrow> dom \<mu> \<star> a\<guillemotright>"
by (metis (no_types, lifting) "0" "1" \<mu> \<phi> hcomp_in_vhom hseqI hseq_char
ide_dom ide_in_hom(2) vconn_implies_hpar(2))
show "\<guillemotleft>\<rr> \<mu> : dom \<mu> \<star> a \<Rightarrow> cod \<mu>\<guillemotright>"
using \<mu> 2 \<rr>.preserves_hom [of \<mu> "dom \<mu>" "cod \<mu>"] arr_simps(2) arr_cod
dom_simp cod_simp
by fastforce
qed
show "dom (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = fst (dom \<mu>, w) \<star> snd (dom \<mu>, w)"
proof -
have "dom (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = dom \<mu> \<star> dom \<phi>"
using \<mu> 3 hcomp_simps(3) dom_comp dom_dom
apply (elim seqE) by auto
also have "... = fst (dom \<mu>, w) \<star> snd (dom \<mu>, w)"
using \<omega>_in_vhom \<phi>
by (metis (no_types, lifting) in_homE prod.sel(1) prod.sel(2))
finally show ?thesis by simp
qed
show "cod (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = map (cod \<mu>)"
proof -
have "seq (\<rr> \<mu>) (dom \<mu> \<star> \<phi>)"
using 3 by simp
hence "cod (\<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)) = cod (\<rr> \<mu>)"
using cod_comp by blast
also have "... = map (cod \<mu>)"
using \<mu> by blast
finally show ?thesis by blast
qed
show "map \<mu> \<cdot> \<rr> (dom \<mu>) \<cdot> (dom (dom \<mu>) \<star> \<phi>) = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "map \<mu> \<cdot> \<rr> (dom \<mu>) \<cdot> (dom (dom \<mu>) \<star> \<phi>) =
(map \<mu> \<cdot> \<rr> (dom \<mu>)) \<cdot> (dom (dom \<mu>) \<star> \<phi>)"
using comp_assoc by simp
also have "... = (map \<mu> \<cdot> \<rr> (dom \<mu>)) \<cdot> (dom \<mu> \<star> \<phi>)"
using \<mu> dom_dom by simp
also have "... = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
using \<mu> \<phi> \<rr>.is_natural_1 by auto
finally show ?thesis by blast
qed
show "(\<rr> (cod \<mu>) \<cdot> (dom (cod \<mu>) \<star> \<phi>)) \<cdot> (fst (\<mu>, w) \<star> snd (\<mu>, w)) = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "(\<rr> (cod \<mu>) \<cdot> (dom (cod \<mu>) \<star> \<phi>)) \<cdot> (fst (\<mu>, w) \<star> snd (\<mu>, w)) =
(\<rr> (cod \<mu>) \<cdot> (B.cod \<mu> \<star> \<phi>)) \<cdot> (\<mu> \<star> w)"
using \<mu> \<phi> dom_char arr_char \<omega>_in_vhom cod_simp by simp
also have "... = \<rr> (cod \<mu>) \<cdot> (B.cod \<mu> \<cdot> \<mu> \<star> \<phi> \<cdot> w)"
proof -
have 2: "seq \<phi> w"
using \<phi> \<omega>_in_vhom w_simps(1) by blast
moreover have "seq (B.cod \<mu>) \<mu>"
using \<mu> seq_char cod_simp by (simp add: comp_cod_arr)
moreover have "src (B.cod \<mu>) = trg \<phi>"
using \<mu> \<phi> 2
using arr_simps(1) calculation(2) seq_char vconn_implies_hpar(2) by force
ultimately show ?thesis
using interchange comp_assoc by simp
qed
also have "... = \<rr> (cod \<mu>) \<cdot> (\<mu> \<star> \<phi>)"
using \<mu> \<phi> \<omega>_in_vhom comp_arr_dom comp_cod_arr cod_simp
apply (elim conjE in_homE) by auto
also have "... = (\<rr> (cod \<mu>) \<cdot> (\<mu> \<star> cod \<phi>)) \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "(\<mu> \<star> cod \<phi>) \<cdot> (dom \<mu> \<star> \<phi>) = \<mu> \<star> \<phi>"
proof -
have "seq \<mu> (dom \<mu>)"
using \<mu> by (simp add: comp_arr_dom)
moreover have "seq (cod \<phi>) \<phi>"
using \<phi> iso_is_arr arr_cod dom_cod by auto
moreover have "src \<mu> = trg (cod \<phi>)"
using \<mu> \<phi> 2
by (metis (no_types, lifting) arr_simps(1) arr_simps(2) calculation(2) seqE)
ultimately show ?thesis
using \<mu> \<phi> iso_is_arr comp_arr_dom comp_cod_arr
interchange [of \<mu> "dom \<mu>" "cod \<phi>" \<phi>]
by simp
qed
thus ?thesis
using comp_assoc by simp
qed
also have "... = \<rr> \<mu> \<cdot> (dom \<mu> \<star> \<phi>)"
proof -
have "\<mu> \<star> cod \<phi> = R \<mu>"
using \<mu> \<phi> arr_simps(1) in_homE by auto
hence "\<rr> (cod \<mu>) \<cdot> (\<mu> \<star> cod \<phi>) = \<rr> \<mu>"
using \<mu> \<phi> \<rr>.is_natural_2 by simp
thus ?thesis by simp
qed
finally show ?thesis by simp
qed
next
show "\<And>f. ide f \<Longrightarrow> iso (\<rr> f \<cdot> (dom f \<star> \<phi>))"
proof -
fix f
assume f: "ide f"
have 1: "iso (\<rr> f)"
using f iso_lunit by simp
moreover have 2: "iso (dom f \<star> \<phi>)"
using \<phi> f src_def trg_def ide_char arr_char
apply (intro iso_hcomp, simp_all)
by (metis (no_types, lifting) in_homE)
moreover have "seq (\<rr> f) (dom f \<star> \<phi>)"
proof (intro seqI')
show "\<guillemotleft>\<rr> f : f \<star> a \<Rightarrow> f\<guillemotright>"
using f runit_in_hom(2) \<rr>_ide_simp ide_char arr_char src_def by simp
show "\<guillemotleft>dom f \<star> \<phi> : f \<star> w \<Rightarrow> f \<star> a\<guillemotright>"
using \<phi> f ide_char arr_char hcomp_def src_def trg_def obj_a ide_in_hom
in_hom_char
by (intro hcomp_in_vhom, auto)
qed
ultimately show "iso (\<rr> f \<cdot> (dom f \<star> \<phi>))"
using isos_compose by simp
qed
qed
show "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (f, w) \<star> snd (f, w))"
using \<r>.natural_isomorphism_axioms R.isomorphic_to_identity_is_equivalence by simp
qed
interpretation R: equivalence_functor \<open>(\<cdot>)\<close> \<open>(\<cdot>)\<close> \<open>\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>)\<close>
proof -
have "(\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>)) = (\<lambda>f. fst (f, w) \<star> snd (f, w))"
using \<omega>_in_vhom by simp
thus "equivalence_functor (\<cdot>) (\<cdot>) (\<lambda>f. fst (f, cod \<omega>) \<star> snd (f, cod \<omega>))"
using R'.equivalence_functor_axioms by simp
qed
interpretation M: monoidal_category \<open>(\<cdot>)\<close> \<open>\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>\<close> \<alpha> \<omega>
proof
show "\<guillemotleft>\<omega> : fst (cod \<omega>, cod \<omega>) \<star> snd (cod \<omega>, cod \<omega>) \<Rightarrow> cod \<omega>\<guillemotright>"
using \<omega>_in_vhom hcomp_def arr_char by auto
show "iso \<omega>"
using \<omega>_is_iso iso_char arr_char inv_char \<omega>_in_vhom by auto
show "\<And>f g h k. \<lbrakk> ide f; ide g; ide h; ide k \<rbrakk> \<Longrightarrow>
(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, hcomp (fst (g, h)) (snd (g, h)), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
\<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> snd (f, g), h, k)"
proof -
fix f g h k
assume f: "ide f" and g: "ide g" and h: "ide h" and k: "ide k"
have 1: "VVV.arr (f, g, h) \<and> VVV.arr (g, h, k)"
using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
by simp
have 2: "VVV.arr (f, g \<star> h, k)"
using f g h k 1 HoHV_def VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp hseqI'
by auto
have 3: "VVV.arr (f, g, h \<star> k)"
using f g h k 1 VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp H.preserves_reflects_arr hseqI'
by auto
have 4: "VVV.arr (f \<star> g, h, k)"
using f g h k VVV.arr_char VV.arr_char src_def trg_def ide_char arr_char hseq_char
VxV.arrI VxVxV.arrI VxVxV_comp_eq_VVV_comp
by force
have "(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, fst (g, h) \<star> snd (g, h), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
(f \<star> \<a>\<^sub>B[g, h, k]) \<cdot> \<a>\<^sub>B[f, g \<star> h, k] \<cdot> (\<a>\<^sub>B[f, g, h] \<star> k)"
unfolding \<alpha>_def by (simp add: 1 2)
also have "... = (f \<star>\<^sub>B \<a>\<^sub>B g h k) \<cdot> \<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot> (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
unfolding hcomp_def
using f g h k src_def trg_def arr_char
using assoc_closed ide_char by auto
also have "... = (f \<star>\<^sub>B \<a>\<^sub>B g h k) \<cdot>\<^sub>B \<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
proof -
have "arr (f \<star>\<^sub>B \<a>\<^sub>B g h k)"
using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f (g \<star>\<^sub>B h) k)"
using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f g h \<star>\<^sub>B k)"
using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B (\<a>\<^sub>B f g h \<star>\<^sub>B k))"
unfolding arr_char
apply (intro conjI)
using ide_char arr_char assoc_closed f g h hcomp_closed k B.HoHV_def B.HoVH_def
apply (intro B.seqI)
apply simp_all
proof -
have 1: "B.arr (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k)"
using f g h k ide_char arr_char B.HoHV_def B.HoVH_def
apply (intro B.seqI)
by auto
show "src\<^sub>B (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k) = a"
using 1 f g h k arr_char B.src_vcomp B.vseq_implies_hpar(1) by fastforce
show "trg\<^sub>B (\<a>\<^sub>B f (g \<star>\<^sub>B h) k \<cdot>\<^sub>B \<a>\<^sub>B f g h \<star>\<^sub>B k) = a"
using "1" arr_char calculation(2-3) by auto
qed
ultimately show ?thesis
using B.ext comp_char by (metis (no_types, lifting))
qed
also have "... = \<a>\<^sub>B f g (h \<star>\<^sub>B k) \<cdot>\<^sub>B \<a>\<^sub>B (f \<star>\<^sub>B g) h k"
using f g h k src_def trg_def arr_char ide_char B.pentagon
using "4" \<alpha>_def hcomp_def by auto
also have "... = \<a>\<^sub>B f g (h \<star>\<^sub>B k) \<cdot> \<a>\<^sub>B (f \<star>\<^sub>B g) h k"
proof -
have "arr (\<a>\<^sub>B (f \<star>\<^sub>B g) h k)"
using ide_char arr_char assoc_closed f g h hcomp_closed k by simp
moreover have "arr (\<a>\<^sub>B f g (h \<star>\<^sub>B k))"
using ide_char arr_char assoc_closed f g h hcomp_closed k by fastforce
ultimately show ?thesis
using B.ext comp_char by auto
qed
also have "... = \<a>\<^sub>B[f, g, fst (h, k) \<star> snd (h, k)] \<cdot> \<a>\<^sub>B[fst (f, g) \<star> snd (f, g), h, k]"
unfolding hcomp_def
using f g h k src_def trg_def arr_char ide_char by simp
also have "... = \<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> snd (f, g), h, k)"
unfolding \<alpha>_def using 1 2 3 4 by simp
finally show "(fst (f, \<alpha> (g, h, k)) \<star> snd (f, \<alpha> (g, h, k))) \<cdot>
\<alpha> (f, fst (g, h) \<star> snd (g, h), k) \<cdot>
(fst (\<alpha> (f, g, h), k) \<star> snd (\<alpha> (f, g, h), k)) =
\<alpha> (f, g, fst (h, k) \<star> snd (h, k)) \<cdot> \<alpha> (fst (f, g) \<star> snd (f, g), h, k)"
by simp
qed
qed
proposition is_monoidal_category:
shows "monoidal_category (\<cdot>) (\<lambda>\<mu>\<nu>. fst \<mu>\<nu> \<star> snd \<mu>\<nu>) \<alpha> \<omega>"
..
end
text \<open>
In a bicategory, the ``objects'' are essentially arbitrarily chosen representatives
of their isomorphism classes. Choosing any other representatives results in an
equivalent structure. Each object \<open>a\<close> is additionally equipped with an arbitrarily chosen
unit isomorphism \<open>\<guillemotleft>\<iota> : a \<star> a \<Rightarrow> a\<guillemotright>\<close>. For any \<open>(a, \<iota>)\<close> and \<open>(a', \<iota>')\<close>,
where \<open>a\<close> and \<open>a'\<close> are isomorphic to the same object, there exists a unique isomorphism
\<open>\<guillemotleft>\<psi>: a \<Rightarrow> a'\<guillemotright>\<close> that is compatible with the chosen unit isomorphisms \<open>\<iota>\<close> and \<open>\<iota>'\<close>.
We have already proved this property for monoidal categories, which are bicategories
with just one ``object''. Here we use that already-proven property to establish its
generalization to arbitary bicategories, by exploiting the fact that if \<open>a\<close> is an object
in a bicategory, then the sub-bicategory consisting of all \<open>\<mu>\<close> such that
\<open>src \<mu> = a = trg \<mu>\<close>, is a monoidal category.
At some point it would potentially be nicer to transfer the proof for monoidal
categories to obtain a direct, ``native'' proof of this fact for bicategories.
\<close>
lemma (in bicategory) unit_unique_upto_unique_iso:
assumes "obj a"
and "isomorphic a w"
and "\<guillemotleft>\<omega> : w \<star> w \<Rightarrow> w\<guillemotright>"
and "iso \<omega>"
shows "\<exists>!\<psi>. \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright> \<and> iso \<psi> \<and> \<psi> \<cdot> \<i>[a] = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
proof -
have \<omega>_in_hhom: "\<guillemotleft>\<omega> : a \<rightarrow> a\<guillemotright>"
using assms
apply (intro in_hhomI)
apply auto
apply (metis src_cod in_homE isomorphic_implies_hpar(3) objE)
by (metis trg_cod in_homE isomorphic_implies_hpar(4) objE)
interpret S: subbicategory V H \<a> \<i> src trg \<open>\<lambda>\<mu>. arr \<mu> \<and> src \<mu> = a \<and> trg \<mu> = a\<close>
using assms iso_unit in_homE isoE isomorphicE VVV.arr_char VV.arr_char
apply unfold_locales
apply auto[7]
proof
fix f g h
assume f: "(arr f \<and> src f = a \<and> trg f = a) \<and> ide f"
and g: "(arr g \<and> src g = a \<and> trg g = a) \<and> ide g"
and h: "(arr h \<and> src h = a \<and> trg h = a) \<and> ide h"
and fg: "src f = trg g" and gh: "src g = trg h"
show "arr (\<a>[f, g, h])"
using assms f g h fg gh by auto
show "src (\<a>[f, g, h]) = a \<and> trg (\<a>[f, g, h]) = a"
using assms f g h fg gh by auto
show "arr (inv (\<a>[f, g, h])) \<and> src (inv (\<a>[f, g, h])) = a \<and> trg (inv (\<a>[f, g, h])) = a"
using assms f g h fg gh \<alpha>.preserves_hom src_dom trg_dom by simp
next
fix f
assume f: "arr f \<and> src f = a \<and> trg f = a"
assume ide_left: "ide f"
show "arr (\<ll> f) \<and> src (\<ll> f) = a \<and> trg (\<ll> f) = a"
using f assms(1) \<ll>.preserves_hom src_cod [of "\<ll> f"] trg_cod [of "\<ll> f"] by simp
show "arr (inv (\<ll> f)) \<and> src (inv (\<ll> f)) = a \<and> trg (inv (\<ll> f)) = a"
using f ide_left assms(1) \<ll>'.preserves_hom src_dom [of "\<ll>'.map f"] trg_dom [of "\<ll>'.map f"]
by simp
show "arr (\<rr> f) \<and> src (\<rr> f) = a \<and> trg (\<rr> f) = a"
using f assms(1) \<rr>.preserves_hom src_cod [of "\<rr> f"] trg_cod [of "\<rr> f"] by simp
show "arr (inv (\<rr> f)) \<and> src (inv (\<rr> f)) = a \<and> trg (inv (\<rr> f)) = a"
using f ide_left assms(1) \<rr>'.preserves_hom src_dom [of "\<rr>'.map f"] trg_dom [of "\<rr>'.map f"]
by simp
qed
interpret S: subbicategory_at_object V H \<a> \<i> src trg a a \<open>\<i>[a]\<close>
proof
show "obj a" by fact
show "isomorphic a a"
using assms(1) isomorphic_reflexive by blast
show "S.in_hom \<i>[a] (a \<star> a) a"
using S.arr_char S.in_hom_char assms(1) by fastforce
show "iso \<i>[a]"
using assms iso_unit by simp
qed
interpret S\<^sub>\<omega>: subbicategory_at_object V H \<a> \<i> src trg a w \<omega>
proof
show "obj a" by fact
show "iso \<omega>" by fact
show "isomorphic a w"
using assms by simp
show "S.in_hom \<omega> (w \<star> w) w"
using assms S.arr_char S.dom_char S.cod_char \<omega>_in_hhom
by (intro S.in_homI, auto)
qed
interpret M: monoidal_category S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<open>\<i>[a]\<close>
using S.is_monoidal_category by simp
interpret M\<^sub>\<omega>: monoidal_category S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<omega>
using S\<^sub>\<omega>.is_monoidal_category by simp
interpret M: monoidal_category_with_alternate_unit
S.comp \<open>\<lambda>\<mu>\<nu>. S.hcomp (fst \<mu>\<nu>) (snd \<mu>\<nu>)\<close> S.\<alpha> \<open>\<i>[a]\<close> \<omega> ..
have 1: "M\<^sub>\<omega>.unity = w"
using assms M\<^sub>\<omega>.unity_def S.cod_char S.arr_char
by (metis (no_types, lifting) S.in_homE S\<^sub>\<omega>.\<omega>_in_vhom)
have 2: "M.unity = a"
using assms M.unity_def S.cod_char S.arr_char by simp
have "\<exists>!\<psi>. S.in_hom \<psi> a w \<and> S.iso \<psi> \<and> S.comp \<psi> \<i>[a] = S.comp \<omega> (M.tensor \<psi> \<psi>)"
using assms 1 2 M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\<omega>.unity_def S.cod_char
by simp
show "\<exists>!\<psi>. \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright> \<and> iso \<psi> \<and> \<psi> \<cdot> \<i>[a] = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
proof -
have 1: "\<And>\<psi>. S.in_hom \<psi> a w \<longleftrightarrow> \<guillemotleft>\<psi> : a \<Rightarrow> w\<guillemotright>"
using assms S.in_hom_char S.arr_char
by (metis (no_types, lifting) S.ideD(1) S.w_simps(1) S\<^sub>\<omega>.w_simps(1) in_homE
src_dom trg_dom)
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.iso \<psi> \<longleftrightarrow> iso \<psi>"
using assms S.in_hom_char S.arr_char S.iso_char by auto
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> M.tensor \<psi> \<psi> = \<psi> \<star> \<psi>"
using assms S.in_hom_char S.arr_char S.hcomp_def by simp
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.comp \<psi> \<i>[a] = \<psi> \<cdot> \<i>[a]"
using assms S.in_hom_char S.comp_char by auto
moreover have "\<And>\<psi>. S.in_hom \<psi> a w \<Longrightarrow> S.comp \<omega> (M.tensor \<psi> \<psi>) = \<omega> \<cdot> (\<psi> \<star> \<psi>)"
using assms S.in_hom_char S.arr_char S.hcomp_def S.comp_char S.dom_char S.cod_char
by (metis (no_types, lifting) M\<^sub>\<omega>.arr_tensor S\<^sub>\<omega>.\<omega>_simps(1) calculation(3) ext)
ultimately show ?thesis
by (metis (no_types, lifting) M.unit_unique_upto_unique_iso M.unity_def M\<^sub>\<omega>.unity_def
S.\<omega>_in_vhom S.in_homE S\<^sub>\<omega>.\<omega>_in_vhom)
qed
qed
end
diff --git a/thys/Binding_Syntax_Theory/Equiv_Relation2.thy b/thys/Binding_Syntax_Theory/Equiv_Relation2.thy
--- a/thys/Binding_Syntax_Theory/Equiv_Relation2.thy
+++ b/thys/Binding_Syntax_Theory/Equiv_Relation2.thy
@@ -1,195 +1,195 @@
-section {* Some preliminaries on equivalence relations and quotients *}
+section \<open>Some preliminaries on equivalence relations and quotients\<close>
theory Equiv_Relation2 imports Preliminaries Pick
begin
-text{* Unary predicates vs. sets: *}
+text\<open>Unary predicates vs. sets:\<close>
definition "S2P A \<equiv> \<lambda> x. x \<in> A"
lemma S2P_app[simp]: "S2P r x \<longleftrightarrow> x \<in> r"
unfolding S2P_def by auto
lemma S2P_Collect[simp]: "S2P (Collect \<phi>) = \<phi>"
apply(rule ext)+ by simp
lemma Collect_S2P[simp]: "Collect (S2P r) = r"
by (metis Collect_mem_eq S2P_Collect)
-text{* Binary predicates vs. relatipons: *}
+text\<open>Binary predicates vs. relatipons:\<close>
definition "P2R \<phi> \<equiv> {(x,y). \<phi> x y}"
definition "R2P r \<equiv> \<lambda> x y. (x,y) \<in> r"
lemma in_P2R[simp]: "xy \<in> P2R \<phi> \<longleftrightarrow> \<phi> (fst xy) (snd xy)"
unfolding P2R_def by auto
lemma in_P2R_pair[simp]: "(x,y) \<in> P2R \<phi> \<longleftrightarrow> \<phi> x y"
by simp
lemma R2P_app[simp]: "R2P r x y \<longleftrightarrow> (x,y) \<in> r"
unfolding R2P_def by auto
lemma R2P_P2R[simp]: "R2P (P2R \<phi>) = \<phi>"
apply(rule ext)+ by simp
lemma P2R_R2P[simp]: "P2R (R2P r) = r"
using Collect_mem_eq P2R_def R2P_P2R case_prod_curry by metis
definition "reflP P \<phi> \<equiv> (\<forall> x y. \<phi> x y \<or> \<phi> y x \<longrightarrow> P x) \<and> (\<forall> x. P x \<longrightarrow> \<phi> x x)"
definition "symP \<phi> \<equiv> \<forall> x y. \<phi> x y \<longrightarrow> \<phi> y x"
definition transP where "transP \<phi> \<equiv> \<forall> x y z. \<phi> x y \<and> \<phi> y z \<longrightarrow> \<phi> x z"
definition "equivP A \<phi> \<equiv> reflP A \<phi> \<and> symP \<phi> \<and> transP \<phi>"
lemma refl_on_P2R[simp]: "refl_on (Collect P) (P2R \<phi>) \<longleftrightarrow> reflP P \<phi>"
unfolding reflP_def refl_on_def by force
lemma reflP_R2P[simp]: "reflP (S2P A) (R2P r) \<longleftrightarrow> refl_on A r"
unfolding reflP_def refl_on_def by auto
lemma sym_P2R[simp]: "sym (P2R \<phi>) \<longleftrightarrow> symP \<phi>"
unfolding symP_def sym_def by auto
lemma symP_R2P[simp]: "symP (R2P r) \<longleftrightarrow> sym r"
unfolding symP_def sym_def by auto
lemma trans_P2R[simp]: "trans (P2R \<phi>) \<longleftrightarrow> transP \<phi>"
unfolding transP_def trans_def by auto
lemma transP_R2P[simp]: "transP (R2P r) \<longleftrightarrow> trans r"
unfolding transP_def trans_def by auto
lemma equiv_P2R[simp]: "equiv (Collect P) (P2R \<phi>) \<longleftrightarrow> equivP P \<phi>"
unfolding equivP_def equiv_def by auto
lemma equivP_R2P[simp]: "equivP (S2P A) (R2P r) \<longleftrightarrow> equiv A r"
unfolding equivP_def equiv_def by auto
lemma in_P2R_Im_singl[simp]: "y \<in> P2R \<phi> `` {x} \<longleftrightarrow> \<phi> x y" by simp
definition proj :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a set" where
"proj \<phi> x \<equiv> {y. \<phi> x y}"
lemma proj_P2R: "proj \<phi> x = P2R \<phi> `` {x}" unfolding proj_def by auto
lemma proj_P2R_raw: "proj \<phi> = (\<lambda> x. P2R \<phi> `` {x})"
apply(rule ext) unfolding proj_P2R ..
definition univ :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a set \<Rightarrow> 'b)"
where "univ f X == f (SOME x. x \<in> X)"
definition quotientP ::
"('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a set \<Rightarrow> bool)" (infixl "'/'/'/" 90)
where "P /// \<phi> \<equiv> S2P ((Collect P) // (P2R \<phi>))"
lemma proj_preserves:
"P x \<Longrightarrow> (P /// \<phi>) (proj \<phi> x)"
unfolding proj_P2R quotientP_def
by (metis S2P_def mem_Collect_eq quotientI)
lemma proj_in_iff:
assumes "equivP P \<phi>"
shows "(P///\<phi>) (proj \<phi> x) \<longleftrightarrow> P x"
using assms unfolding quotientP_def proj_def
by (metis (mono_tags) Collect_mem_eq Equiv_Relation2.proj_def
Equiv_Relation2.proj_preserves S2P_Collect empty_Collect_eq equivP_def
equiv_P2R in_quotient_imp_non_empty quotientP_def reflP_def)
lemma proj_iff[simp]:
"\<lbrakk>equivP P \<phi>; P x; P y\<rbrakk> \<Longrightarrow> proj \<phi> x = proj \<phi> y \<longleftrightarrow> \<phi> x y"
unfolding proj_P2R
by (metis (full_types) equiv_P2R equiv_class_eq_iff equiv_class_self
in_P2R_pair mem_Collect_eq proj_P2R proj_def)
lemma in_proj[simp]: "\<lbrakk>equivP P \<phi>; P x\<rbrakk> \<Longrightarrow> x \<in> proj \<phi> x"
unfolding proj_P2R equiv_def refl_on_def equiv_P2R[symmetric]
by auto
lemma proj_image[simp]: "(proj \<phi>) ` (Collect P) = Collect (P///\<phi>)"
unfolding proj_P2R_raw quotientP_def quotient_def by auto
lemma in_quotientP_imp_non_empty:
assumes "equivP P \<phi>" and "(P///\<phi>) X"
shows "X \<noteq> {}"
by (metis R2P_P2R S2P_Collect S2P_def assms equivP_R2P
in_quotient_imp_non_empty quotientP_def)
lemma in_quotientP_imp_in_rel:
"\<lbrakk>equivP P \<phi>; (P///\<phi>) X; x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> \<phi> x y"
unfolding equiv_P2R[symmetric] quotientP_def quotient_eq_iff
by (metis S2P_def in_P2R_pair quotient_eq_iff)
lemma in_quotientP_imp_closed:
"\<lbrakk>equivP P \<phi>; (P///\<phi>) X; x \<in> X; \<phi> x y\<rbrakk> \<Longrightarrow> y \<in> X"
using S2P_Collect S2P_def equivP_def proj_P2R_raw proj_def
quotientE quotientP_def transP_def
by metis
lemma in_quotientP_imp_subset:
assumes "equivP P \<phi>" and "(P///\<phi>) X"
shows "X \<subseteq> Collect P"
by (metis (mono_tags, lifting) CollectI assms equivP_def in_quotientP_imp_in_rel reflP_def subsetI)
lemma equivP_pick_in:
assumes "equivP P \<phi> " and "(P///\<phi>) X"
shows "pick X \<in> X"
by (metis assms in_quotientP_imp_non_empty pick_NE)
lemma equivP_pick_preserves:
assumes "equivP P \<phi> " and "(P///\<phi>) X"
shows "P (pick X)"
by (metis assms equivP_pick_in in_quotientP_imp_subset mem_Collect_eq set_rev_mp)
lemma proj_pick:
assumes \<phi>: "equivP P \<phi>" and X: "(P///\<phi>) X"
shows "proj \<phi> (pick X) = X"
by (smt proj_def Equiv_Relation2.proj_iff Equiv_Relation2.proj_image X
\<phi> equivP_pick_in equivP_pick_preserves image_iff mem_Collect_eq)
lemma pick_proj:
assumes "equivP P \<phi>" and "P x"
shows "\<phi> (pick (proj \<phi> x)) x"
by (metis assms equivP_def in_proj mem_Collect_eq pick proj_def symP_def)
lemma equivP_pick_iff[simp]:
assumes \<phi>: "equivP P \<phi>" and X: "(P///\<phi>) X" and Y: "(P///\<phi>) Y"
shows "\<phi> (pick X) (pick Y) \<longleftrightarrow> X = Y"
by (metis Equiv_Relation2.proj_iff X Y \<phi> equivP_pick_preserves proj_pick)
lemma equivP_pick_inj_on:
assumes "equivP P \<phi>"
shows "inj_on pick (Collect (P///\<phi>))"
using assms unfolding inj_on_def
by (metis assms equivP_pick_iff mem_Collect_eq)
definition congruentP where
"congruentP \<phi> f \<equiv> \<forall> x y. \<phi> x y \<longrightarrow> f x = f y"
abbreviation RESPECTS_P (infixr "respectsP" 80) where
"f respectsP r == congruentP r f"
lemma congruent_P2R: "congruent (P2R \<phi>) f = congruentP \<phi> f"
unfolding congruent_def congruentP_def by auto
lemma univ_commute[simp]:
assumes "equivP P \<phi>" and "f respectsP \<phi>" and "P x"
shows "(univ f) (proj \<phi> x) = f x"
unfolding congruent_P2R[symmetric]
by (metis (full_types) assms pick_def congruentP_def pick_proj univ_def)
lemma univ_unique:
assumes "equivP P \<phi>" and "f respectsP \<phi>" and "\<And> x. P x \<Longrightarrow> G (proj \<phi> x) = f x"
shows "\<forall> X. (P///\<phi>) X \<longrightarrow> G X = univ f X"
by (metis assms equivP_pick_preserves proj_pick univ_commute)
lemma univ_preserves:
assumes "equivP P \<phi> " and "f respectsP \<phi>" and "\<And> x. P x \<Longrightarrow> f x \<in> B"
shows "\<forall> X. (P///\<phi>) X \<longrightarrow> univ f X \<in> B"
by (metis Equiv_Relation2.univ_commute assms
equivP_pick_preserves proj_pick)
end
diff --git a/thys/Binding_Syntax_Theory/Iteration.thy b/thys/Binding_Syntax_Theory/Iteration.thy
--- a/thys/Binding_Syntax_Theory/Iteration.thy
+++ b/thys/Binding_Syntax_Theory/Iteration.thy
@@ -1,4186 +1,4186 @@
-section {* Iteration *}
+section \<open>Iteration\<close>
theory Iteration imports Well_Sorted_Terms
begin
-text{* In this section, we introduce first-order models (models, for short).
+text\<open>In this section, we introduce first-order models (models, for short).
These are structures having operators that
match those for terms (including variable-injection, binding operations, freshness,
swapping and substitution) and satisfy some clauses,
and show that terms form initial models. This gives iteration principles.
As a matter of notation: the prefix
``g" will stand for ``generalized" -- elements of models are referred to as ``generalized terms".
The actual full prefix will be ``ig" (where ``i" stands for ``iteration"), symbolizing the fact that
the models from this section support iteration, and not general recursion.
The latter is dealt with by the models introduced in the next section, for which we
use the simple prefix ``g".
-*}
-
-subsection {* Models *}
-
-text{* We have two basic kinds of models:
+\<close>
+
+subsection \<open>Models\<close>
+
+text\<open>We have two basic kinds of models:
\\- fresh-swap (FSw) models, featuring operations corresponding to
the concrete syntactic constructs (``Var", ``Op", ``Abs"),
henceforth referred to simply as {\em the constructs}, and to fresh and swap;
\\- fresh-swap-subst (FSb) models, featuring substitution instead of swapping.
We also consider two combinations of the above, FSwSb-models and FSbSw-models.
To keep things structurally
simple, we use one single Isabelle for all the 4 kinds models,
allowing the most generous signature.
Since terms are the main actors of our theory, models being considered only
for the sake of recursive definitions, we call the items inhabiting these models
``generalized" terms, abstractions and inputs, and correspondingly
the operations; hence the prefix ``g" from the names of the type parameters and
operators.
(However,
we refer to the generalized items using the same notations as for
``concrete items": X, A, etc.)
%
Indeed, a model can be regarded as implementing
a generalization/axiomatization of the term structure, where now the objects are
not terms, but do have term-like properties.
-*}
-
-subsubsection {* Raw models *}
+\<close>
+
+subsubsection \<open>Raw models\<close>
record ('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model =
igWls :: "'sort \<Rightarrow> 'gTerm \<Rightarrow> bool"
igWlsAbs :: "'varSort \<times> 'sort \<Rightarrow> 'gAbs \<Rightarrow> bool"
(* *)
igVar :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'gTerm"
igAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'gTerm \<Rightarrow> 'gAbs"
igOp :: "'opSym \<Rightarrow> ('index,'gTerm)input \<Rightarrow> ('bindex,'gAbs)input \<Rightarrow> 'gTerm"
(* *)
igFresh :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'gTerm \<Rightarrow> bool"
igFreshAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'gAbs \<Rightarrow> bool"
(* *)
igSwap :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'gTerm \<Rightarrow> 'gTerm"
igSwapAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'gAbs \<Rightarrow> 'gAbs"
(* *)
igSubst :: "'varSort \<Rightarrow> 'gTerm \<Rightarrow> 'var \<Rightarrow> 'gTerm \<Rightarrow> 'gTerm"
igSubstAbs :: "'varSort \<Rightarrow> 'gTerm \<Rightarrow> 'var \<Rightarrow> 'gAbs \<Rightarrow> 'gAbs"
-text{* \
+text\<open>\
\\- ``igSwap MOD zs z1 z2 X" swaps in X z1 and z2 (assumed of sorts zs).
-\\- ``igSubst MOD ys Y x X" substitutes, in X, Y with y (assumed of sort ys). *}
+\\- ``igSubst MOD ys Y x X" substitutes, in X, Y with y (assumed of sort ys).\<close>
definition igFreshInp where
"igFreshInp MOD ys y inp == liftAll (igFresh MOD ys y) inp"
definition igFreshBinp where
"igFreshBinp MOD ys y binp == liftAll (igFreshAbs MOD ys y) binp"
definition igSwapInp where
"igSwapInp MOD zs z1 z2 inp == lift (igSwap MOD zs z1 z2) inp"
definition igSwapBinp where
"igSwapBinp MOD zs z1 z2 binp == lift (igSwapAbs MOD zs z1 z2) binp"
definition igSubstInp where
"igSubstInp MOD ys Y y inp == lift (igSubst MOD ys Y y) inp"
definition igSubstBinp where
"igSubstBinp MOD ys Y y binp == lift (igSubstAbs MOD ys Y y) binp"
(* *************************************************** *)
context FixSyn
begin
(* In this theory, new type variables are introduced into the context,
corresponding to 'gTerm and 'gAbs (making the locale polymorphic). *)
-subsubsection {* Well-sorted models of various kinds*}
-
-text{* We define the following kinds of well-sorted models
+subsubsection \<open>Well-sorted models of various kinds\<close>
+
+text\<open>We define the following kinds of well-sorted models
\\- fresh-swap models (predicate ``iwlsFSw");
\\- fresh-subst models (``iwlsFSb");
\\- fresh-swap-subst models (``iwlsFSwSb");
\\- fresh-subst-swap models (``iwlsFSbSw").
All of these models are defined as raw models subject to various Horn conditions:
\\- For ``iwlsFSw":
\\--- definition-like clauses for ``fresh" and ``swap" in terms of the
construct operators;
\\--- congruence for abstraction based on fresh and swap (mirroring the abstraction case in
the definition of alpha-equivalence for quasi-terms).
%
\footnote{Here, by ``congruence for abstraction" we do not mean the standard notion of congrunece
(satisfied by any operator once or ever), but a {\em stronger} notion: in order for two abstractions
to be equal, it is not required that their ariguments be equal, but that they be in a
``permutative" relationship based either on swapping or on substitution.}
%
\\- For ``iwlsFSb": the same as for ``iwlsFSw", except that:
\\--- ``swap" is replaced by ``subst";
%
\footnote{
Note that traditionally alpha-equivalence is defined using ``subst", not ``swap".
}
%
\\--- The [fresh and swap]-based congrunce clause is replaced by an ``abstraction-renaming" clause,
which is stronger than the corresponding [fresh and subst]-based congruence clause.
%
\footnote{
We also define the [fresh and subst]-based congruence clause, although we do not
employ it directly in the definition of any kind of model.
}
%
\\- For ``iwlsFSwSb": the clauses for ``iwlsFSw", plus some of the definition-like clauses for ``subst".
%
\footnote{Not all the ``subst" definition-like clauses from ``iwlsFSb" are required
for ``iwlsFSwSb" -- namely, the clause that
we call ``igSubstIGAbsCls2" is not required here.
}
%
\\- For ``iwlsFSbSw": the clauses for ``iwlsFSb", plus definition-like clauses for ``swap".
Thus, a fresh-swap-subst model is also a fresh-swap model, and
a fresh-subst-swap model is also a fresh-subst model.
For convenience, all these 4 kinds of models are defined on one single type, that of {\em raw models},
which interpret the most generous signature, comprizing all the operations and relations required by all
4 kinds of models.
Note that, although some operations (namely, ``subst" or ``swap") may not be involved in the clauses for certain kinds
of models, the extra structure is harmless to the development of their theory.
Note that for the models operations and relations we do not actually write ``fresh", ``swap" and ``subst", but
``igFresh", ``igSwap" and ``igSubst".
As usual, we shall have not only term versions, but also abstraction versions of the above
operations.
- *}
+\<close>
definition igWlsInp where
"igWlsInp MOD delta inp ==
wlsOpS delta \<and> sameDom (arOf delta) inp \<and> liftAll2 (igWls MOD) (arOf delta) inp"
lemmas igWlsInp_defs = igWlsInp_def sameDom_def liftAll2_def
definition igWlsBinp where
"igWlsBinp MOD delta binp ==
wlsOpS delta \<and> sameDom (barOf delta) binp \<and> liftAll2 (igWlsAbs MOD) (barOf delta) binp"
lemmas igWlsBinp_defs = igWlsBinp_def sameDom_def liftAll2_def
-text{* Domain disjointness: *}
+text\<open>Domain disjointness:\<close>
definition igWlsDisj where
"igWlsDisj MOD == \<forall> s s' X. igWls MOD s X \<and> igWls MOD s' X \<longrightarrow> s = s'"
definition igWlsAbsDisj where
"igWlsAbsDisj MOD ==
\<forall> xs s xs' s' A.
isInBar (xs,s) \<and> isInBar (xs',s') \<and>
igWlsAbs MOD (xs,s) A \<and> igWlsAbs MOD (xs',s') A
\<longrightarrow> xs = xs' \<and> s = s'"
definition igWlsAllDisj where
"igWlsAllDisj MOD ==
igWlsDisj MOD \<and> igWlsAbsDisj MOD"
lemmas igWlsAllDisj_defs =
igWlsAllDisj_def
igWlsDisj_def igWlsAbsDisj_def
-text {* Abstration domains inhabited only within bound arities: *}
+text \<open>Abstration domains inhabited only within bound arities:\<close>
definition igWlsAbsIsInBar where
"igWlsAbsIsInBar MOD ==
\<forall> us s A. igWlsAbs MOD (us,s) A \<longrightarrow> isInBar (us,s)"
-text{* Domain preservation by the operators: weak (``if") versions and strong (``iff") versions
-(for the latter, we use the suffix ``STR"): *}
-
-text{* The constructs preserve the domains: *}
+text\<open>Domain preservation by the operators: weak (``if") versions and strong (``iff") versions
+(for the latter, we use the suffix ``STR"):\<close>
+
+text\<open>The constructs preserve the domains:\<close>
definition igVarIPresIGWls where
"igVarIPresIGWls MOD ==
\<forall> xs x. igWls MOD (asSort xs) (igVar MOD xs x)"
definition igAbsIPresIGWls where
"igAbsIPresIGWls MOD ==
\<forall> xs s x X. isInBar (xs,s) \<and> igWls MOD s X \<longrightarrow>
igWlsAbs MOD (xs,s) (igAbs MOD xs x X)"
definition igAbsIPresIGWlsSTR where
"igAbsIPresIGWlsSTR MOD ==
\<forall> xs s x X. isInBar (xs,s) \<longrightarrow>
igWlsAbs MOD (xs,s) (igAbs MOD xs x X) =
igWls MOD s X"
lemma igAbsIPresIGWlsSTR_imp_igAbsIPresIGWls:
"igAbsIPresIGWlsSTR MOD \<Longrightarrow> igAbsIPresIGWls MOD"
unfolding igAbsIPresIGWlsSTR_def igAbsIPresIGWls_def by simp
definition igOpIPresIGWls where
"igOpIPresIGWls MOD ==
\<forall> delta inp binp.
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp
\<longrightarrow> igWls MOD (stOf delta) (igOp MOD delta inp binp)"
definition igOpIPresIGWlsSTR where
"igOpIPresIGWlsSTR MOD ==
\<forall> delta inp binp.
igWls MOD (stOf delta) (igOp MOD delta inp binp) =
(igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp)"
lemma igOpIPresIGWlsSTR_imp_igOpIPresIGWls:
"igOpIPresIGWlsSTR MOD \<Longrightarrow> igOpIPresIGWls MOD"
unfolding igOpIPresIGWlsSTR_def igOpIPresIGWls_def by simp
definition igConsIPresIGWls where
"igConsIPresIGWls MOD ==
igVarIPresIGWls MOD \<and>
igAbsIPresIGWls MOD \<and>
igOpIPresIGWls MOD"
lemmas igConsIPresIGWls_defs = igConsIPresIGWls_def
igVarIPresIGWls_def
igAbsIPresIGWls_def
igOpIPresIGWls_def
definition igConsIPresIGWlsSTR where
"igConsIPresIGWlsSTR MOD ==
igVarIPresIGWls MOD \<and>
igAbsIPresIGWlsSTR MOD \<and>
igOpIPresIGWlsSTR MOD"
lemmas igConsIPresIGWlsSTR_defs = igConsIPresIGWlsSTR_def
igVarIPresIGWls_def
igAbsIPresIGWlsSTR_def
igOpIPresIGWlsSTR_def
lemma igConsIPresIGWlsSTR_imp_igConsIPresIGWls:
"igConsIPresIGWlsSTR MOD \<Longrightarrow> igConsIPresIGWls MOD"
unfolding igConsIPresIGWlsSTR_def igConsIPresIGWls_def
using
igAbsIPresIGWlsSTR_imp_igAbsIPresIGWls
igOpIPresIGWlsSTR_imp_igOpIPresIGWls
by auto
(* The notion of ``fresh" preserving well-sorted-ness does not make sense, since
``fresh" is a relation. *)
-text{* ``swap" preserves the domains: *}
+text\<open>``swap" preserves the domains:\<close>
definition igSwapIPresIGWls where
"igSwapIPresIGWls MOD ==
\<forall> zs z1 z2 s X. igWls MOD s X \<longrightarrow>
igWls MOD s (igSwap MOD zs z1 z2 X)"
definition igSwapIPresIGWlsSTR where
"igSwapIPresIGWlsSTR MOD ==
\<forall> zs z1 z2 s X. igWls MOD s (igSwap MOD zs z1 z2 X) =
igWls MOD s X"
lemma igSwapIPresIGWlsSTR_imp_igSwapIPresIGWls:
"igSwapIPresIGWlsSTR MOD \<Longrightarrow> igSwapIPresIGWls MOD"
unfolding igSwapIPresIGWlsSTR_def igSwapIPresIGWls_def by simp
definition igSwapAbsIPresIGWlsAbs where
"igSwapAbsIPresIGWlsAbs MOD ==
\<forall> zs z1 z2 us s A.
isInBar (us,s) \<and> igWlsAbs MOD (us,s) A \<longrightarrow>
igWlsAbs MOD (us,s) (igSwapAbs MOD zs z1 z2 A)"
definition igSwapAbsIPresIGWlsAbsSTR where
"igSwapAbsIPresIGWlsAbsSTR MOD ==
\<forall> zs z1 z2 us s A.
igWlsAbs MOD (us,s) (igSwapAbs MOD zs z1 z2 A) =
igWlsAbs MOD (us,s) A"
lemma igSwapAbsIPresIGWlsAbsSTR_imp_igSwapAbsIPresIGWlsAbs:
"igSwapAbsIPresIGWlsAbsSTR MOD \<Longrightarrow> igSwapAbsIPresIGWlsAbs MOD"
unfolding igSwapAbsIPresIGWlsAbsSTR_def igSwapAbsIPresIGWlsAbs_def by simp
definition igSwapAllIPresIGWlsAll where
"igSwapAllIPresIGWlsAll MOD ==
igSwapIPresIGWls MOD \<and> igSwapAbsIPresIGWlsAbs MOD"
lemmas igSwapAllIPresIGWlsAll_defs = igSwapAllIPresIGWlsAll_def
igSwapIPresIGWls_def igSwapAbsIPresIGWlsAbs_def
definition igSwapAllIPresIGWlsAllSTR where
"igSwapAllIPresIGWlsAllSTR MOD ==
igSwapIPresIGWlsSTR MOD \<and> igSwapAbsIPresIGWlsAbsSTR MOD"
lemmas igSwapAllIPresIGWlsAllSTR_defs = igSwapAllIPresIGWlsAllSTR_def
igSwapIPresIGWlsSTR_def igSwapAbsIPresIGWlsAbsSTR_def
lemma igSwapAllIPresIGWlsAllSTR_imp_igSwapAllIPresIGWlsAll:
"igSwapAllIPresIGWlsAllSTR MOD \<Longrightarrow> igSwapAllIPresIGWlsAll MOD"
unfolding igSwapAllIPresIGWlsAllSTR_def igSwapAllIPresIGWlsAll_def
using
igSwapIPresIGWlsSTR_imp_igSwapIPresIGWls
igSwapAbsIPresIGWlsAbsSTR_imp_igSwapAbsIPresIGWlsAbs
by auto
-text{* ``subst" preserves the domains: *}
+text\<open>``subst" preserves the domains:\<close>
definition igSubstIPresIGWls where
"igSubstIPresIGWls MOD ==
\<forall> ys Y y s X. igWls MOD (asSort ys) Y \<and> igWls MOD s X \<longrightarrow>
igWls MOD s (igSubst MOD ys Y y X)"
definition igSubstIPresIGWlsSTR where
"igSubstIPresIGWlsSTR MOD ==
\<forall> ys Y y s X.
igWls MOD s (igSubst MOD ys Y y X) =
(igWls MOD (asSort ys) Y \<and> igWls MOD s X)"
lemma igSubstIPresIGWlsSTR_imp_igSubstIPresIGWls:
"igSubstIPresIGWlsSTR MOD \<Longrightarrow> igSubstIPresIGWls MOD"
unfolding igSubstIPresIGWlsSTR_def igSubstIPresIGWls_def by simp
definition igSubstAbsIPresIGWlsAbs where
"igSubstAbsIPresIGWlsAbs MOD ==
\<forall> ys Y y us s A.
isInBar (us,s) \<and> igWls MOD (asSort ys) Y \<and> igWlsAbs MOD (us,s) A \<longrightarrow>
igWlsAbs MOD (us,s) (igSubstAbs MOD ys Y y A)"
definition igSubstAbsIPresIGWlsAbsSTR where
"igSubstAbsIPresIGWlsAbsSTR MOD ==
\<forall> ys Y y us s A.
igWlsAbs MOD (us,s) (igSubstAbs MOD ys Y y A) =
(igWls MOD (asSort ys) Y \<and> igWlsAbs MOD (us,s) A)"
lemma igSubstAbsIPresIGWlsAbsSTR_imp_igSubstAbsIPresIGWlsAbs:
"igSubstAbsIPresIGWlsAbsSTR MOD \<Longrightarrow> igSubstAbsIPresIGWlsAbs MOD"
unfolding igSubstAbsIPresIGWlsAbsSTR_def igSubstAbsIPresIGWlsAbs_def by simp
definition igSubstAllIPresIGWlsAll where
"igSubstAllIPresIGWlsAll MOD ==
igSubstIPresIGWls MOD \<and> igSubstAbsIPresIGWlsAbs MOD"
lemmas igSubstAllIPresIGWlsAll_defs = igSubstAllIPresIGWlsAll_def
igSubstIPresIGWls_def igSubstAbsIPresIGWlsAbs_def
definition igSubstAllIPresIGWlsAllSTR where
"igSubstAllIPresIGWlsAllSTR MOD ==
igSubstIPresIGWlsSTR MOD \<and> igSubstAbsIPresIGWlsAbsSTR MOD"
lemmas igSubstAllIPresIGWlsAllSTR_defs = igSubstAllIPresIGWlsAllSTR_def
igSubstIPresIGWlsSTR_def igSubstAbsIPresIGWlsAbsSTR_def
lemma igSubstAllIPresIGWlsAllSTR_imp_igSubstAllIPresIGWlsAll:
"igSubstAllIPresIGWlsAllSTR MOD \<Longrightarrow> igSubstAllIPresIGWlsAll MOD"
unfolding igSubstAllIPresIGWlsAllSTR_def igSubstAllIPresIGWlsAll_def
using
igSubstIPresIGWlsSTR_imp_igSubstIPresIGWls
igSubstAbsIPresIGWlsAbsSTR_imp_igSubstAbsIPresIGWlsAbs
by auto
-text{* Clauses for fresh: fully conditional versions and less conditional,
-stronger versions (the latter having suffix ``STR"). *}
+text\<open>Clauses for fresh: fully conditional versions and less conditional,
+stronger versions (the latter having suffix ``STR").\<close>
definition igFreshIGVar where
"igFreshIGVar MOD ==
\<forall> ys y xs x.
ys \<noteq> xs \<or> y \<noteq> x \<longrightarrow>
igFresh MOD ys y (igVar MOD xs x)"
definition igFreshIGAbs1 where
"igFreshIGAbs1 MOD ==
\<forall> ys y s X.
isInBar (ys,s) \<and> igWls MOD s X \<longrightarrow>
igFreshAbs MOD ys y (igAbs MOD ys y X)"
definition igFreshIGAbs1STR where
"igFreshIGAbs1STR MOD ==
\<forall> ys y X. igFreshAbs MOD ys y (igAbs MOD ys y X)"
lemma igFreshIGAbs1STR_imp_igFreshIGAbs1:
"igFreshIGAbs1STR MOD \<Longrightarrow> igFreshIGAbs1 MOD"
unfolding igFreshIGAbs1STR_def igFreshIGAbs1_def by simp
definition igFreshIGAbs2 where
"igFreshIGAbs2 MOD ==
\<forall> ys y xs x s X.
isInBar (xs,s) \<and> igWls MOD s X \<longrightarrow>
igFresh MOD ys y X \<longrightarrow> igFreshAbs MOD ys y (igAbs MOD xs x X)"
definition igFreshIGAbs2STR where
"igFreshIGAbs2STR MOD ==
\<forall> ys y xs x X.
igFresh MOD ys y X \<longrightarrow> igFreshAbs MOD ys y (igAbs MOD xs x X)"
lemma igFreshIGAbs2STR_imp_igFreshIGAbs2:
"igFreshIGAbs2STR MOD \<Longrightarrow> igFreshIGAbs2 MOD"
unfolding igFreshIGAbs2STR_def igFreshIGAbs2_def by simp
definition igFreshIGOp where
"igFreshIGOp MOD ==
\<forall> ys y delta inp binp.
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp \<longrightarrow>
(igFreshInp MOD ys y inp \<and> igFreshBinp MOD ys y binp) \<longrightarrow>
igFresh MOD ys y (igOp MOD delta inp binp)"
definition igFreshIGOpSTR where
"igFreshIGOpSTR MOD ==
\<forall> ys y delta inp binp.
igFreshInp MOD ys y inp \<and> igFreshBinp MOD ys y binp \<longrightarrow>
igFresh MOD ys y (igOp MOD delta inp binp)"
lemma igFreshIGOpSTR_imp_igFreshIGOp:
"igFreshIGOpSTR MOD \<Longrightarrow> igFreshIGOp MOD"
unfolding igFreshIGOpSTR_def igFreshIGOp_def by simp
definition igFreshCls where
"igFreshCls MOD ==
igFreshIGVar MOD \<and>
igFreshIGAbs1 MOD \<and> igFreshIGAbs2 MOD \<and>
igFreshIGOp MOD"
lemmas igFreshCls_defs = igFreshCls_def
igFreshIGVar_def
igFreshIGAbs1_def igFreshIGAbs2_def
igFreshIGOp_def
definition igFreshClsSTR where
"igFreshClsSTR MOD ==
igFreshIGVar MOD \<and>
igFreshIGAbs1STR MOD \<and> igFreshIGAbs2STR MOD \<and>
igFreshIGOpSTR MOD"
lemmas igFreshClsSTR_defs = igFreshClsSTR_def
igFreshIGVar_def
igFreshIGAbs1STR_def igFreshIGAbs2STR_def
igFreshIGOpSTR_def
lemma igFreshClsSTR_imp_igFreshCls:
"igFreshClsSTR MOD \<Longrightarrow> igFreshCls MOD"
unfolding igFreshClsSTR_def igFreshCls_def
using
igFreshIGAbs1STR_imp_igFreshIGAbs1 igFreshIGAbs2STR_imp_igFreshIGAbs2
igFreshIGOpSTR_imp_igFreshIGOp
by auto
(* Clauses for swap: fully-conditional versions and less-conditional,
stronger versions (suffix ``STR") *)
definition igSwapIGVar where
"igSwapIGVar MOD ==
\<forall> zs z1 z2 xs x.
igSwap MOD zs z1 z2 (igVar MOD xs x) = igVar MOD xs (x @xs[z1 \<and> z2]_zs)"
definition igSwapIGAbs where
"igSwapIGAbs MOD ==
\<forall> zs z1 z2 xs x s X.
isInBar (xs,s) \<and> igWls MOD s X \<longrightarrow>
igSwapAbs MOD zs z1 z2 (igAbs MOD xs x X) =
igAbs MOD xs (x @xs[z1 \<and> z2]_zs) (igSwap MOD zs z1 z2 X)"
definition igSwapIGAbsSTR where
"igSwapIGAbsSTR MOD ==
\<forall> zs z1 z2 xs x X.
igSwapAbs MOD zs z1 z2 (igAbs MOD xs x X) =
igAbs MOD xs (x @xs[z1 \<and> z2]_zs) (igSwap MOD zs z1 z2 X)"
lemma igSwapIGAbsSTR_imp_igSwapIGAbs:
"igSwapIGAbsSTR MOD \<Longrightarrow> igSwapIGAbs MOD"
unfolding igSwapIGAbsSTR_def igSwapIGAbs_def by simp
definition igSwapIGOp where
"igSwapIGOp MOD ==
\<forall> zs z1 z2 delta inp binp.
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp \<longrightarrow>
igSwap MOD zs z1 z2 (igOp MOD delta inp binp) =
igOp MOD delta (igSwapInp MOD zs z1 z2 inp) (igSwapBinp MOD zs z1 z2 binp)"
definition igSwapIGOpSTR where
"igSwapIGOpSTR MOD ==
\<forall> zs z1 z2 delta inp binp.
igSwap MOD zs z1 z2 (igOp MOD delta inp binp) =
igOp MOD delta (igSwapInp MOD zs z1 z2 inp) (igSwapBinp MOD zs z1 z2 binp)"
lemma igSwapIGOpSTR_imp_igSwapIGOp:
"igSwapIGOpSTR MOD \<Longrightarrow> igSwapIGOp MOD"
unfolding igSwapIGOpSTR_def igSwapIGOp_def by simp
definition igSwapCls where
"igSwapCls MOD ==
igSwapIGVar MOD \<and>
igSwapIGAbs MOD \<and>
igSwapIGOp MOD"
lemmas igSwapCls_defs = igSwapCls_def
igSwapIGVar_def
igSwapIGAbs_def
igSwapIGOp_def
definition igSwapClsSTR where
"igSwapClsSTR MOD ==
igSwapIGVar MOD \<and>
igSwapIGAbsSTR MOD \<and>
igSwapIGOpSTR MOD"
lemmas igSwapClsSTR_defs = igSwapClsSTR_def
igSwapIGVar_def
igSwapIGAbsSTR_def
igSwapIGOpSTR_def
lemma igSwapClsSTR_imp_igSwapCls:
"igSwapClsSTR MOD \<Longrightarrow> igSwapCls MOD"
unfolding igSwapClsSTR_def igSwapCls_def
using
igSwapIGAbsSTR_imp_igSwapIGAbs
igSwapIGOpSTR_imp_igSwapIGOp
by auto
(* Clauses for subst: fully-conditional versions and less-conditional,
stronger versions (suffix ``STR") *)
definition igSubstIGVar1 where
"igSubstIGVar1 MOD ==
\<forall> ys y Y xs x.
igWls MOD (asSort ys) Y \<longrightarrow>
(ys \<noteq> xs \<or> y \<noteq> x) \<longrightarrow>
igSubst MOD ys Y y (igVar MOD xs x) = igVar MOD xs x"
definition igSubstIGVar1STR where
"igSubstIGVar1STR MOD ==
(\<forall> ys y y1 xs x.
(ys \<noteq> xs \<or> x \<noteq> y) \<longrightarrow>
igSubst MOD ys (igVar MOD ys y1) y (igVar MOD xs x) = igVar MOD xs x)
\<and>
(\<forall> ys y Y xs x.
igWls MOD (asSort ys) Y \<longrightarrow>
(ys \<noteq> xs \<or> y \<noteq> x) \<longrightarrow>
igSubst MOD ys Y y (igVar MOD xs x) = igVar MOD xs x)"
lemma igSubstIGVar1STR_imp_igSubstIGVar1:
"igSubstIGVar1STR MOD \<Longrightarrow> igSubstIGVar1 MOD"
unfolding igSubstIGVar1STR_def igSubstIGVar1_def by simp
definition igSubstIGVar2 where
"igSubstIGVar2 MOD ==
\<forall> ys y Y.
igWls MOD (asSort ys) Y \<longrightarrow>
igSubst MOD ys Y y (igVar MOD ys y) = Y"
definition igSubstIGVar2STR where
"igSubstIGVar2STR MOD ==
(\<forall> ys y y1.
igSubst MOD ys (igVar MOD ys y1) y (igVar MOD ys y) = igVar MOD ys y1)
\<and>
(\<forall> ys y Y.
igWls MOD (asSort ys) Y \<longrightarrow>
igSubst MOD ys Y y (igVar MOD ys y) = Y)"
lemma igSubstIGVar2STR_imp_igSubstIGVar2:
"igSubstIGVar2STR MOD \<Longrightarrow> igSubstIGVar2 MOD"
unfolding igSubstIGVar2STR_def igSubstIGVar2_def by simp
definition igSubstIGAbs where
"igSubstIGAbs MOD ==
\<forall> ys y Y xs x s X.
isInBar (xs,s) \<and> igWls MOD (asSort ys) Y \<and> igWls MOD s X \<longrightarrow>
(xs \<noteq> ys \<or> x \<noteq> y) \<and> igFresh MOD xs x Y \<longrightarrow>
igSubstAbs MOD ys Y y (igAbs MOD xs x X) =
igAbs MOD xs x (igSubst MOD ys Y y X)"
definition igSubstIGAbsSTR where
"igSubstIGAbsSTR MOD ==
\<forall> ys y Y xs x X.
(xs \<noteq> ys \<or> x \<noteq> y) \<and> igFresh MOD xs x Y \<longrightarrow>
igSubstAbs MOD ys Y y (igAbs MOD xs x X) =
igAbs MOD xs x (igSubst MOD ys Y y X)"
lemma igSubstIGAbsSTR_imp_igSubstIGAbs:
"igSubstIGAbsSTR MOD \<Longrightarrow> igSubstIGAbs MOD"
unfolding igSubstIGAbsSTR_def igSubstIGAbs_def by simp
definition igSubstIGOp where
"igSubstIGOp MOD ==
\<forall> ys y Y delta inp binp.
igWls MOD (asSort ys) Y \<and>
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp \<longrightarrow>
igSubst MOD ys Y y (igOp MOD delta inp binp) =
igOp MOD delta (igSubstInp MOD ys Y y inp) (igSubstBinp MOD ys Y y binp)"
definition igSubstIGOpSTR where
"igSubstIGOpSTR MOD ==
(\<forall> ys y y1 delta inp binp.
igSubst MOD ys (igVar MOD ys y1) y (igOp MOD delta inp binp) =
igOp MOD delta (igSubstInp MOD ys (igVar MOD ys y1) y inp)
(igSubstBinp MOD ys (igVar MOD ys y1) y binp))
\<and>
(\<forall> ys y Y delta inp binp.
igWls MOD (asSort ys) Y \<longrightarrow>
igSubst MOD ys Y y (igOp MOD delta inp binp) =
igOp MOD delta (igSubstInp MOD ys Y y inp) (igSubstBinp MOD ys Y y binp))"
lemma igSubstIGOpSTR_imp_igSubstIGOp:
"igSubstIGOpSTR MOD \<Longrightarrow> igSubstIGOp MOD"
unfolding igSubstIGOpSTR_def igSubstIGOp_def by simp
definition igSubstCls where
"igSubstCls MOD ==
igSubstIGVar1 MOD \<and> igSubstIGVar2 MOD \<and>
igSubstIGAbs MOD \<and>
igSubstIGOp MOD"
lemmas igSubstCls_defs = igSubstCls_def
igSubstIGVar1_def igSubstIGVar2_def
igSubstIGAbs_def
igSubstIGOp_def
definition igSubstClsSTR where
"igSubstClsSTR MOD ==
igSubstIGVar1STR MOD \<and> igSubstIGVar2STR MOD \<and>
igSubstIGAbsSTR MOD \<and>
igSubstIGOpSTR MOD"
lemmas igSubstClsSTR_defs = igSubstClsSTR_def
igSubstIGVar1STR_def igSubstIGVar2STR_def
igSubstIGAbsSTR_def
igSubstIGOpSTR_def
lemma igSubstClsSTR_imp_igSubstCls:
"igSubstClsSTR MOD \<Longrightarrow> igSubstCls MOD"
unfolding igSubstClsSTR_def igSubstCls_def
using
igSubstIGVar1STR_imp_igSubstIGVar1
igSubstIGVar2STR_imp_igSubstIGVar2
igSubstIGAbsSTR_imp_igSubstIGAbs
igSubstIGOpSTR_imp_igSubstIGOp
by auto
(* Freshness-based congruences for abstractions: *)
(* ... employing swap: *)
definition igAbsCongS where
"igAbsCongS MOD ==
\<forall> xs x x' y s X X'.
isInBar (xs,s) \<and> igWls MOD s X \<and> igWls MOD s X' \<longrightarrow>
igFresh MOD xs y X \<and> igFresh MOD xs y X' \<and> igSwap MOD xs y x X = igSwap MOD xs y x' X' \<longrightarrow>
igAbs MOD xs x X = igAbs MOD xs x' X'"
definition igAbsCongSSTR where
"igAbsCongSSTR MOD ==
\<forall> xs x x' y X X'.
igFresh MOD xs y X \<and> igFresh MOD xs y X' \<and> igSwap MOD xs y x X = igSwap MOD xs y x' X' \<longrightarrow>
igAbs MOD xs x X = igAbs MOD xs x' X'"
lemma igAbsCongSSTR_imp_igAbsCongS:
"igAbsCongSSTR MOD \<Longrightarrow> igAbsCongS MOD"
unfolding igAbsCongSSTR_def igAbsCongS_def by auto
(* ... employing subst: *)
definition igAbsCongU where
"igAbsCongU MOD ==
\<forall> xs x x' y s X X'.
isInBar (xs,s) \<and> igWls MOD s X \<and> igWls MOD s X' \<longrightarrow>
igFresh MOD xs y X \<and> igFresh MOD xs y X' \<and>
igSubst MOD xs (igVar MOD xs y) x X = igSubst MOD xs (igVar MOD xs y) x' X' \<longrightarrow>
igAbs MOD xs x X = igAbs MOD xs x' X'"
definition igAbsCongUSTR where
"igAbsCongUSTR MOD ==
\<forall> xs x x' y X X'.
igFresh MOD xs y X \<and> igFresh MOD xs y X' \<and>
igSubst MOD xs (igVar MOD xs y) x X = igSubst MOD xs (igVar MOD xs y) x' X' \<longrightarrow>
igAbs MOD xs x X = igAbs MOD xs x' X'"
lemma igAbsCongUSTR_imp_igAbsCongU:
"igAbsCongUSTR MOD \<Longrightarrow> igAbsCongU MOD"
unfolding igAbsCongUSTR_def igAbsCongU_def by auto
(* (Subst-based) renaming of the bound variable in abstractions (``abstraction renaming", for short): *)
definition igAbsRen where
"igAbsRen MOD ==
\<forall> xs y x s X.
isInBar (xs,s) \<and> igWls MOD s X \<longrightarrow>
igFresh MOD xs y X \<longrightarrow>
igAbs MOD xs y (igSubst MOD xs (igVar MOD xs y) x X) = igAbs MOD xs x X"
definition igAbsRenSTR where
"igAbsRenSTR MOD ==
\<forall> xs y x X.
igFresh MOD xs y X \<longrightarrow>
igAbs MOD xs y (igSubst MOD xs (igVar MOD xs y) x X) = igAbs MOD xs x X"
lemma igAbsRenSTR_imp_igAbsRen:
"igAbsRenSTR MOD \<Longrightarrow> igAbsRen MOD"
unfolding igAbsRenSTR_def igAbsRen_def by simp
(* igAbsRenSTR is stronger than igAbsCongUSTR: *)
lemma igAbsRenSTR_imp_igAbsCongUSTR:
"igAbsRenSTR MOD \<Longrightarrow> igAbsCongUSTR MOD"
unfolding igAbsCongUSTR_def igAbsRenSTR_def by metis
-text {* Well-sorted fresh-swap models: *}
+text \<open>Well-sorted fresh-swap models:\<close>
definition iwlsFSw where
"iwlsFSw MOD ==
igWlsAllDisj MOD \<and> igWlsAbsIsInBar MOD \<and>
igConsIPresIGWls MOD \<and> igSwapAllIPresIGWlsAll MOD \<and>
igFreshCls MOD \<and> igSwapCls MOD \<and> igAbsCongS MOD"
lemmas iwlsFSw_defs1 = iwlsFSw_def
igWlsAllDisj_def igWlsAbsIsInBar_def
igConsIPresIGWls_def igSwapAllIPresIGWlsAll_def
igFreshCls_def igSwapCls_def igAbsCongS_def
lemmas iwlsFSw_defs = iwlsFSw_def
igWlsAllDisj_defs igWlsAbsIsInBar_def
igConsIPresIGWls_defs igSwapAllIPresIGWlsAll_defs
igFreshCls_defs igSwapCls_defs igAbsCongS_def
definition iwlsFSwSTR where
"iwlsFSwSTR MOD ==
igWlsAllDisj MOD \<and> igWlsAbsIsInBar MOD \<and>
igConsIPresIGWlsSTR MOD \<and> igSwapAllIPresIGWlsAllSTR MOD \<and>
igFreshClsSTR MOD \<and> igSwapClsSTR MOD \<and> igAbsCongSSTR MOD"
lemmas iwlsFSwSTR_defs1 = iwlsFSwSTR_def
igWlsAllDisj_def igWlsAbsIsInBar_def
igConsIPresIGWlsSTR_def igSwapAllIPresIGWlsAllSTR_def
igFreshClsSTR_def igSwapClsSTR_def igAbsCongSSTR_def
lemmas iwlsFSwSTR_defs = iwlsFSwSTR_def
igWlsAllDisj_defs igWlsAbsIsInBar_def
igConsIPresIGWlsSTR_defs igSwapAllIPresIGWlsAllSTR_defs
igFreshClsSTR_defs igSwapClsSTR_defs igAbsCongSSTR_def
lemma iwlsFSwSTR_imp_iwlsFSw:
"iwlsFSwSTR MOD \<Longrightarrow> iwlsFSw MOD"
unfolding iwlsFSwSTR_def iwlsFSw_def
using
igConsIPresIGWlsSTR_imp_igConsIPresIGWls
igSwapAllIPresIGWlsAllSTR_imp_igSwapAllIPresIGWlsAll
igFreshClsSTR_imp_igFreshCls
igSwapClsSTR_imp_igSwapCls
igAbsCongSSTR_imp_igAbsCongS
by auto
-text {* Well-sorted fresh-subst models: *}
+text \<open>Well-sorted fresh-subst models:\<close>
definition iwlsFSb where
"iwlsFSb MOD ==
igWlsAllDisj MOD \<and> igWlsAbsIsInBar MOD \<and>
igConsIPresIGWls MOD \<and> igSubstAllIPresIGWlsAll MOD \<and>
igFreshCls MOD \<and> igSubstCls MOD \<and> igAbsRen MOD"
lemmas iwlsFSb_defs1 = iwlsFSb_def
igWlsAllDisj_def igWlsAbsIsInBar_def
igConsIPresIGWls_def igSubstAllIPresIGWlsAll_def
igFreshCls_def igSubstCls_def igAbsRen_def
lemmas iwlsFSb_defs = iwlsFSb_def
igWlsAllDisj_defs igWlsAbsIsInBar_def
igConsIPresIGWls_defs igSubstAllIPresIGWlsAll_defs
igFreshCls_defs igSubstCls_defs igAbsRen_def
definition iwlsFSbSwTR where
"iwlsFSbSwTR MOD ==
igWlsAllDisj MOD \<and> igWlsAbsIsInBar MOD \<and>
igConsIPresIGWlsSTR MOD \<and> igSubstAllIPresIGWlsAllSTR MOD \<and>
igFreshClsSTR MOD \<and> igSubstClsSTR MOD \<and> igAbsRenSTR MOD"
lemmas wlsFSbSwSTR_defs1 = iwlsFSbSwTR_def
igWlsAllDisj_def igWlsAbsIsInBar_def
igConsIPresIGWlsSTR_def igSwapAllIPresIGWlsAllSTR_def
igFreshClsSTR_def igSwapClsSTR_def igAbsRenSTR_def
lemmas iwlsFSbSwTR_defs = iwlsFSbSwTR_def
igWlsAllDisj_defs igWlsAbsIsInBar_def
igConsIPresIGWlsSTR_defs igSwapAllIPresIGWlsAllSTR_defs
igFreshClsSTR_defs igSwapClsSTR_defs igAbsRenSTR_def
lemma iwlsFSbSwTR_imp_iwlsFSb:
"iwlsFSbSwTR MOD \<Longrightarrow> iwlsFSb MOD"
unfolding iwlsFSbSwTR_def iwlsFSb_def
using
igConsIPresIGWlsSTR_imp_igConsIPresIGWls
igSubstAllIPresIGWlsAllSTR_imp_igSubstAllIPresIGWlsAll
igFreshClsSTR_imp_igFreshCls
igSubstClsSTR_imp_igSubstCls
igAbsRenSTR_imp_igAbsRen
by auto
-text {* Well-sorted fresh-swap-subst-models *}
+text \<open>Well-sorted fresh-swap-subst-models\<close>
(* "strong" versions not required for this kind of models *)
definition iwlsFSwSb where
"iwlsFSwSb MOD ==
iwlsFSw MOD \<and> igSubstAllIPresIGWlsAll MOD \<and> igSubstCls MOD"
lemmas iwlsFSwSb_defs1 = iwlsFSwSb_def
iwlsFSw_def igSubstAllIPresIGWlsAll_def igSubstCls_def
lemmas iwlsFSwSb_defs = iwlsFSwSb_def
iwlsFSw_def igSubstAllIPresIGWlsAll_defs igSubstCls_defs
-text {* Well-sorted fresh-subst-swap-models *}
+text \<open>Well-sorted fresh-subst-swap-models\<close>
(* "strong" versions not required for this kind of models *)
definition iwlsFSbSw where
"iwlsFSbSw MOD ==
iwlsFSb MOD \<and> igSwapAllIPresIGWlsAll MOD \<and> igSwapCls MOD"
lemmas iwlsFSbSw_defs1 = iwlsFSbSw_def
iwlsFSw_def igSwapAllIPresIGWlsAll_def igSwapCls_def
lemmas iwlsFSbSw_defs = iwlsFSbSw_def
iwlsFSw_def igSwapAllIPresIGWlsAll_defs igSwapCls_defs
-text{* Extension of domain preservation (by swap and subst) to inputs: *}
-
-text {* First for free inputs: *}
+text\<open>Extension of domain preservation (by swap and subst) to inputs:\<close>
+
+text \<open>First for free inputs:\<close>
definition igSwapInpIPresIGWlsInp where
"igSwapInpIPresIGWlsInp MOD ==
\<forall> zs z1 z2 delta inp.
igWlsInp MOD delta inp \<longrightarrow>
igWlsInp MOD delta (igSwapInp MOD zs z1 z2 inp)"
definition igSwapInpIPresIGWlsInpSTR where
"igSwapInpIPresIGWlsInpSTR MOD ==
\<forall> zs z1 z2 delta inp.
igWlsInp MOD delta (igSwapInp MOD zs z1 z2 inp) =
igWlsInp MOD delta inp"
definition igSubstInpIPresIGWlsInp where
"igSubstInpIPresIGWlsInp MOD ==
\<forall> ys y Y delta inp.
igWls MOD (asSort ys) Y \<and> igWlsInp MOD delta inp \<longrightarrow>
igWlsInp MOD delta (igSubstInp MOD ys Y y inp)"
definition igSubstInpIPresIGWlsInpSTR where
"igSubstInpIPresIGWlsInpSTR MOD ==
\<forall> ys y Y delta inp.
igWls MOD (asSort ys) Y \<longrightarrow>
igWlsInp MOD delta (igSubstInp MOD ys Y y inp) =
igWlsInp MOD delta inp"
lemma imp_igSwapInpIPresIGWlsInp:
"igSwapIPresIGWls MOD \<Longrightarrow> igSwapInpIPresIGWlsInp MOD"
by (simp add:
igSwapInpIPresIGWlsInp_def igWlsInp_def liftAll2_def
igSwapIPresIGWls_def igSwapAbsIPresIGWlsAbs_def igSwapInp_def lift_def
sameDom_def split: option.splits)
lemma imp_igSwapInpIPresIGWlsInpSTR:
"igSwapIPresIGWlsSTR MOD \<Longrightarrow> igSwapInpIPresIGWlsInpSTR MOD"
by (simp add:
igSwapIPresIGWlsSTR_def igWlsInp_def liftAll2_def
igSwapIPresIGWls_def igSwapInpIPresIGWlsInpSTR_def igSwapInp_def lift_def
sameDom_def split: option.splits)
(smt option.distinct(1) option.exhaust)
lemma imp_igSubstInpIPresIGWlsInp:
"igSubstIPresIGWls MOD \<Longrightarrow> igSubstInpIPresIGWlsInp MOD"
by (simp add : igSubstInp_def
igSubstIPresIGWls_def igSubstInpIPresIGWlsInp_def igWlsInp_def liftAll2_def
lift_def sameDom_def split: option.splits)
lemma imp_igSubstInpIPresIGWlsInpSTR:
"igSubstIPresIGWlsSTR MOD \<Longrightarrow> igSubstInpIPresIGWlsInpSTR MOD"
by(simp add:
igSubstInpIPresIGWlsInpSTR_def igSubstIPresIGWlsSTR_def igSubstInp_def
igWlsInp_def liftAll2_def lift_def sameDom_def
split: option.splits) (smt option.distinct(1) option.exhaust)
-text {* Then for bound inputs: *}
+text \<open>Then for bound inputs:\<close>
definition igSwapBinpIPresIGWlsBinp where
"igSwapBinpIPresIGWlsBinp MOD ==
\<forall> zs z1 z2 delta binp.
igWlsBinp MOD delta binp \<longrightarrow>
igWlsBinp MOD delta (igSwapBinp MOD zs z1 z2 binp)"
definition igSwapBinpIPresIGWlsBinpSTR where
"igSwapBinpIPresIGWlsBinpSTR MOD ==
\<forall> zs z1 z2 delta binp.
igWlsBinp MOD delta (igSwapBinp MOD zs z1 z2 binp) =
igWlsBinp MOD delta binp"
definition igSubstBinpIPresIGWlsBinp where
"igSubstBinpIPresIGWlsBinp MOD ==
\<forall> ys y Y delta binp.
igWls MOD (asSort ys) Y \<and> igWlsBinp MOD delta binp \<longrightarrow>
igWlsBinp MOD delta (igSubstBinp MOD ys Y y binp)"
definition igSubstBinpIPresIGWlsBinpSTR where
"igSubstBinpIPresIGWlsBinpSTR MOD ==
\<forall> ys y Y delta binp.
igWls MOD (asSort ys) Y \<longrightarrow>
igWlsBinp MOD delta (igSubstBinp MOD ys Y y binp) =
igWlsBinp MOD delta binp"
lemma imp_igSwapBinpIPresIGWlsBinp:
"igSwapAbsIPresIGWlsAbs MOD \<Longrightarrow> igSwapBinpIPresIGWlsBinp MOD"
by(auto simp add:
igSwapBinpIPresIGWlsBinp_def igSwapAbsIPresIGWlsAbs_def igSwapBinp_def
igWlsBinp_def liftAll2_def lift_def sameDom_def
split: option.splits)
lemma imp_igSwapBinpIPresIGWlsBinpSTR:
"igSwapAbsIPresIGWlsAbsSTR MOD \<Longrightarrow> igSwapBinpIPresIGWlsBinpSTR MOD"
by (simp add:
igSwapBinpIPresIGWlsBinpSTR_def igSwapAbsIPresIGWlsAbsSTR_def igSwapBinp_def
igWlsBinp_def liftAll2_def lift_def sameDom_def
split: option.splits) (smt option.distinct(1) option.exhaust surj_pair)
lemma imp_igSubstBinpIPresIGWlsBinp:
"igSubstAbsIPresIGWlsAbs MOD \<Longrightarrow> igSubstBinpIPresIGWlsBinp MOD"
by (auto simp add:
igSubstBinpIPresIGWlsBinp_def igSubstAbsIPresIGWlsAbs_def igSubstBinp_def
igWlsBinp_def liftAll2_def lift_def sameDom_def
split: option.splits)
lemma imp_igSubstBinpIPresIGWlsBinpSTR:
"igSubstAbsIPresIGWlsAbsSTR MOD \<Longrightarrow> igSubstBinpIPresIGWlsBinpSTR MOD"
by (simp add:
igSubstAbsIPresIGWlsAbsSTR_def igSubstBinpIPresIGWlsBinpSTR_def igSubstBinp_def
igWlsBinp_def liftAll2_def lift_def sameDom_def
split: option.splits) (smt option.distinct(1) option.exhaust surj_pair)
-subsection {* Morphisms of models *}
-
-text{*
+subsection \<open>Morphisms of models\<close>
+
+text\<open>
The morphisms between models shall be the usual first-order-logic morphisms, i.e,, functions
commuting with the operations and preserving the (freshness) relations. Because they involve the
same signature, the morphisms for fresh-swap-subst models (called fresh-swap-subst morphisms)
will be the same as those for fresh-subst-swap-models.
-*}
-
-subsubsection {* Preservation of the domains *}
+\<close>
+
+subsubsection \<open>Preservation of the domains\<close>
definition ipresIGWls where
"ipresIGWls h MOD MOD' ==
\<forall> s X. igWls MOD s X \<longrightarrow> igWls MOD' s (h X)"
definition ipresIGWlsAbs where
"ipresIGWlsAbs hA MOD MOD' ==
\<forall> us s A. igWlsAbs MOD (us,s) A \<longrightarrow> igWlsAbs MOD' (us,s) (hA A)"
definition ipresIGWlsAll where
"ipresIGWlsAll h hA MOD MOD' ==
ipresIGWls h MOD MOD' \<and> ipresIGWlsAbs hA MOD MOD'"
lemmas ipresIGWlsAll_defs = ipresIGWlsAll_def
ipresIGWls_def ipresIGWlsAbs_def
-subsubsection {* Preservation of the constructs *}
+subsubsection \<open>Preservation of the constructs\<close>
definition ipresIGVar where
"ipresIGVar h MOD MOD' ==
\<forall> xs x. h (igVar MOD xs x) = igVar MOD' xs x"
definition ipresIGAbs where
"ipresIGAbs h hA MOD MOD' ==
\<forall> xs x s X. isInBar (xs,s) \<and> igWls MOD s X \<longrightarrow>
hA (igAbs MOD xs x X) = igAbs MOD' xs x (h X)"
definition ipresIGOp
where
"ipresIGOp h hA MOD MOD' ==
\<forall> delta inp binp.
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp \<longrightarrow>
h (igOp MOD delta inp binp) = igOp MOD' delta (lift h inp) (lift hA binp)"
definition ipresIGCons where
"ipresIGCons h hA MOD MOD' ==
ipresIGVar h MOD MOD' \<and>
ipresIGAbs h hA MOD MOD' \<and>
ipresIGOp h hA MOD MOD'"
lemmas ipresIGCons_defs = ipresIGCons_def
ipresIGVar_def
ipresIGAbs_def
ipresIGOp_def
-subsubsection {* Preservation of freshness *}
+subsubsection \<open>Preservation of freshness\<close>
definition ipresIGFresh where
"ipresIGFresh h MOD MOD' ==
\<forall> ys y s X.
igWls MOD s X \<longrightarrow>
igFresh MOD ys y X \<longrightarrow> igFresh MOD' ys y (h X)"
definition ipresIGFreshAbs where
"ipresIGFreshAbs hA MOD MOD' ==
\<forall> ys y us s A.
igWlsAbs MOD (us,s) A \<longrightarrow>
igFreshAbs MOD ys y A \<longrightarrow> igFreshAbs MOD' ys y (hA A)"
definition ipresIGFreshAll where
"ipresIGFreshAll h hA MOD MOD' ==
ipresIGFresh h MOD MOD' \<and> ipresIGFreshAbs hA MOD MOD'"
lemmas ipresIGFreshAll_defs = ipresIGFreshAll_def
ipresIGFresh_def ipresIGFreshAbs_def
-subsubsection {* Preservation of swapping *}
+subsubsection \<open>Preservation of swapping\<close>
definition ipresIGSwap where
"ipresIGSwap h MOD MOD' ==
\<forall> zs z1 z2 s X.
igWls MOD s X \<longrightarrow>
h (igSwap MOD zs z1 z2 X) = igSwap MOD' zs z1 z2 (h X)"
definition ipresIGSwapAbs where
"ipresIGSwapAbs hA MOD MOD' ==
\<forall> zs z1 z2 us s A.
igWlsAbs MOD (us,s) A \<longrightarrow>
hA (igSwapAbs MOD zs z1 z2 A) = igSwapAbs MOD' zs z1 z2 (hA A)"
definition ipresIGSwapAll where
"ipresIGSwapAll h hA MOD MOD' ==
ipresIGSwap h MOD MOD' \<and> ipresIGSwapAbs hA MOD MOD'"
lemmas ipresIGSwapAll_defs = ipresIGSwapAll_def
ipresIGSwap_def ipresIGSwapAbs_def
-subsubsection {* Preservation of subst *}
+subsubsection \<open>Preservation of subst\<close>
definition ipresIGSubst where
"ipresIGSubst h MOD MOD' ==
\<forall> ys Y y s X.
igWls MOD (asSort ys) Y \<and> igWls MOD s X \<longrightarrow>
h (igSubst MOD ys Y y X) = igSubst MOD' ys (h Y) y (h X)"
definition ipresIGSubstAbs where
"ipresIGSubstAbs h hA MOD MOD' ==
\<forall> ys Y y us s A.
igWls MOD (asSort ys) Y \<and> igWlsAbs MOD (us,s) A \<longrightarrow>
hA (igSubstAbs MOD ys Y y A) = igSubstAbs MOD' ys (h Y) y (hA A)"
definition ipresIGSubstAll where
"ipresIGSubstAll h hA MOD MOD' ==
ipresIGSubst h MOD MOD' \<and>
ipresIGSubstAbs h hA MOD MOD'"
lemmas ipresIGSubstAll_defs = ipresIGSubstAll_def
ipresIGSubst_def ipresIGSubstAbs_def
-subsubsection {* Fresh-swap morphisms *}
+subsubsection \<open>Fresh-swap morphisms\<close>
definition FSwImorph where
"FSwImorph h hA MOD MOD' ==
ipresIGWlsAll h hA MOD MOD' \<and> ipresIGCons h hA MOD MOD' \<and>
ipresIGFreshAll h hA MOD MOD' \<and> ipresIGSwapAll h hA MOD MOD'"
lemmas FSwImorph_defs1 = FSwImorph_def
ipresIGWlsAll_def ipresIGCons_def
ipresIGFreshAll_def ipresIGSwapAll_def
lemmas FSwImorph_defs = FSwImorph_def
ipresIGWlsAll_defs ipresIGCons_defs
ipresIGFreshAll_defs ipresIGSwapAll_defs
-subsubsection {* Fresh-subst morphisms *}
+subsubsection \<open>Fresh-subst morphisms\<close>
definition FSbImorph where
"FSbImorph h hA MOD MOD' ==
ipresIGWlsAll h hA MOD MOD' \<and> ipresIGCons h hA MOD MOD' \<and>
ipresIGFreshAll h hA MOD MOD' \<and> ipresIGSubstAll h hA MOD MOD'"
lemmas FSbImorph_defs1 = FSbImorph_def
ipresIGWlsAll_def ipresIGCons_def
ipresIGFreshAll_def ipresIGSubstAll_def
lemmas FSbImorph_defs = FSbImorph_def
ipresIGWlsAll_defs ipresIGCons_defs
ipresIGFreshAll_defs ipresIGSubstAll_defs
-subsubsection {* Fresh-swap-subst morphisms *}
+subsubsection \<open>Fresh-swap-subst morphisms\<close>
(* Note that FSwSb-morphisms are also igood for FSbSw-models. *)
definition FSwSbImorph where
"FSwSbImorph h hA MOD MOD' ==
FSwImorph h hA MOD MOD' \<and> ipresIGSubstAll h hA MOD MOD'"
lemmas FSwSbImorph_defs1 = FSwSbImorph_def
FSwImorph_def ipresIGSubstAll_def
lemmas FSwSbImorph_defs = FSwSbImorph_def
FSwImorph_defs ipresIGSubstAll_defs
-subsubsection {* Basic facts *}
-
-text {* FSwSb morphisms are the same as FSbSw morphisms: *}
+subsubsection \<open>Basic facts\<close>
+
+text \<open>FSwSb morphisms are the same as FSbSw morphisms:\<close>
lemma FSwSbImorph_iff:
"FSwSbImorph h hA MOD MOD' =
(FSbImorph h hA MOD MOD' \<and> ipresIGSwapAll h hA MOD MOD')"
unfolding FSwSbImorph_def FSbImorph_def FSwImorph_def by auto
-text {* Some facts for free inpus: *}
+text \<open>Some facts for free inpus:\<close>
lemma igSwapInp_None[simp]:
"(igSwapInp MOD zs z1 z2 inp i = None) = (inp i = None)"
unfolding igSwapInp_def by(simp add: lift_None)
lemma igSubstInp_None[simp]:
"(igSubstInp MOD ys Y y inp i = None) = (inp i = None)"
unfolding igSubstInp_def by(simp add: lift_None)
lemma imp_igWlsInp:
"igWlsInp MOD delta inp \<Longrightarrow> ipresIGWls h MOD MOD'
\<Longrightarrow> igWlsInp MOD' delta (lift h inp)"
by (simp add: igWlsInp_def ipresIGWls_def liftAll2_def lift_def
sameDom_def split: option.splits)
corollary FSwImorph_igWlsInp:
assumes "igWlsInp MOD delta inp" and "FSwImorph h hA MOD MOD'"
shows "igWlsInp MOD' delta (lift h inp)"
using assms unfolding FSwImorph_def ipresIGWlsAll_def
using imp_igWlsInp by auto
corollary FSbImorph_igWlsInp:
assumes "igWlsInp MOD delta inp" and "FSbImorph h hA MOD MOD'"
shows "igWlsInp MOD' delta (lift h inp)"
using assms unfolding FSbImorph_def ipresIGWlsAll_def
using imp_igWlsInp by auto
lemma FSwSbImorph_igWlsInp:
assumes "igWlsInp MOD delta inp" and "FSwSbImorph h hA MOD MOD'"
shows "igWlsInp MOD' delta (lift h inp)"
using assms unfolding FSwSbImorph_def using FSwImorph_igWlsInp by auto
-text {* Similar facts for bound inpus: *}
+text \<open>Similar facts for bound inpus:\<close>
lemma igSwapBinp_None[simp]:
"(igSwapBinp MOD zs z1 z2 binp i = None) = (binp i = None)"
unfolding igSwapBinp_def by(simp add: lift_None)
lemma igSubstBinp_None[simp]:
"(igSubstBinp MOD ys Y y binp i = None) = (binp i = None)"
unfolding igSubstBinp_def by(simp add: lift_None)
lemma imp_igWlsBinp:
assumes *: "igWlsBinp MOD delta binp"
and **: "ipresIGWlsAbs hA MOD MOD'"
shows "igWlsBinp MOD' delta (lift hA binp)"
using assms by (simp add: igWlsBinp_def ipresIGWlsAbs_def liftAll2_def lift_def
sameDom_def split: option.splits)
corollary FSwImorph_igWlsBinp:
assumes "igWlsBinp MOD delta binp" and "FSwImorph h hA MOD MOD'"
shows "igWlsBinp MOD' delta (lift hA binp)"
using assms unfolding FSwImorph_def ipresIGWlsAll_def
using imp_igWlsBinp by auto
corollary FSbImorph_igWlsBinp:
assumes "igWlsBinp MOD delta binp" and "FSbImorph h hA MOD MOD'"
shows "igWlsBinp MOD' delta (lift hA binp)"
using assms unfolding FSbImorph_def ipresIGWlsAll_def
using imp_igWlsBinp by auto
lemma FSwSbImorph_igWlsBinp:
assumes "igWlsBinp MOD delta binp" and "FSwSbImorph h hA MOD MOD'"
shows "igWlsBinp MOD' delta (lift hA binp)"
using assms unfolding FSwSbImorph_def using FSwImorph_igWlsBinp by auto
lemmas input_igSwap_igSubst_None =
igSwapInp_None igSubstInp_None
igSwapBinp_None igSubstBinp_None
-subsubsection {* Identity and composition *}
+subsubsection \<open>Identity and composition\<close>
lemma id_FSwImorph: "FSwImorph id id MOD MOD"
unfolding FSwImorph_defs by auto
lemma id_FSbImorph: "FSbImorph id id MOD MOD"
unfolding FSbImorph_defs by auto
lemma id_FSwSbImorph: "FSwSbImorph id id MOD MOD"
unfolding FSwSbImorph_def apply(auto simp add: id_FSwImorph)
unfolding ipresIGSubstAll_defs by auto
lemma comp_ipresIGWls:
assumes "ipresIGWls h MOD MOD'" and "ipresIGWls h' MOD' MOD''"
shows "ipresIGWls (h' o h) MOD MOD''"
using assms unfolding ipresIGWls_def by auto
lemma comp_ipresIGWlsAbs:
assumes "ipresIGWlsAbs hA MOD MOD'" and "ipresIGWlsAbs hA' MOD' MOD''"
shows "ipresIGWlsAbs (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAbs_def by auto
lemma comp_ipresIGWlsAll:
assumes "ipresIGWlsAll h hA MOD MOD'" and "ipresIGWlsAll h' hA' MOD' MOD''"
shows "ipresIGWlsAll (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAll_def
using comp_ipresIGWls comp_ipresIGWlsAbs by auto
lemma comp_ipresIGVar:
assumes "ipresIGVar h MOD MOD'" and "ipresIGVar h' MOD' MOD''"
shows "ipresIGVar (h' o h) MOD MOD''"
using assms unfolding ipresIGVar_def by auto
lemma comp_ipresIGAbs:
assumes "ipresIGWls h MOD MOD'"
and "ipresIGAbs h hA MOD MOD'" and "ipresIGAbs h' hA' MOD' MOD''"
shows "ipresIGAbs (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWls_def ipresIGAbs_def by fastforce
lemma comp_ipresIGOp:
assumes ipres: "ipresIGWls h MOD MOD'" and ipresAbs: "ipresIGWlsAbs hA MOD MOD'"
and h: "ipresIGOp h hA MOD MOD'" and h': "ipresIGOp h' hA' MOD' MOD''"
shows "ipresIGOp (h' o h) (hA' o hA) MOD MOD''"
using assms by (auto simp: imp_igWlsInp imp_igWlsBinp ipresIGOp_def lift_comp)
lemma comp_ipresIGCons:
assumes "ipresIGWlsAll h hA MOD MOD'"
and "ipresIGCons h hA MOD MOD'" and "ipresIGCons h' hA' MOD' MOD''"
shows "ipresIGCons (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAll_def ipresIGCons_def
using comp_ipresIGVar comp_ipresIGAbs comp_ipresIGOp by auto
lemma comp_ipresIGFresh:
assumes "ipresIGWls h MOD MOD'"
and "ipresIGFresh h MOD MOD'" and "ipresIGFresh h' MOD' MOD''"
shows "ipresIGFresh (h' o h) MOD MOD''"
using assms unfolding ipresIGWls_def ipresIGFresh_def by fastforce
lemma comp_ipresIGFreshAbs:
assumes "ipresIGWlsAbs hA MOD MOD'"
and "ipresIGFreshAbs hA MOD MOD'" and "ipresIGFreshAbs hA' MOD' MOD''"
shows "ipresIGFreshAbs (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAbs_def ipresIGFreshAbs_def by fastforce
lemma comp_ipresIGFreshAll:
assumes "ipresIGWlsAll h hA MOD MOD'"
and "ipresIGFreshAll h hA MOD MOD'" and "ipresIGFreshAll h' hA' MOD' MOD''"
shows "ipresIGFreshAll (h' o h) (hA' o hA) MOD MOD''"
using assms
unfolding ipresIGWlsAll_def ipresIGFreshAll_def
using comp_ipresIGFresh comp_ipresIGFreshAbs by auto
lemma comp_ipresIGSwap:
assumes "ipresIGWls h MOD MOD'"
and "ipresIGSwap h MOD MOD'" and "ipresIGSwap h' MOD' MOD''"
shows "ipresIGSwap (h' o h) MOD MOD''"
using assms unfolding ipresIGWls_def ipresIGSwap_def by fastforce
lemma comp_ipresIGSwapAbs:
assumes "ipresIGWlsAbs hA MOD MOD'"
and "ipresIGSwapAbs hA MOD MOD'" and "ipresIGSwapAbs hA' MOD' MOD''"
shows "ipresIGSwapAbs (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAbs_def ipresIGSwapAbs_def by fastforce
lemma comp_ipresIGSwapAll:
assumes "ipresIGWlsAll h hA MOD MOD'"
and "ipresIGSwapAll h hA MOD MOD'" and "ipresIGSwapAll h' hA' MOD' MOD''"
shows "ipresIGSwapAll (h' o h) (hA' o hA) MOD MOD''"
using assms
unfolding ipresIGWlsAll_def ipresIGSwapAll_def
using comp_ipresIGSwap comp_ipresIGSwapAbs by auto
lemma comp_ipresIGSubst:
assumes "ipresIGWls h MOD MOD'"
and "ipresIGSubst h MOD MOD'" and "ipresIGSubst h' MOD' MOD''"
shows "ipresIGSubst (h' o h) MOD MOD''"
using assms unfolding ipresIGWls_def ipresIGSubst_def
apply auto by blast
lemma comp_ipresIGSubstAbs:
assumes *: "igWlsAbsIsInBar MOD"
and h: "ipresIGWls h MOD MOD'" and hA: "ipresIGWlsAbs hA MOD MOD'"
and hhA: "ipresIGSubstAbs h hA MOD MOD'" and h'hA': "ipresIGSubstAbs h' hA' MOD' MOD''"
shows "ipresIGSubstAbs (h' o h) (hA' o hA) MOD MOD''"
using assms by(fastforce simp: igWlsAbsIsInBar_def
ipresIGSubstAbs_def ipresIGWls_def ipresIGWlsAbs_def)
lemma comp_ipresIGSubstAll:
assumes "igWlsAbsIsInBar MOD"
and "ipresIGWlsAll h hA MOD MOD'"
and "ipresIGSubstAll h hA MOD MOD'" and "ipresIGSubstAll h' hA' MOD' MOD''"
shows "ipresIGSubstAll (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding ipresIGWlsAll_def ipresIGSubstAll_def
using comp_ipresIGSubst comp_ipresIGSubstAbs by auto
lemma comp_FSwImorph:
assumes *: "FSwImorph h hA MOD MOD'" and **: "FSwImorph h' hA' MOD' MOD''"
shows "FSwImorph (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding FSwImorph_def
using comp_ipresIGWlsAll comp_ipresIGCons
comp_ipresIGFreshAll comp_ipresIGSwapAll by auto
lemma comp_FSbImorph:
assumes "igWlsAbsIsInBar MOD"
and "FSbImorph h hA MOD MOD'" and "FSbImorph h' hA' MOD' MOD''"
shows "FSbImorph (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding FSbImorph_def
using comp_ipresIGWlsAll comp_ipresIGCons
comp_ipresIGFreshAll comp_ipresIGSubstAll by auto
lemma comp_FSwSbImorph:
assumes "igWlsAbsIsInBar MOD"
and "FSwSbImorph h hA MOD MOD'" and "FSwSbImorph h' hA' MOD' MOD''"
shows "FSwSbImorph (h' o h) (hA' o hA) MOD MOD''"
using assms unfolding FSwSbImorph_def
using comp_FSwImorph FSwImorph_def comp_ipresIGSubstAll FixSyn_axioms by blast
-subsection {* The term model *}
-
-text {* We show that terms form fresh-swap-subst and fresh-subst-swap models. *}
-
-subsubsection {* Definitions and simplification rules *}
+subsection \<open>The term model\<close>
+
+text \<open>We show that terms form fresh-swap-subst and fresh-subst-swap models.\<close>
+
+subsubsection \<open>Definitions and simplification rules\<close>
definition termMOD where
"termMOD ==
\<lparr>igWls = wls, igWlsAbs = wlsAbs,
igVar = Var, igAbs = Abs, igOp = Op,
igFresh = fresh, igFreshAbs = freshAbs,
igSwap = swap, igSwapAbs = swapAbs,
igSubst = subst, igSubstAbs = substAbs\<rparr>"
lemma igWls_termMOD[simp]: "igWls termMOD = wls"
unfolding termMOD_def by simp
lemma igWlsAbs_termMOD[simp]: "igWlsAbs termMOD = wlsAbs"
unfolding termMOD_def by simp
lemma igWlsInp_termMOD_wlsInp[simp]:
"igWlsInp termMOD delta inp = wlsInp delta inp"
unfolding igWlsInp_def wlsInp_iff by simp
lemma igWlsBinp_termMOD_wlsBinp[simp]:
"igWlsBinp termMOD delta binp = wlsBinp delta binp"
unfolding igWlsBinp_def wlsBinp_iff by simp
lemmas igWlsAll_termMOD_simps =
igWls_termMOD igWlsAbs_termMOD
igWlsInp_termMOD_wlsInp igWlsBinp_termMOD_wlsBinp
lemma igVar_termMOD[simp]: "igVar termMOD = Var"
unfolding termMOD_def by simp
lemma igAbs_termMOD[simp]: "igAbs termMOD = Abs"
unfolding termMOD_def by simp
lemma igOp_termMOD[simp]: "igOp termMOD = Op"
unfolding termMOD_def by simp
lemmas igCons_termMOD_simps =
igVar_termMOD igAbs_termMOD igOp_termMOD
lemma igFresh_termMOD[simp]: "igFresh termMOD = fresh"
unfolding termMOD_def by simp
lemma igFreshAbs_termMOD[simp]: "igFreshAbs termMOD = freshAbs"
unfolding termMOD_def by simp
lemma igFreshInp_termMOD[simp]: "igFreshInp termMOD = freshInp"
unfolding igFreshInp_def[abs_def] freshInp_def[abs_def] by simp
lemma igFreshBinp_termMOD[simp]: "igFreshBinp termMOD = freshBinp"
unfolding igFreshBinp_def[abs_def] freshBinp_def[abs_def] by simp
lemmas igFreshAll_termMOD_simps =
igFresh_termMOD igFreshAbs_termMOD
igFreshInp_termMOD igFreshBinp_termMOD
lemma igSwap_termMOD[simp]: "igSwap termMOD = swap"
unfolding termMOD_def by simp
lemma igSwapAbs_termMOD[simp]: "igSwapAbs termMOD = swapAbs"
unfolding termMOD_def by simp
lemma igSwapInp_termMOD[simp]: "igSwapInp termMOD = swapInp"
unfolding igSwapInp_def[abs_def] swapInp_def[abs_def] by simp
lemma igSwapBinp_termMOD[simp]: "igSwapBinp termMOD = swapBinp"
unfolding igSwapBinp_def[abs_def] swapBinp_def[abs_def] by simp
lemmas igSwapAll_termMOD_simps =
igSwap_termMOD igSwapAbs_termMOD
igSwapInp_termMOD igSwapBinp_termMOD
lemma igSubst_termMOD[simp]: "igSubst termMOD = subst"
unfolding termMOD_def by simp
lemma igSubstAbs_termMOD[simp]: "igSubstAbs termMOD = substAbs"
unfolding termMOD_def by simp
lemma igSubstInp_termMOD[simp]: "igSubstInp termMOD = substInp"
by (simp add: igSubstInp_def[abs_def] substInp_def[abs_def]
psubstInp_def[abs_def] subst_def)
lemma igSubstBinp_termMOD[simp]: "igSubstBinp termMOD = substBinp"
by (simp add: igSubstBinp_def[abs_def] substBinp_def[abs_def]
psubstBinp_def[abs_def] substAbs_def)
lemmas igSubstAll_termMOD_simps =
igSubst_termMOD igSubstAbs_termMOD
igSubstInp_termMOD igSubstBinp_termMOD
lemmas structure_termMOD_simps =
igWlsAll_termMOD_simps
igFreshAll_termMOD_simps
igSwapAll_termMOD_simps
igSubstAll_termMOD_simps
-subsubsection {* Well-sortedness of the term model *}
-
-text{* Domains are disjoint: *}
+subsubsection \<open>Well-sortedness of the term model\<close>
+
+text\<open>Domains are disjoint:\<close>
lemma termMOD_igWlsDisj: "igWlsDisj termMOD"
unfolding igWlsDisj_def using wls_disjoint by auto
lemma termMOD_igWlsAbsDisj: "igWlsAbsDisj termMOD"
unfolding igWlsAbsDisj_def using wlsAbs_disjoint by auto
lemma termMOD_igWlsAllDisj: "igWlsAllDisj termMOD"
unfolding igWlsAllDisj_def
using termMOD_igWlsDisj termMOD_igWlsAbsDisj by simp
-text {* Abstraction domains inhabited only within bound arities: *}
+text \<open>Abstraction domains inhabited only within bound arities:\<close>
lemma termMOD_igWlsAbsIsInBar: "igWlsAbsIsInBar termMOD"
unfolding igWlsAbsIsInBar_def using wlsAbs_nchotomy by simp
-text{* The syntactic constructs preserve the domains: *}
+text\<open>The syntactic constructs preserve the domains:\<close>
lemma termMOD_igVarIPresIGWls: "igVarIPresIGWls termMOD"
unfolding igVarIPresIGWls_def by simp
lemma termMOD_igAbsIPresIGWls: "igAbsIPresIGWls termMOD"
unfolding igAbsIPresIGWls_def by simp
lemma termMOD_igOpIPresIGWls: "igOpIPresIGWls termMOD"
unfolding igOpIPresIGWls_def by simp
lemma termMOD_igConsIPresIGWls: "igConsIPresIGWls termMOD"
unfolding igConsIPresIGWls_def
using termMOD_igVarIPresIGWls termMOD_igAbsIPresIGWls termMOD_igOpIPresIGWls
by auto
-text{* Swap preserves the domains: *}
+text\<open>Swap preserves the domains:\<close>
lemma termMOD_igSwapIPresIGWls: "igSwapIPresIGWls termMOD"
unfolding igSwapIPresIGWls_def by simp
lemma termMOD_igSwapAbsIPresIGWlsAbs: "igSwapAbsIPresIGWlsAbs termMOD"
unfolding igSwapAbsIPresIGWlsAbs_def by simp
lemma termMOD_igSwapAllIPresIGWlsAll: "igSwapAllIPresIGWlsAll termMOD"
unfolding igSwapAllIPresIGWlsAll_def
using termMOD_igSwapIPresIGWls termMOD_igSwapAbsIPresIGWlsAbs by auto
-text{* ``Subst" preserves the domains: *}
+text\<open>``Subst" preserves the domains:\<close>
lemma termMOD_igSubstIPresIGWls: "igSubstIPresIGWls termMOD"
unfolding igSubstIPresIGWls_def by simp
lemma termMOD_igSubstAbsIPresIGWlsAbs: "igSubstAbsIPresIGWlsAbs termMOD"
unfolding igSubstAbsIPresIGWlsAbs_def by simp
lemma termMOD_igSubstAllIPresIGWlsAll: "igSubstAllIPresIGWlsAll termMOD"
unfolding igSubstAllIPresIGWlsAll_def
using termMOD_igSubstIPresIGWls termMOD_igSubstAbsIPresIGWlsAbs by auto
-text{* The ``fresh" clauses hold: *}
+text\<open>The ``fresh" clauses hold:\<close>
lemma termMOD_igFreshIGVar: "igFreshIGVar termMOD"
unfolding igFreshIGVar_def by simp
lemma termMOD_igFreshIGAbs1: "igFreshIGAbs1 termMOD"
unfolding igFreshIGAbs1_def by auto
lemma termMOD_igFreshIGAbs2: "igFreshIGAbs2 termMOD"
unfolding igFreshIGAbs2_def by auto
lemma termMOD_igFreshIGOp: "igFreshIGOp termMOD"
unfolding igFreshIGOp_def by simp
lemma termMOD_igFreshCls: "igFreshCls termMOD"
unfolding igFreshCls_def
using termMOD_igFreshIGVar termMOD_igFreshIGAbs1 termMOD_igFreshIGAbs2 termMOD_igFreshIGOp
by simp
-text{* The ``swap" clauses hold: *}
+text\<open>The ``swap" clauses hold:\<close>
lemma termMOD_igSwapIGVar: "igSwapIGVar termMOD"
unfolding igSwapIGVar_def by simp
lemma termMOD_igSwapIGAbs: "igSwapIGAbs termMOD"
unfolding igSwapIGAbs_def by auto
lemma termMOD_igSwapIGOp: "igSwapIGOp termMOD"
unfolding igSwapIGOp_def by simp
lemma termMOD_igSwapCls: "igSwapCls termMOD"
unfolding igSwapCls_def
using termMOD_igSwapIGVar termMOD_igSwapIGAbs termMOD_igSwapIGOp by simp
-text{* The ``subst" clauses hold: *}
+text\<open>The ``subst" clauses hold:\<close>
lemma termMOD_igSubstIGVar1: "igSubstIGVar1 termMOD"
unfolding igSubstIGVar1_def by auto
lemma termMOD_igSubstIGVar2: "igSubstIGVar2 termMOD"
unfolding igSubstIGVar2_def by auto
lemma termMOD_igSubstIGAbs: "igSubstIGAbs termMOD"
unfolding igSubstIGAbs_def by auto
lemma termMOD_igSubstIGOp: "igSubstIGOp termMOD"
unfolding igSubstIGOp_def by simp
lemma termMOD_igSubstCls: "igSubstCls termMOD"
unfolding igSubstCls_def
using termMOD_igSubstIGVar1 termMOD_igSubstIGVar2
termMOD_igSubstIGAbs termMOD_igSubstIGOp by simp
-text{* The swap-congruence clause for abstractions holds: *}
+text\<open>The swap-congruence clause for abstractions holds:\<close>
lemma termMOD_igAbsCongS: "igAbsCongS termMOD"
unfolding igAbsCongS_def using wls_Abs_swap_cong
by (metis igAbs_termMOD igFresh_termMOD igSwap_termMOD igWls_termMOD)
-text{* The subst-renaming clause for abstractions holds: *}
+text\<open>The subst-renaming clause for abstractions holds:\<close>
lemma termMOD_igAbsRen: "igAbsRen termMOD"
unfolding igAbsRen_def by auto
lemma termMOD_iwlsFSw: "iwlsFSw termMOD"
unfolding iwlsFSw_def
using
termMOD_igWlsAllDisj termMOD_igWlsAbsIsInBar
termMOD_igConsIPresIGWls termMOD_igSwapAllIPresIGWlsAll
termMOD_igFreshCls termMOD_igSwapCls termMOD_igAbsCongS
by auto
lemma termMOD_iwlsFSb: "iwlsFSb termMOD"
unfolding iwlsFSb_def
using
termMOD_igWlsAllDisj termMOD_igWlsAbsIsInBar
termMOD_igConsIPresIGWls termMOD_igSubstAllIPresIGWlsAll
termMOD_igFreshCls termMOD_igSubstCls termMOD_igAbsRen
by auto
lemma termMOD_iwlsFSwSb: "iwlsFSwSb termMOD"
unfolding iwlsFSwSb_def
using termMOD_iwlsFSw termMOD_igSubstAllIPresIGWlsAll termMOD_igSubstCls
by simp
lemma termMOD_iwlsFSbSw: "iwlsFSbSw termMOD"
unfolding iwlsFSbSw_def
using termMOD_iwlsFSb termMOD_igSwapAllIPresIGWlsAll termMOD_igSwapCls
by simp
-subsubsection {* Direct description of morphisms from the term models *}
+subsubsection \<open>Direct description of morphisms from the term models\<close>
(* We merely employ predicates referring directly to terms rather than
mediating through the model structure of terms. *)
definition ipresWls where
"ipresWls h MOD ==
\<forall> s X. wls s X \<longrightarrow> igWls MOD s (h X)"
lemma ipresIGWls_termMOD[simp]:
"ipresIGWls h termMOD MOD = ipresWls h MOD"
unfolding ipresIGWls_def ipresWls_def by simp
definition ipresWlsAbs where
"ipresWlsAbs hA MOD ==
\<forall> us s A. wlsAbs (us,s) A \<longrightarrow> igWlsAbs MOD (us,s) (hA A)"
lemma ipresIGWlsAbs_termMOD[simp]:
"ipresIGWlsAbs hA termMOD MOD = ipresWlsAbs hA MOD"
unfolding ipresIGWlsAbs_def ipresWlsAbs_def by simp
definition ipresWlsAll where
"ipresWlsAll h hA MOD ==
ipresWls h MOD \<and> ipresWlsAbs hA MOD"
lemmas ipresWlsAll_defs = ipresWlsAll_def
ipresWls_def ipresWlsAbs_def
lemma ipresIGWlsAll_termMOD[simp]:
"ipresIGWlsAll h hA termMOD MOD = ipresWlsAll h hA MOD"
unfolding ipresIGWlsAll_def ipresWlsAll_def by simp
lemmas ipresIGWlsAll_termMOD_simps =
ipresIGWls_termMOD ipresIGWlsAbs_termMOD ipresIGWlsAll_termMOD
definition ipresVar where
"ipresVar h MOD ==
\<forall> xs x. h (Var xs x) = igVar MOD xs x"
lemma ipresIGVar_termMOD[simp]:
"ipresIGVar h termMOD MOD = ipresVar h MOD"
unfolding ipresIGVar_def ipresVar_def by simp
definition ipresAbs where
"ipresAbs h hA MOD ==
\<forall> xs x s X. isInBar (xs,s) \<and> wls s X \<longrightarrow> hA (Abs xs x X) = igAbs MOD xs x (h X)"
lemma ipresIGAbs_termMOD[simp]:
"ipresIGAbs h hA termMOD MOD = ipresAbs h hA MOD"
unfolding ipresIGAbs_def ipresAbs_def by simp
definition ipresOp where
"ipresOp h hA MOD ==
\<forall> delta inp binp.
wlsInp delta inp \<and> wlsBinp delta binp \<longrightarrow>
h (Op delta inp binp) =
igOp MOD delta (lift h inp) (lift hA binp)"
lemma ipresIGOp_termMOD[simp]:
"ipresIGOp h hA termMOD MOD = ipresOp h hA MOD"
unfolding ipresIGOp_def ipresOp_def by simp
definition ipresCons where
"ipresCons h hA MOD ==
ipresVar h MOD \<and>
ipresAbs h hA MOD \<and>
ipresOp h hA MOD"
lemmas ipresCons_defs = ipresCons_def
ipresVar_def
ipresAbs_def
ipresOp_def
lemma ipresIGCons_termMOD[simp]:
"ipresIGCons h hA termMOD MOD = ipresCons h hA MOD"
unfolding ipresIGCons_def ipresCons_def by simp
lemmas ipresIGCons_termMOD_simps =
ipresIGVar_termMOD ipresIGAbs_termMOD ipresIGOp_termMOD
ipresIGCons_termMOD
definition ipresFresh where
"ipresFresh h MOD ==
\<forall> ys y s X.
wls s X \<longrightarrow>
fresh ys y X \<longrightarrow> igFresh MOD ys y (h X)"
lemma ipresIGFresh_termMOD[simp]:
"ipresIGFresh h termMOD MOD = ipresFresh h MOD"
unfolding ipresIGFresh_def ipresFresh_def by simp
definition ipresFreshAbs where
"ipresFreshAbs hA MOD ==
\<forall> ys y us s A.
wlsAbs (us,s) A \<longrightarrow>
freshAbs ys y A \<longrightarrow> igFreshAbs MOD ys y (hA A)"
lemma ipresIGFreshAbs_termMOD[simp]:
"ipresIGFreshAbs hA termMOD MOD = ipresFreshAbs hA MOD"
unfolding ipresIGFreshAbs_def ipresFreshAbs_def by simp
definition ipresFreshAll where
"ipresFreshAll h hA MOD ==
ipresFresh h MOD \<and> ipresFreshAbs hA MOD"
lemmas ipresFreshAll_defs = ipresFreshAll_def
ipresFresh_def ipresFreshAbs_def
lemma ipresIGFreshAll_termMOD[simp]:
"ipresIGFreshAll h hA termMOD MOD = ipresFreshAll h hA MOD"
unfolding ipresIGFreshAll_def ipresFreshAll_def by simp
lemmas ipresIGFreshAll_termMOD_simps =
ipresIGFresh_termMOD ipresIGFreshAbs_termMOD ipresIGFreshAll_termMOD
definition ipresSwap where
"ipresSwap h MOD ==
\<forall> zs z1 z2 s X.
wls s X \<longrightarrow>
h (X #[z1 \<and> z2]_zs) = igSwap MOD zs z1 z2 (h X)"
lemma ipresIGSwap_termMOD[simp]:
"ipresIGSwap h termMOD MOD = ipresSwap h MOD"
unfolding ipresIGSwap_def ipresSwap_def by simp
definition ipresSwapAbs where
"ipresSwapAbs hA MOD ==
\<forall> zs z1 z2 us s A.
wlsAbs (us,s) A \<longrightarrow>
hA (A $[z1 \<and> z2]_zs) = igSwapAbs MOD zs z1 z2 (hA A)"
lemma ipresIGSwapAbs_termMOD[simp]:
"ipresIGSwapAbs hA termMOD MOD = ipresSwapAbs hA MOD"
unfolding ipresIGSwapAbs_def ipresSwapAbs_def by simp
definition ipresSwapAll where
"ipresSwapAll h hA MOD ==
ipresSwap h MOD \<and> ipresSwapAbs hA MOD"
lemmas ipresSwapAll_defs = ipresSwapAll_def
ipresSwap_def ipresSwapAbs_def
lemma ipresIGSwapAll_termMOD[simp]:
"ipresIGSwapAll h hA termMOD MOD = ipresSwapAll h hA MOD"
unfolding ipresIGSwapAll_def ipresSwapAll_def by simp
lemmas ipresIGSwapAll_termMOD_simps =
ipresIGSwap_termMOD ipresIGSwapAbs_termMOD ipresIGSwapAll_termMOD
definition ipresSubst where
"ipresSubst h MOD ==
\<forall> ys Y y s X.
wls (asSort ys) Y \<and> wls s X \<longrightarrow>
h (subst ys Y y X) = igSubst MOD ys (h Y) y (h X)"
lemma ipresIGSubst_termMOD[simp]:
"ipresIGSubst h termMOD MOD = ipresSubst h MOD"
unfolding ipresIGSubst_def ipresSubst_def by simp
definition ipresSubstAbs where
"ipresSubstAbs h hA MOD ==
\<forall> ys Y y us s A.
wls (asSort ys) Y \<and> wlsAbs (us,s) A \<longrightarrow>
hA (A $[Y / y]_ys) = igSubstAbs MOD ys (h Y) y (hA A)"
lemma ipresIGSubstAbs_termMOD[simp]:
"ipresIGSubstAbs h hA termMOD MOD = ipresSubstAbs h hA MOD"
unfolding ipresIGSubstAbs_def ipresSubstAbs_def by simp
definition ipresSubstAll where
"ipresSubstAll h hA MOD ==
ipresSubst h MOD \<and> ipresSubstAbs h hA MOD"
lemmas ipresSubstAll_defs = ipresSubstAll_def
ipresSubst_def ipresSubstAbs_def
lemma ipresIGSubstAll_termMOD[simp]:
"ipresIGSubstAll h hA termMOD MOD = ipresSubstAll h hA MOD"
unfolding ipresIGSubstAll_def ipresSubstAll_def by simp
lemmas ipresIGSubstAll_termMOD_simps =
ipresIGSubst_termMOD ipresIGSubstAbs_termMOD ipresIGSubstAll_termMOD
definition termFSwImorph where
"termFSwImorph h hA MOD ==
ipresWlsAll h hA MOD \<and> ipresCons h hA MOD \<and>
ipresFreshAll h hA MOD \<and> ipresSwapAll h hA MOD"
lemmas termFSwImorph_defs1 = termFSwImorph_def
ipresWlsAll_def ipresCons_def
ipresFreshAll_def ipresSwapAll_def
lemmas termFSwImorph_defs = termFSwImorph_def
ipresWlsAll_defs ipresCons_defs
ipresFreshAll_defs ipresSwapAll_defs
lemma FSwImorph_termMOD[simp]:
"FSwImorph h hA termMOD MOD = termFSwImorph h hA MOD"
unfolding FSwImorph_def termFSwImorph_def by simp
definition termFSbImorph where
"termFSbImorph h hA MOD ==
ipresWlsAll h hA MOD \<and> ipresCons h hA MOD \<and>
ipresFreshAll h hA MOD \<and> ipresSubstAll h hA MOD"
lemmas termFSbImorph_defs1 = termFSbImorph_def
ipresWlsAll_def ipresCons_def
ipresFreshAll_def ipresSubstAll_def
lemmas termFSbImorph_defs = termFSbImorph_def
ipresWlsAll_defs ipresCons_defs
ipresFreshAll_defs ipresSubstAll_defs
lemma FSbImorph_termMOD[simp]:
"FSbImorph h hA termMOD MOD = termFSbImorph h hA MOD"
unfolding FSbImorph_def termFSbImorph_def by simp
definition termFSwSbImorph where
"termFSwSbImorph h hA MOD ==
termFSwImorph h hA MOD \<and> ipresSubstAll h hA MOD"
lemmas termFSwSbImorph_defs1 = termFSwSbImorph_def
termFSwImorph_def ipresSubstAll_def
lemmas termFSwSbImorph_defs = termFSwSbImorph_def
termFSwImorph_defs ipresSubstAll_defs
-text {* Term FSwSb morphisms are the same as FSbSw morphisms: *}
+text \<open>Term FSwSb morphisms are the same as FSbSw morphisms:\<close>
lemma termFSwSbImorph_iff:
"termFSwSbImorph h hA MOD =
(termFSbImorph h hA MOD \<and> ipresSwapAll h hA MOD)"
unfolding termFSwSbImorph_def termFSwImorph_def termFSbImorph_def ipresSubstAll_def
unfolding FSwSbImorph_def FSbImorph_def FSwImorph_def by auto
lemma FSwSbImorph_termMOD[simp]:
"FSwSbImorph h hA termMOD MOD = termFSwSbImorph h hA MOD"
unfolding FSwSbImorph_def termFSwSbImorph_def by simp
lemma ipresWls_wlsInp:
assumes "wlsInp delta inp" and "ipresWls h MOD"
shows "igWlsInp MOD delta (lift h inp)"
using assms imp_igWlsInp[of termMOD delta inp h MOD] by auto
lemma termFSwImorph_wlsInp:
assumes "wlsInp delta inp" and "termFSwImorph h hA MOD"
shows "igWlsInp MOD delta (lift h inp)"
using assms FSwImorph_igWlsInp[of termMOD delta inp h hA MOD] by auto
lemma termFSwSbImorph_wlsInp:
assumes "wlsInp delta inp" and "termFSwSbImorph h hA MOD"
shows "igWlsInp MOD delta (lift h inp)"
using assms FSwSbImorph_igWlsInp[of termMOD delta inp h hA MOD] by auto
lemma ipresWls_wlsBinp:
assumes "wlsBinp delta binp" and "ipresWlsAbs hA MOD"
shows "igWlsBinp MOD delta (lift hA binp)"
using assms imp_igWlsBinp[of termMOD delta binp hA MOD] by auto
lemma termFSwImorph_wlsBinp:
assumes "wlsBinp delta binp" and "termFSwImorph h hA MOD"
shows "igWlsBinp MOD delta (lift hA binp)"
using assms FSwImorph_igWlsBinp[of termMOD delta binp h hA MOD] by auto
lemma termFSwSbImorph_wlsBinp:
assumes "wlsBinp delta binp" and "termFSwSbImorph h hA MOD"
shows "igWlsBinp MOD delta (lift hA binp)"
using assms FSwSbImorph_igWlsBinp[of termMOD delta binp h hA MOD] by auto
lemma id_termFSwImorph: "termFSwImorph id id termMOD"
using id_FSwImorph[of termMOD] by simp
lemma id_termFSbImorph: "termFSbImorph id id termMOD"
using id_FSbImorph[of termMOD] by simp
lemma id_termFSwSbImorph: "termFSwSbImorph id id termMOD"
using id_FSwSbImorph[of termMOD] by simp
lemma comp_termFSwImorph:
assumes *: "termFSwImorph h hA MOD" and **: "FSwImorph h' hA' MOD MOD'"
shows "termFSwImorph (h' o h) (hA' o hA) MOD'"
using assms comp_FSwImorph[of h hA termMOD MOD h' hA' MOD'] by auto
lemma comp_termFSbImorph:
assumes *: "termFSbImorph h hA MOD" and **: "FSbImorph h' hA' MOD MOD'"
shows "termFSbImorph (h' o h) (hA' o hA) MOD'"
using assms comp_FSbImorph[of termMOD h hA MOD h' hA' MOD']
termMOD_igWlsAbsIsInBar by auto
lemma comp_termFSwSbImorph:
assumes *: "termFSwSbImorph h hA MOD" and **: "FSwSbImorph h' hA' MOD MOD'"
shows "termFSwSbImorph (h' o h) (hA' o hA) MOD'"
using assms comp_FSwSbImorph[of termMOD h hA MOD h' hA' MOD']
termMOD_igWlsAbsIsInBar by auto
lemmas mapFrom_termMOD_simps =
ipresIGWlsAll_termMOD_simps
ipresIGCons_termMOD_simps
ipresIGFreshAll_termMOD_simps
ipresIGSwapAll_termMOD_simps
ipresIGSubstAll_termMOD_simps
FSwImorph_termMOD FSbImorph_termMOD FSwSbImorph_termMOD
lemmas termMOD_simps =
structure_termMOD_simps mapFrom_termMOD_simps
subsubsection
-{* Sufficient criteria for being a morphism
- to a well-sorted model (of various kinds) *}
-
-text{* In a nutshell: in these cases, we only need to check preservation of the
- syntactic constructs, ``ipresCons". *}
+\<open>Sufficient criteria for being a morphism
+ to a well-sorted model (of various kinds)\<close>
+
+text\<open>In a nutshell: in these cases, we only need to check preservation of the
+ syntactic constructs, ``ipresCons".\<close>
lemma ipresCons_imp_ipresWlsAll:
assumes *: "ipresCons h hA MOD" and **: "igConsIPresIGWls MOD"
shows "ipresWlsAll h hA MOD"
proof-
{fix s X us s' A
have "(wls s X \<longrightarrow> igWls MOD s (h X)) \<and>
(wlsAbs (us,s') A \<longrightarrow> igWlsAbs MOD (us,s') (hA A))"
proof(induction rule: wls_rawInduct)
case (Var xs x)
then show ?case
by (metis assms igConsIPresIGWls_def igVarIPresIGWls_def ipresCons_def ipresVar_def)
next
case (Op delta inp binp)
have "igWlsInp MOD delta (lift h inp) \<and> igWlsBinp MOD delta (lift hA binp)"
using Op unfolding igWlsInp_def igWlsBinp_def wlsInp_iff wlsBinp_iff
by simp (simp add: liftAll2_def lift_def split: option.splits)
hence "igWls MOD (stOf delta) (igOp MOD delta (lift h inp) (lift hA binp))"
using ** unfolding igConsIPresIGWls_def igOpIPresIGWls_def by simp
thus ?case using Op * unfolding ipresCons_def ipresOp_def by simp
next
case (Abs s xs x X)
then show ?case
by (metis assms igAbsIPresIGWls_def igConsIPresIGWls_def ipresAbs_def ipresCons_def)
qed
}
thus ?thesis unfolding ipresWlsAll_defs by simp
qed
lemma ipresCons_imp_ipresFreshAll:
assumes *: "ipresCons h hA MOD" and **: "igFreshCls MOD"
and "igConsIPresIGWls MOD"
shows "ipresFreshAll h hA MOD"
proof-
have ***: "ipresWlsAll h hA MOD"
using assms ipresCons_imp_ipresWlsAll by auto
hence ****:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> igWlsInp MOD delta (lift h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> igWlsBinp MOD delta (lift hA binp)"
unfolding ipresWlsAll_def using ipresWls_wlsInp ipresWls_wlsBinp by auto
(* *)
{fix s X us s' A ys y
have "(wls s X \<longrightarrow> fresh ys y X \<longrightarrow> igFresh MOD ys y (h X)) \<and>
(wlsAbs (us,s') A \<longrightarrow> freshAbs ys y A \<longrightarrow> igFreshAbs MOD ys y (hA A))"
proof(induction rule: wls_rawInduct)
case (Var xs x)
then show ?case
by (metis * ** fresh_Var_simp igFreshCls_def igFreshIGVar_def ipresCons_def ipresVar_def)
next
case (Op delta inp binp)
show ?case proof safe
assume y_fresh: "fresh ys y (Op delta inp binp)"
{fix i X assume inp: "inp i = Some X"
then obtain s where "arOf delta i = Some s"
using Op unfolding wlsInp_iff sameDom_def by fastforce
hence "igFresh MOD ys y (h X)"
using Op.IH y_fresh inp unfolding freshInp_def liftAll_def liftAll2_def
by (metis freshInp_def liftAll_def wls_fresh_Op_simp)
}
moreover
{fix i A assume binp: "binp i = Some A"
then obtain us_s where "barOf delta i = Some us_s"
using Op unfolding wlsBinp_iff sameDom_def by force
hence "igFreshAbs MOD ys y (hA A)"
using Op.IH y_fresh binp unfolding freshBinp_def liftAll_def liftAll2_def
by simp (metis (no_types, hide_lams) freshBinp_def liftAll_def old.prod.exhaust)
}
ultimately have "igFreshInp MOD ys y (lift h inp) \<and> igFreshBinp MOD ys y (lift hA binp)"
unfolding igFreshInp_def igFreshBinp_def liftAll_lift_comp unfolding liftAll_def by auto
moreover have "igWlsInp MOD delta (lift h inp) \<and> igWlsBinp MOD delta (lift hA binp)"
using Op **** by simp
ultimately have "igFresh MOD ys y (igOp MOD delta (lift h inp) (lift hA binp))"
using ** unfolding igFreshCls_def igFreshIGOp_def by simp
thus "igFresh MOD ys y (h (Op delta inp binp))"
using Op * unfolding ipresCons_def ipresOp_def by simp
qed
next
case (Abs s xs x X)
hence hX_wls: "igWls MOD s (h X)"
using *** unfolding ipresWlsAll_def ipresWls_def by simp
thus ?case
using Abs assms by (cases "ys = xs \<and> y = x")
(simp_all add: igFreshCls_def igFreshIGAbs1_def igFreshIGAbs2_def ipresAbs_def ipresCons_def)
qed
}
thus ?thesis unfolding ipresFreshAll_defs by auto
qed
lemma ipresCons_imp_ipresSwapAll:
assumes *: "ipresCons h hA MOD" and **: "igSwapCls MOD"
and "igConsIPresIGWls MOD"
shows "ipresSwapAll h hA MOD"
proof-
have ***: "ipresWlsAll h hA MOD"
using assms ipresCons_imp_ipresWlsAll by auto
hence ****:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> igWlsInp MOD delta (lift h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> igWlsBinp MOD delta (lift hA binp)"
unfolding ipresWlsAll_def using ipresWls_wlsInp ipresWls_wlsBinp by auto
(* *)
{fix s X us s' A zs z1 z2
have "(wls s X \<longrightarrow> h (swap zs z1 z2 X) = igSwap MOD zs z1 z2 (h X)) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA (swapAbs zs z1 z2 A) = igSwapAbs MOD zs z1 z2 (hA A))"
proof(induction rule: wls_rawInduct)
case (Var xs x)
then show ?case
by (metis "*" "**" igSwapCls_def igSwapIGVar_def ipresCons_def ipresVar_def swap_Var_simp)
next
case (Op delta inp binp)
let ?inpsw = "swapInp zs z1 z2 inp" let ?binpsw = "swapBinp zs z1 z2 binp"
let ?Left = "h (Op delta ?inpsw ?binpsw)"
let ?Right = "igSwap MOD zs z1 z2 (h (Op delta inp binp))"
have wlsLiftInp:
"igWlsInp MOD delta (lift h inp) \<and> igWlsBinp MOD delta (lift hA binp)"
using Op **** by simp
have "wlsInp delta ?inpsw \<and> wlsBinp delta ?binpsw"
using Op by simp
hence "?Left = igOp MOD delta (lift h ?inpsw) (lift hA ?binpsw)"
using * unfolding ipresCons_def ipresOp_def by simp
moreover
have "lift h ?inpsw = igSwapInp MOD zs z1 z2 (lift h inp) \<and>
lift hA ?binpsw = igSwapBinp MOD zs z1 z2 (lift hA binp)"
using Op * not_None_eq
by (simp add: igSwapCls_def igSwapIGOp_def wlsInp_iff wlsBinp_iff
swapInp_def swapBinp_def igSwapInp_def igSwapBinp_def
lift_comp fun_eq_iff liftAll2_def lift_def sameDom_def split: option.splits)
(metis not_None_eq old.prod.exhaust)
moreover
have "igOp MOD delta (igSwapInp MOD zs z1 z2 (lift h inp))
(igSwapBinp MOD zs z1 z2 (lift hA binp)) =
igSwap MOD zs z1 z2 (igOp MOD delta (lift h inp) (lift hA binp))"
using wlsLiftInp ** unfolding igSwapCls_def igSwapIGOp_def by simp
moreover
have "igSwap MOD zs z1 z2 (igOp MOD delta (lift h inp) (lift hA binp)) = ?Right"
using Op * unfolding ipresCons_def ipresOp_def by simp
ultimately have "?Left = ?Right" by simp
then show ?case by (simp add: Op)
next
case (Abs s xs x X)
let ?Xsw = "swap zs z1 z2 X" let ?xsw = "x @xs[z1 \<and> z2]_zs"
have hX: "igWls MOD s (h X)" using Abs.IH *** unfolding ipresWlsAll_def ipresWls_def by simp
let ?Left = "hA (Abs xs ?xsw ?Xsw)"
let ?Right = "igSwapAbs MOD zs z1 z2 (hA (Abs xs x X))"
have "wls s (swap zs z1 z2 X)" using Abs by simp
hence "?Left = igAbs MOD xs ?xsw (h ?Xsw)"
using Abs * unfolding ipresCons_def ipresAbs_def by blast
also note Abs(3)
also have "igAbs MOD xs ?xsw (igSwap MOD zs z1 z2 (h X)) =
igSwapAbs MOD zs z1 z2 (igAbs MOD xs x (h X))"
using Abs hX ** by (auto simp: igSwapCls_def igSwapIGAbs_def)
also have "\<dots> = ?Right" using Abs * by (auto simp: ipresCons_def ipresAbs_def)
finally have "?Left = ?Right" .
then show ?case using Abs(2) by auto
qed
}
thus ?thesis unfolding ipresSwapAll_defs by auto
qed
lemma ipresCons_imp_ipresSubstAll_aux:
assumes *: "ipresCons h hA MOD" and **: "igSubstCls MOD"
and "igConsIPresIGWls MOD" and "igFreshCls MOD"
assumes P: "wlsPar P"
shows
"(wls s X \<longrightarrow>
(\<forall> ys y Y. y \<in> varsOfS P ys \<and> Y \<in> termsOfS P (asSort ys) \<longrightarrow>
h (X #[Y / y]_ys) = igSubst MOD ys (h Y) y (h X)))
\<and>
(wlsAbs (us,s') A \<longrightarrow>
(\<forall> ys y Y. y \<in> varsOfS P ys \<and> Y \<in> termsOfS P (asSort ys) \<longrightarrow>
hA (A $[Y / y]_ys) = igSubstAbs MOD ys (h Y) y (hA A)))"
proof-
have ***: "ipresWlsAll h hA MOD"
using assms ipresCons_imp_ipresWlsAll by auto
hence ****:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> igWlsInp MOD delta (lift h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> igWlsBinp MOD delta (lift hA binp)"
unfolding ipresWlsAll_def using ipresWls_wlsInp ipresWls_wlsBinp by auto
have *****: "ipresFreshAll h hA MOD"
using assms ipresCons_imp_ipresFreshAll by auto
(* *)
show ?thesis
proof(induction rule: wls_induct_fresh[of P])
case Par
then show ?case using P by auto
next
case (Var xs x)
then show ?case using assms
by (simp add: ipresWlsAll_def ipresWls_def igSubstCls_def igSubstIGVar2_def
ipresCons_def ipresVar_def)
(metis "***" FixSyn.ipresWlsAll_defs(1) FixSyn.ipresWlsAll_defs(2) FixSyn_axioms
igSubstIGVar1_def wlsPar_def wls_subst_Var_simp1 wls_subst_Var_simp2)
next
case (Op delta inp binp)
show ?case proof safe
fix ys y Y
assume yP: "y \<in> varsOfS P ys" and YP: "Y \<in> termsOfS P (asSort ys)"
hence Y: "wls (asSort ys) Y" using P by auto
hence hY: "igWls MOD (asSort ys) (h Y)"
using *** unfolding ipresWlsAll_def ipresWls_def by simp
have sinp: "wlsInp delta (substInp ys Y y inp) \<and>
wlsBinp delta (substBinp ys Y y binp)" using Y Op by simp
have liftInp: "igWlsInp MOD delta (lift h inp) \<and>
igWlsBinp MOD delta (lift hA binp)"
using Op **** by simp
let ?Left = "h ((Op delta inp binp) #[Y / y]_ys)"
let ?Right = "igSubst MOD ys (h Y) y (h (Op delta inp binp))"
have "?Left = igOp MOD delta (lift h (substInp ys Y y inp))
(lift hA (substBinp ys Y y binp))"
using sinp * unfolding ipresCons_def ipresOp_def
by (simp add: Op.IH(1) Op.IH(2) Y)
moreover
have "lift h (substInp ys Y y inp) = igSubstInp MOD ys (h Y) y (lift h inp) \<and>
lift hA (substBinp ys Y y binp) = igSubstBinp MOD ys (h Y) y (lift hA binp)"
using Op YP yP by (simp add: substInp_def2 igSubstInp_def substBinp_def2 igSubstBinp_def lift_comp
lift_def liftAll2_def fun_eq_iff wlsInp_iff wlsBinp_iff sameDom_def split: option.splits)
(metis (no_types, hide_lams) not_Some_eq option.distinct(1) sinp wlsBinp.simps)
moreover
have "igOp MOD delta (igSubstInp MOD ys (h Y) y (lift h inp))
(igSubstBinp MOD ys (h Y) y (lift hA binp)) =
igSubst MOD ys (h Y) y (igOp MOD delta (lift h inp) (lift hA binp))"
using hY liftInp ** unfolding igSubstCls_def igSubstIGOp_def by simp
moreover have "\<dots> = ?Right" using Op * unfolding ipresCons_def ipresOp_def by simp
ultimately show "?Left = ?Right" by simp
qed
next
case (Abs s xs x X)
show ?case proof safe
fix ys y Y
assume yP: "y \<in> varsOfS P ys" and YP: "Y \<in> termsOfS P (asSort ys)"
hence x_diff: "ys \<noteq> xs \<or> y \<noteq> x"
and Y: "wls (asSort ys) Y" and x_fresh: "fresh xs x Y" using P Abs by auto
hence hY: "igWls MOD (asSort ys) (h Y)"
using *** unfolding ipresWlsAll_def ipresWls_def by simp
have hX: "igWls MOD s (h X)"
using Abs *** unfolding ipresWlsAll_def ipresWls_def by simp
let ?Xsb = "subst ys Y y X"
have Xsb: "wls s ?Xsb" using Y Abs by simp
have x_igFresh: "igFresh MOD xs x (h Y)"
using Y x_fresh ***** unfolding ipresFreshAll_def ipresFresh_def by simp
let ?Left = "hA (Abs xs x X $[Y / y]_ys)"
let ?Right = "igSubstAbs MOD ys (h Y) y (hA (Abs xs x X))"
have "?Left = hA (Abs xs x ?Xsb)" using Y Abs x_diff x_fresh by auto
also have "\<dots> = igAbs MOD xs x (h ?Xsb)"
using Abs Xsb * unfolding ipresCons_def ipresAbs_def by fastforce
also have "\<dots> = igAbs MOD xs x (igSubst MOD ys (h Y) y (h X))"
using yP YP Abs.IH by simp
also have "\<dots> = igSubstAbs MOD ys (h Y) y (igAbs MOD xs x (h X))"
using Abs hY hX x_diff x_igFresh **
by (auto simp: igSubstCls_def igSubstIGAbs_def)
also have "\<dots> = ?Right" using Abs * by (auto simp: ipresCons_def ipresAbs_def)
finally show "?Left = ?Right" .
qed
qed
qed
lemma ipresCons_imp_ipresSubst:
assumes *: "ipresCons h hA MOD" and **: "igSubstCls MOD"
and "igConsIPresIGWls MOD" and "igFreshCls MOD"
shows "ipresSubst h MOD"
unfolding ipresSubst_def apply clarify
subgoal for ys Y y s X
using assms ipresCons_imp_ipresSubstAll_aux
[of h hA MOD
"ParS (\<lambda>zs. if zs = ys then [y] else [])
(\<lambda>s'. if s' = asSort ys then [Y] else [])
(\<lambda>_. [])
[]"]
unfolding wlsPar_def by auto .
lemma ipresCons_imp_ipresSubstAbs:
assumes *: "ipresCons h hA MOD" and **: "igSubstCls MOD"
and "igConsIPresIGWls MOD" and "igFreshCls MOD"
shows "ipresSubstAbs h hA MOD"
unfolding ipresSubstAbs_def apply clarify
subgoal for ys Y y us s A
using assms ipresCons_imp_ipresSubstAll_aux
[of h hA MOD
"ParS (\<lambda>zs. if zs = ys then [y] else [])
(\<lambda>s'. if s' = asSort ys then [Y] else [])
(\<lambda>_. [])
[]"]
unfolding wlsPar_def by auto .
lemma ipresCons_imp_ipresSubstAll:
assumes *: "ipresCons h hA MOD" and **: "igSubstCls MOD"
and "igConsIPresIGWls MOD" and "igFreshCls MOD"
shows "ipresSubstAll h hA MOD"
unfolding ipresSubstAll_def using assms
ipresCons_imp_ipresSubst ipresCons_imp_ipresSubstAbs by auto
lemma iwlsFSw_termFSwImorph_iff:
"iwlsFSw MOD \<Longrightarrow> termFSwImorph h hA MOD = ipresCons h hA MOD"
unfolding iwlsFSw_def termFSwImorph_def
using ipresCons_imp_ipresWlsAll
ipresCons_imp_ipresFreshAll ipresCons_imp_ipresSwapAll by auto
corollary iwlsFSwSTR_termFSwImorph_iff:
"iwlsFSwSTR MOD \<Longrightarrow> termFSwImorph h hA MOD = ipresCons h hA MOD"
using iwlsFSwSTR_imp_iwlsFSw iwlsFSw_termFSwImorph_iff by fastforce
lemma iwlsFSb_termFSbImorph_iff:
"iwlsFSb MOD \<Longrightarrow> termFSbImorph h hA MOD = ipresCons h hA MOD"
unfolding iwlsFSb_def termFSbImorph_def
using ipresCons_imp_ipresWlsAll
ipresCons_imp_ipresFreshAll ipresCons_imp_ipresSubstAll
unfolding igSubstCls_def by fastforce+
corollary iwlsFSbSwTR_termFSbImorph_iff:
"iwlsFSbSwTR MOD \<Longrightarrow> termFSbImorph h hA MOD = ipresCons h hA MOD"
using iwlsFSbSwTR_imp_iwlsFSb iwlsFSb_termFSbImorph_iff by fastforce
lemma iwlsFSwSb_termFSwSbImorph_iff:
"iwlsFSwSb MOD \<Longrightarrow> termFSwSbImorph h hA MOD = ipresCons h hA MOD"
unfolding termFSwSbImorph_def iwlsFSwSb_def
apply(simp add: iwlsFSw_termFSwImorph_iff)
unfolding iwlsFSw_def using ipresCons_imp_ipresSubstAll by auto
lemma iwlsFSbSw_termFSwSbImorph_iff:
"iwlsFSbSw MOD \<Longrightarrow> termFSwSbImorph h hA MOD = ipresCons h hA MOD"
unfolding termFSwSbImorph_iff iwlsFSbSw_def
apply(simp add: iwlsFSb_termFSbImorph_iff)
unfolding iwlsFSb_def using ipresCons_imp_ipresSwapAll by auto
end (* context FixSyn *)
-subsection{* The ``error" model of associated to a model *}
-
-text{* The error model will have the operators act like the original ones
+subsection\<open>The ``error" model of associated to a model\<close>
+
+text\<open>The error model will have the operators act like the original ones
on well-formed terms, except that will return ``ERR" (error) or ``True" (in the case of fresh)
whenever one of the inputs (variables, terms or abstractions) is ``ERR" or
is not well-formed.
The error model is more convenient than the original one, since
one can define more easily a map from the model of terms to the former. This map shall be defined
by the universal property of quotients, via a map from quasi-terms whose kernel
includes the alpha-equivalence relation. The latter property (of including
the alpha-equivalence would not be achievable with the original model as tariget, since
alpha is defined unsortedly and the model clauses hold sortedly.
-We shall only need error models associated to fresh-swap and to fresh-subst models. *}
-
-subsubsection {* Preliminaries *}
+We shall only need error models associated to fresh-swap and to fresh-subst models.\<close>
+
+subsubsection \<open>Preliminaries\<close>
(* I prefer defining a new type to using the option type, since
I already use options for inputs: *)
datatype 'a withERR = ERR | OK 'a
(* *************************************************** *)
context FixSyn (* scope all throuighout the file *)
begin
definition OKI where
"OKI inp = lift OK inp"
definition check where
"check eX == THE X. eX = OK X"
definition checkI where
"checkI einp == lift check einp"
lemma check_ex_unique:
"eX \<noteq> ERR \<Longrightarrow> (EX! X. eX = OK X)"
by(cases eX, auto)
lemma check_OK[simp]:
"check (OK X) = X"
unfolding check_def using check_ex_unique theI' by auto
lemma OK_check[simp]:
"eX \<noteq> ERR \<Longrightarrow> OK (check eX) = eX"
unfolding check_def using check_ex_unique theI' by auto
lemma checkI_OKI[simp]:
"checkI (OKI inp) = inp"
unfolding OKI_def checkI_def lift_def apply(rule ext)
by(case_tac "inp i", auto)
lemma OKI_checkI[simp]:
assumes "liftAll (\<lambda> X. X \<noteq> ERR) einp"
shows "OKI (checkI einp) = einp"
unfolding OKI_def checkI_def lift_def apply(rule ext)
using assms unfolding liftAll_def by (case_tac "einp i", auto)
lemma OKI_inj[simp]:
fixes inp inp' :: "('index,'gTerm)input"
shows "(OKI inp = OKI inp') = (inp = inp')"
apply(auto) unfolding OKI_def
using lift_preserves_inj[of OK]
unfolding inj_on_def by auto
lemmas OK_OKI_simps =
check_OK OK_check checkI_OKI OKI_checkI OKI_inj
-subsubsection {* Definitions and notations *}
+subsubsection \<open>Definitions and notations\<close>
definition errMOD ::
"('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model \<Rightarrow>
('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm withERR,'gAbs withERR)model"
where
"errMOD MOD ==
\<lparr>igWls = \<lambda> s eX. case eX of ERR \<Rightarrow> False | OK X \<Rightarrow> igWls MOD s X,
igWlsAbs = \<lambda> (us,s) eA. case eA of ERR \<Rightarrow> False | OK A \<Rightarrow> igWlsAbs MOD (us,s) A,
igVar = \<lambda> xs x. OK (igVar MOD xs x),
igAbs = \<lambda>xs x eX.
if (eX \<noteq> ERR \<and> (\<exists> s. isInBar (xs,s) \<and> igWls MOD s (check eX)))
then OK (igAbs MOD xs x (check eX))
else ERR,
igOp = \<lambda>delta einp ebinp.
if liftAll (\<lambda> X. X \<noteq> ERR) einp \<and> liftAll (\<lambda> A. A \<noteq> ERR) ebinp
\<and> igWlsInp MOD delta (checkI einp) \<and> igWlsBinp MOD delta (checkI ebinp)
then OK (igOp MOD delta (checkI einp) (checkI ebinp))
else ERR,
igFresh = \<lambda>ys y eX.
if eX \<noteq> ERR \<and> (\<exists> s. igWls MOD s (check eX))
then igFresh MOD ys y (check eX)
else True,
igFreshAbs = \<lambda>ys y eA.
if eA \<noteq> ERR \<and> (\<exists> us s. igWlsAbs MOD (us,s) (check eA))
then igFreshAbs MOD ys y (check eA)
else True,
igSwap = \<lambda>zs z1 z2 eX.
if eX \<noteq> ERR \<and> (\<exists> s. igWls MOD s (check eX))
then OK (igSwap MOD zs z1 z2 (check eX))
else ERR,
igSwapAbs = \<lambda>zs z1 z2 eA.
if eA \<noteq> ERR \<and> (\<exists> us s. igWlsAbs MOD (us,s) (check eA))
then OK (igSwapAbs MOD zs z1 z2 (check eA))
else ERR,
igSubst = \<lambda>ys eY y eX.
if eY \<noteq> ERR \<and> igWls MOD (asSort ys) (check eY)
\<and> eX \<noteq> ERR \<and> (\<exists> s. igWls MOD s (check eX))
then OK (igSubst MOD ys (check eY) y (check eX))
else ERR,
igSubstAbs = \<lambda>ys eY y eA.
if eY \<noteq> ERR \<and> igWls MOD (asSort ys) (check eY)
\<and> eA \<noteq> ERR \<and> (\<exists> us s. igWlsAbs MOD (us,s) (check eA))
then OK (igSubstAbs MOD ys (check eY) y (check eA))
else ERR
\<rparr>"
abbreviation eWls where "eWls MOD == igWls (errMOD MOD)"
abbreviation eWlsAbs where "eWlsAbs MOD == igWlsAbs (errMOD MOD)"
abbreviation eWlsInp where "eWlsInp MOD == igWlsInp (errMOD MOD)"
abbreviation eWlsBinp where "eWlsBinp MOD == igWlsBinp (errMOD MOD)"
abbreviation eVar where "eVar MOD == igVar (errMOD MOD)"
abbreviation eAbs where "eAbs MOD == igAbs (errMOD MOD)"
abbreviation eOp where "eOp MOD == igOp (errMOD MOD)"
abbreviation eFresh where "eFresh MOD == igFresh (errMOD MOD)"
abbreviation eFreshAbs where "eFreshAbs MOD == igFreshAbs (errMOD MOD)"
abbreviation eFreshInp where "eFreshInp MOD == igFreshInp (errMOD MOD)"
abbreviation eFreshBinp where "eFreshBinp MOD == igFreshBinp (errMOD MOD)"
abbreviation eSwap where "eSwap MOD == igSwap (errMOD MOD)"
abbreviation eSwapAbs where "eSwapAbs MOD == igSwapAbs (errMOD MOD)"
abbreviation eSwapInp where "eSwapInp MOD == igSwapInp (errMOD MOD)"
abbreviation eSwapBinp where "eSwapBinp MOD == igSwapBinp (errMOD MOD)"
abbreviation eSubst where "eSubst MOD == igSubst (errMOD MOD)"
abbreviation eSubstAbs where "eSubstAbs MOD == igSubstAbs (errMOD MOD)"
abbreviation eSubstInp where "eSubstInp MOD == igSubstInp (errMOD MOD)"
abbreviation eSubstBinp where "eSubstBinp MOD == igSubstBinp (errMOD MOD)"
-subsubsection {* Simplification rules *}
+subsubsection \<open>Simplification rules\<close>
lemma eWls_simp1[simp]:
"eWls MOD s (OK X) = igWls MOD s X"
unfolding errMOD_def by simp
lemma eWls_simp2[simp]:
"eWls MOD s ERR = False"
unfolding errMOD_def by simp
lemma eWlsAbs_simp1[simp]:
"eWlsAbs MOD (us,s) (OK A) = igWlsAbs MOD (us,s) A"
unfolding errMOD_def by simp
lemma eWlsAbs_simp2[simp]:
"eWlsAbs MOD (us,s) ERR = False"
unfolding errMOD_def by simp
lemma eWlsInp_simp1[simp]:
"eWlsInp MOD delta (OKI inp) = igWlsInp MOD delta inp"
by (fastforce simp: OKI_def sameDom_def liftAll2_def lift_def igWlsInp_def
split: option.splits)
lemma eWlsInp_simp2[simp]:
"\<not> liftAll (\<lambda> eX. eX \<noteq> ERR) einp \<Longrightarrow> \<not> eWlsInp MOD delta einp"
by (force simp: sameDom_def liftAll_def liftAll2_def lift_def igWlsInp_def)
corollary eWlsInp_simp3[simp]:
"\<not> eWlsInp MOD delta (\<lambda>i. Some ERR)"
by (auto simp: liftAll_def)
lemma eWlsBinp_simp1[simp]:
"eWlsBinp MOD delta (OKI binp) = igWlsBinp MOD delta binp"
by (fastforce simp: OKI_def sameDom_def liftAll2_def lift_def igWlsBinp_def
split: option.splits)
lemma eWlsBinp_simp2[simp]:
"\<not> liftAll (\<lambda> eA. eA \<noteq> ERR) ebinp \<Longrightarrow> \<not> eWlsBinp MOD delta ebinp"
by (force simp: sameDom_def liftAll_def liftAll2_def lift_def igWlsBinp_def)
corollary eWlsBinp_simp3[simp]:
"\<not> eWlsBinp MOD delta (\<lambda>i. Some ERR)"
by (auto simp: liftAll_def)
lemmas eWlsAll_simps =
eWls_simp1 eWls_simp2
eWlsAbs_simp1 eWlsAbs_simp2
eWlsInp_simp1 eWlsInp_simp2 eWlsInp_simp3
eWlsBinp_simp1 eWlsBinp_simp2 eWlsBinp_simp3
lemma eVar_simp[simp]:
"eVar MOD xs x = OK (igVar MOD xs x)"
unfolding errMOD_def by simp
lemma eAbs_simp1[simp]:
"\<lbrakk>isInBar (xs,s); igWls MOD s X\<rbrakk> \<Longrightarrow> eAbs MOD xs x (OK X) = OK (igAbs MOD xs x X)"
unfolding errMOD_def by auto
lemma eAbs_simp2[simp]:
"\<forall> s. \<not> (isInBar (xs,s) \<and> igWls MOD s X) \<Longrightarrow> eAbs MOD xs x (OK X) = ERR"
unfolding errMOD_def by auto
lemma eAbs_simp3[simp]:
"eAbs MOD xs x ERR = ERR"
unfolding errMOD_def by auto
lemma eOp_simp1[simp]:
assumes "igWlsInp MOD delta inp" and "igWlsBinp MOD delta binp"
shows "eOp MOD delta (OKI inp) (OKI binp) = OK (igOp MOD delta inp binp)"
unfolding errMOD_def apply simp
unfolding liftAll_def OKI_def lift_def
using assms by (auto split: option.splits)
lemma eOp_simp2[simp]:
assumes "\<not> igWlsInp MOD delta inp"
shows "eOp MOD delta (OKI inp) ebinp = ERR"
using assms unfolding errMOD_def by auto
lemma eOp_simp3[simp]:
assumes "\<not> igWlsBinp MOD delta binp"
shows "eOp MOD delta einp (OKI binp) = ERR"
using assms unfolding errMOD_def by auto
lemma eOp_simp4[simp]:
assumes "\<not> liftAll (\<lambda> eX. eX \<noteq> ERR) einp"
shows "eOp MOD delta einp ebinp = ERR"
using assms unfolding errMOD_def by auto
corollary eOp_simp5[simp]:
"eOp MOD delta (\<lambda>i. Some ERR) ebinp = ERR"
by (auto simp: liftAll_def)
lemma eOp_simp6[simp]:
assumes "\<not> liftAll (\<lambda> eA. eA \<noteq> ERR) ebinp"
shows "eOp MOD delta einp ebinp = ERR"
using assms unfolding errMOD_def by auto
corollary eOp_simp7[simp]:
"eOp MOD delta einp (\<lambda>i. Some ERR) = ERR"
by (auto simp: liftAll_def)
lemmas eCons_simps =
eVar_simp
eAbs_simp1 eAbs_simp2 eAbs_simp3
eOp_simp1 eOp_simp2 eOp_simp3 eOp_simp4 eOp_simp5 eOp_simp6 eOp_simp7
lemma eFresh_simp1[simp]:
"igWls MOD s X \<Longrightarrow> eFresh MOD ys y (OK X) = igFresh MOD ys y X"
unfolding errMOD_def by auto
lemma eFresh_simp2[simp]:
"\<forall> s. \<not> igWls MOD s X \<Longrightarrow> eFresh MOD ys y (OK X)"
unfolding errMOD_def by auto
lemma eFresh_simp3[simp]:
"eFresh MOD ys y ERR"
unfolding errMOD_def by auto
lemma eFreshAbs_simp1[simp]:
"igWlsAbs MOD (us,s) A \<Longrightarrow> eFreshAbs MOD ys y (OK A) = igFreshAbs MOD ys y A"
unfolding errMOD_def by auto
lemma eFreshAbs_simp2[simp]:
"\<forall> us s. \<not> igWlsAbs MOD (us,s) A \<Longrightarrow> eFreshAbs MOD ys y (OK A)"
unfolding errMOD_def by auto
lemma eFreshAbs_simp3[simp]:
"eFreshAbs MOD ys y ERR"
unfolding errMOD_def by auto
lemma eFreshInp_simp[simp]:
"igWlsInp MOD delta inp
\<Longrightarrow> eFreshInp MOD ys y (OKI inp) = igFreshInp MOD ys y inp"
by (force simp: igFreshInp_def OKI_def liftAll_lift_comp igWlsInp_defs intro!: liftAll_cong)
lemma eFreshBinp_simp[simp]:
"igWlsBinp MOD delta binp
\<Longrightarrow> eFreshBinp MOD ys y (OKI binp) = igFreshBinp MOD ys y binp"
by (force simp: igFreshBinp_def OKI_def liftAll_lift_comp igWlsBinp_defs intro!: liftAll_cong)
lemmas eFreshAll_simps =
eFresh_simp1 eFresh_simp2 eFresh_simp3
eFreshAbs_simp1 eFreshAbs_simp2 eFreshAbs_simp3
eFreshInp_simp
eFreshBinp_simp
lemma eSwap_simp1[simp]:
"igWls MOD s X
\<Longrightarrow> eSwap MOD zs z1 z2 (OK X) = OK (igSwap MOD zs z1 z2 X)"
unfolding errMOD_def by auto
lemma eSwap_simp2[simp]:
"\<forall> s. \<not> igWls MOD s X \<Longrightarrow> eSwap MOD zs z1 z2 (OK X) = ERR"
unfolding errMOD_def by auto
lemma eSwap_simp3[simp]:
"eSwap MOD zs z1 z2 ERR = ERR"
unfolding errMOD_def by auto
lemma eSwapAbs_simp1[simp]:
"igWlsAbs MOD (us,s) A
\<Longrightarrow> eSwapAbs MOD zs z1 z2 (OK A) = OK (igSwapAbs MOD zs z1 z2 A)"
unfolding errMOD_def by auto
lemma eSwapAbs_simp2[simp]:
"\<forall> us s. \<not> igWlsAbs MOD (us,s) A \<Longrightarrow> eSwapAbs MOD zs z1 z2 (OK A) = ERR"
unfolding errMOD_def by auto
lemma eSwapAbs_simp3[simp]:
"eSwapAbs MOD zs z1 z2 ERR = ERR"
unfolding errMOD_def by auto
lemma eSwapInp_simp1[simp]:
"igWlsInp MOD delta inp
\<Longrightarrow> eSwapInp MOD zs z1 z2 (OKI inp) = OKI (igSwapInp MOD zs z1 z2 inp)"
by (force simp: igSwapInp_def OKI_def lift_comp igWlsInp_defs intro!: lift_cong)
lemma eSwapInp_simp2[simp]:
assumes "\<not> liftAll (\<lambda> eX. eX \<noteq> ERR) einp"
shows "\<not> liftAll (\<lambda> eX. eX \<noteq> ERR) (eSwapInp MOD zs z1 z2 einp)"
using assms unfolding liftAll_def igSwapInp_def lift_def by (auto split: option.splits)
lemma eSwapBinp_simp1[simp]:
"igWlsBinp MOD delta binp
\<Longrightarrow> eSwapBinp MOD zs z1 z2 (OKI binp) = OKI (igSwapBinp MOD zs z1 z2 binp)"
by (force simp: igSwapBinp_def OKI_def lift_comp igWlsBinp_defs intro!: lift_cong)
lemma eSwapBinp_simp2[simp]:
assumes "\<not> liftAll (\<lambda> eA. eA \<noteq> ERR) ebinp"
shows "\<not> liftAll (\<lambda> eA. eA \<noteq> ERR) (eSwapBinp MOD zs z1 z2 ebinp)"
using assms unfolding liftAll_def igSwapBinp_def lift_def by (auto split: option.splits)
lemmas eSwapAll_simps =
eSwap_simp1 eSwap_simp2 eSwap_simp3
eSwapAbs_simp1 eSwapAbs_simp2 eSwapAbs_simp3
eSwapInp_simp1 eSwapInp_simp2
eSwapBinp_simp1 eSwapBinp_simp2
lemma eSubst_simp1[simp]:
"\<lbrakk>igWls MOD (asSort ys) Y; igWls MOD s X\<rbrakk>
\<Longrightarrow> eSubst MOD ys (OK Y) y (OK X) = OK (igSubst MOD ys Y y X)"
unfolding errMOD_def by auto
lemma eSubst_simp2[simp]:
"\<not> igWls MOD (asSort ys) Y \<Longrightarrow> eSubst MOD ys (OK Y) y eX = ERR"
unfolding errMOD_def by auto
lemma eSubst_simp3[simp]:
"\<forall> s. \<not> igWls MOD s X \<Longrightarrow> eSubst MOD ys eY y (OK X) = ERR"
unfolding errMOD_def by auto
lemma eSubst_simp4[simp]:
"eSubst MOD ys eY y ERR = ERR"
unfolding errMOD_def by auto
lemma eSubst_simp5[simp]:
"eSubst MOD ys ERR y eX = ERR"
unfolding errMOD_def by auto
lemma eSubstAbs_simp1[simp]:
"\<lbrakk>igWls MOD (asSort ys) Y; igWlsAbs MOD (us,s) A\<rbrakk>
\<Longrightarrow> eSubstAbs MOD ys (OK Y) y (OK A) = OK (igSubstAbs MOD ys Y y A)"
unfolding errMOD_def by auto
lemma eSubstAbs_simp2[simp]:
"\<not> igWls MOD (asSort ys) Y \<Longrightarrow> eSubstAbs MOD ys (OK Y) y eA = ERR"
unfolding errMOD_def by auto
lemma eSubstAbs_simp3[simp]:
"\<forall> us s. \<not> igWlsAbs MOD (us,s) A \<Longrightarrow> eSubstAbs MOD ys eY y (OK A) = ERR"
unfolding errMOD_def by auto
lemma eSubstAbs_simp4[simp]:
"eSubstAbs MOD ys eY y ERR = ERR"
unfolding errMOD_def by auto
lemma eSubstAbs_simp5[simp]:
"eSubstAbs MOD ys ERR y eA = ERR"
unfolding errMOD_def by auto
lemma eSubstInp_simp1[simp]:
"\<lbrakk>igWls MOD (asSort ys) Y; igWlsInp MOD delta inp\<rbrakk>
\<Longrightarrow> eSubstInp MOD ys (OK Y) y (OKI inp) = OKI (igSubstInp MOD ys Y y inp)"
by (force simp: igSubstInp_def OKI_def lift_comp igWlsInp_defs intro!: lift_cong)
lemma eSubstInp_simp2[simp]:
assumes "\<not> liftAll (\<lambda>eX. eX \<noteq> ERR) einp"
shows "\<not> liftAll (\<lambda>eX. eX \<noteq> ERR) (eSubstInp MOD ys eY y einp)"
using assms unfolding lift_def igSubstInp_def liftAll_def by (auto split: option.splits)
lemma eSubstInp_simp3[simp]:
assumes *: "\<not> igWls MOD (asSort ys) Y" and **: "\<not> einp = (\<lambda> i. None)"
shows "\<not> liftAll (\<lambda>eX. eX \<noteq> ERR) (eSubstInp MOD ys (OK Y) y einp)"
using assms by (auto simp: igSubstInp_def liftAll_lift_comp lift_def liftAll_def
split: option.splits)
lemma eSubstInp_simp4[simp]:
assumes "\<not> einp = (\<lambda> i. None)"
shows "\<not> liftAll (\<lambda>eX. eX \<noteq> ERR) (eSubstInp MOD ys ERR y einp)"
using assms by (auto simp: igSubstInp_def liftAll_lift_comp lift_def liftAll_def
split: option.splits)
lemma eSubstBinp_simp1[simp]:
"\<lbrakk>igWls MOD (asSort ys) Y; igWlsBinp MOD delta binp\<rbrakk>
\<Longrightarrow> eSubstBinp MOD ys (OK Y) y (OKI binp) = OKI (igSubstBinp MOD ys Y y binp)"
by (force simp: igSubstBinp_def OKI_def lift_comp igWlsBinp_defs intro!: lift_cong)
lemma eSubstBinp_simp2[simp]:
assumes "\<not> liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp"
shows "\<not> liftAll (\<lambda>eA. eA \<noteq> ERR) (eSubstBinp MOD ys eY y ebinp)"
using assms by (auto simp: igSubstBinp_def liftAll_lift_comp lift_def liftAll_def
split: option.splits)
lemma eSubstBinp_simp3[simp]:
assumes *: "\<not> igWls MOD (asSort ys) Y" and **: "\<not> ebinp = (\<lambda> i. None)"
shows "\<not> liftAll (\<lambda>eA. eA \<noteq> ERR) (eSubstBinp MOD ys (OK Y) y ebinp)"
using assms by (auto simp: igSubstBinp_def liftAll_lift_comp lift_def liftAll_def
split: option.splits)
lemma eSubstBinp_simp4[simp]:
assumes "\<not> ebinp = (\<lambda> i. None)"
shows "\<not> liftAll (\<lambda>eA. eA \<noteq> ERR) (eSubstBinp MOD ys ERR y ebinp)"
using assms by (auto simp: igSubstBinp_def liftAll_lift_comp lift_def liftAll_def
split: option.splits)
lemmas eSubstAll_simps =
eSubst_simp1 eSubst_simp2 eSubst_simp3 eSubst_simp4 eSubst_simp5
eSubstAbs_simp1 eSubstAbs_simp2 eSubstAbs_simp3 eSubstAbs_simp4 eSubstAbs_simp5
eSubstInp_simp1 eSubstInp_simp2 eSubstInp_simp3 eSubstInp_simp4
eSubstBinp_simp1 eSubstBinp_simp2 eSubstBinp_simp3 eSubstBinp_simp4
lemmas error_model_simps =
OK_OKI_simps
eWlsAll_simps
eCons_simps
eFreshAll_simps
eSwapAll_simps
eSubstAll_simps
-subsubsection {* Nchotomies *}
+subsubsection \<open>Nchotomies\<close>
lemma eWls_nchotomy:
"(\<exists> X. eX = OK X \<and> igWls MOD s X) \<or> \<not> eWls MOD s eX"
unfolding errMOD_def by(cases eX) auto
lemma eWlsAbs_nchotomy:
"(\<exists> A. eA = OK A \<and> igWlsAbs MOD (us,s) A) \<or> \<not> eWlsAbs MOD (us,s) eA"
unfolding errMOD_def by(cases eA) auto
lemma eAbs_nchotomy:
"((\<exists> s X. eX = OK X \<and> isInBar (xs,s) \<and> igWls MOD s X)) \<or> (eAbs MOD xs x eX = ERR)"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eOp_nchotomy:
"(\<exists> inp binp. einp = OKI inp \<and> igWlsInp MOD delta inp \<and>
ebinp = OKI binp \<and> igWlsBinp MOD delta binp)
\<or>
(eOp MOD delta einp ebinp = ERR)"
unfolding errMOD_def apply simp using OKI_checkI by force
lemma eFresh_nchotomy:
"(\<exists> s X. eX = OK X \<and> igWls MOD s X) \<or> eFresh MOD ys y eX"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eFreshAbs_nchotomy:
"(\<exists> us s A. eA = OK A \<and> igWlsAbs MOD (us,s) A)
\<or> eFreshAbs MOD ys y eA"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eSwap_nchotomy:
"(\<exists> s X. eX = OK X \<and> igWls MOD s X) \<or>
(eSwap MOD zs z1 z2 eX = ERR)"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eSwapAbs_nchotomy:
"(\<exists> us s A. eA = OK A \<and> igWlsAbs MOD (us,s) A) \<or>
(eSwapAbs MOD zs z1 z2 eA = ERR)"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eSubst_nchotomy:
"(\<exists> Y. eY = OK Y \<and>
igWls MOD (asSort ys) Y) \<and> (\<exists> s X. eX = OK X \<and> igWls MOD s X)
\<or>
(eSubst MOD ys eY y eX = ERR)"
unfolding errMOD_def apply simp using OK_check by fastforce
lemma eSubstAbs_nchotomy:
"(\<exists> Y. eY = OK Y \<and> igWls MOD (asSort ys) Y) \<and>
(\<exists> us s A. eA = OK A \<and> igWlsAbs MOD (us,s) A)
\<or>
(eSubstAbs MOD ys eY y eA = ERR)"
unfolding errMOD_def apply simp using OK_check by fastforce
-subsubsection {* Inversion rules *}
+subsubsection \<open>Inversion rules\<close>
lemma eWls_invert:
assumes "eWls MOD s eX"
shows "\<exists> X. eX = OK X \<and> igWls MOD s X"
using assms eWls_nchotomy by blast
lemma eWlsAbs_invert:
assumes "eWlsAbs MOD (us,s) eA"
shows "\<exists> A. eA = OK A \<and> igWlsAbs MOD (us,s) A"
using assms eWlsAbs_nchotomy by blast
lemma eWlsInp_invert:
assumes "eWlsInp MOD delta einp"
shows "\<exists> inp. igWlsInp MOD delta inp \<and> einp = OKI inp"
proof
let ?inp = "checkI einp"
have "wlsOpS delta" using assms unfolding igWlsInp_def by simp
moreover have "sameDom (arOf delta) ?inp"
using assms unfolding igWlsInp_def checkI_def by simp
moreover have "liftAll2 (igWls MOD) (arOf delta) ?inp"
using assms eWls_invert
by (fastforce simp: igWlsInp_def checkI_def liftAll2_def lift_def sameDom_def
split: option.splits)
ultimately have "igWlsInp MOD delta ?inp" unfolding igWlsInp_def by simp
moreover
{have "liftAll (\<lambda>eX. eX \<noteq> ERR) einp"
using assms using eWlsInp_simp2 by blast
hence "einp = OKI ?inp" by simp
}
ultimately show "igWlsInp MOD delta ?inp \<and> einp = OKI ?inp" by simp
qed
lemma eWlsBinp_invert:
assumes "eWlsBinp MOD delta ebinp"
shows "\<exists> binp. igWlsBinp MOD delta binp \<and> ebinp = OKI binp"
proof
let ?binp = "checkI ebinp"
have "wlsOpS delta" using assms unfolding igWlsBinp_def by simp
moreover have "sameDom (barOf delta) ?binp"
using assms unfolding igWlsBinp_def checkI_def by simp
moreover have "liftAll2 (igWlsAbs MOD) (barOf delta) ?binp"
using assms eWlsAbs_invert
by (fastforce simp: igWlsBinp_def checkI_def liftAll2_def lift_def sameDom_def
split: option.splits)
ultimately have "igWlsBinp MOD delta ?binp" unfolding igWlsBinp_def by simp
moreover
{have "liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp"
using assms using eWlsBinp_simp2 by blast
hence "ebinp = OKI ?binp" by simp
}
ultimately show "igWlsBinp MOD delta ?binp \<and> ebinp = OKI ?binp" by simp
qed
lemma eAbs_invert:
assumes "eAbs MOD xs x eX = OK A"
shows "\<exists> s X. eX = OK X \<and> isInBar (xs,s) \<and> A = igAbs MOD xs x X \<and> igWls MOD s X"
proof-
have 1: "eAbs MOD xs x eX \<noteq> ERR" using assms by auto
then obtain s X where *: "eX = OK X"
and **: "isInBar (xs,s)" and ***: "igWls MOD s X"
using eAbs_nchotomy[of eX] by fastforce
hence "eAbs MOD xs x eX = OK (igAbs MOD xs x X)" by simp
thus ?thesis using assms * ** *** by auto
qed
lemma eOp_invert:
assumes "eOp MOD delta einp ebinp = OK X"
shows
"\<exists> inp binp. einp = OKI inp \<and> ebinp = OKI binp \<and>
X = igOp MOD delta inp binp \<and>
igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp"
proof-
have "eOp MOD delta einp ebinp \<noteq> ERR" using assms by auto
then obtain inp binp where *: "einp = OKI inp" "ebinp = OKI binp"
"igWlsInp MOD delta inp" "igWlsBinp MOD delta binp"
using eOp_nchotomy by blast
hence "eOp MOD delta einp ebinp = OK (igOp MOD delta inp binp)" by simp
thus ?thesis using assms * by auto
qed
lemma eFresh_invert:
assumes "\<not> eFresh MOD ys y eX"
shows "\<exists> s X. eX = OK X \<and> \<not> igFresh MOD ys y X \<and> igWls MOD s X"
proof-
obtain s X where *: "eX = OK X" and **: "igWls MOD s X"
using assms eFresh_nchotomy[of eX] by fastforce
hence "eFresh MOD ys y eX = igFresh MOD ys y X" by simp
thus ?thesis using assms * ** by auto
qed
lemma eFreshAbs_invert:
assumes "\<not> eFreshAbs MOD ys y eA"
shows "\<exists> us s A. eA = OK A \<and> \<not> igFreshAbs MOD ys y A \<and> igWlsAbs MOD (us,s) A"
proof-
obtain us s A where *: "eA = OK A" and **: "igWlsAbs MOD (us,s) A"
using assms eFreshAbs_nchotomy[of eA] by fastforce
hence "eFreshAbs MOD ys y eA = igFreshAbs MOD ys y A" by simp
thus ?thesis using assms * ** by auto
qed
lemma eSwap_invert:
assumes "eSwap MOD zs z1 z2 eX = OK Y"
shows "\<exists> s X. eX = OK X \<and> Y = igSwap MOD zs z1 z2 X \<and> igWls MOD s X"
proof-
have 1: "eSwap MOD zs z1 z2 eX \<noteq> ERR" using assms by auto
then obtain s X where *: "eX = OK X" and **: "igWls MOD s X"
using eSwap_nchotomy[of eX] by fastforce
hence "eSwap MOD zs z1 z2 eX = OK (igSwap MOD zs z1 z2 X)" by simp
thus ?thesis using assms * ** by auto
qed
lemma eSwapAbs_invert:
assumes "eSwapAbs MOD zs z1 z2 eA = OK B"
shows "\<exists> us s A. eA = OK A \<and> B = igSwapAbs MOD zs z1 z2 A \<and> igWlsAbs MOD (us,s) A"
proof-
have 1: "eSwapAbs MOD zs z1 z2 eA \<noteq> ERR" using assms by auto
then obtain us s A where *: "eA = OK A" and **: "igWlsAbs MOD (us,s) A"
using eSwapAbs_nchotomy[of eA] by fastforce
hence "eSwapAbs MOD zs z1 z2 eA = OK (igSwapAbs MOD zs z1 z2 A)" by simp
thus ?thesis using assms * ** by auto
qed
lemma eSubst_invert:
assumes "eSubst MOD ys eY y eX = OK Z"
shows
"\<exists> s X Y. eY = OK Y \<and> eX = OK X \<and> igWls MOD s X \<and> igWls MOD (asSort ys) Y \<and>
Z = igSubst MOD ys Y y X"
proof-
have 1: "eSubst MOD ys eY y eX \<noteq> ERR" using assms by auto
then obtain s X Y where *: "eX = OK X" "eY = OK Y"
"igWls MOD s X" "igWls MOD (asSort ys) Y"
using eSubst_nchotomy[of eY _ _ eX] by fastforce
hence "eSubst MOD ys eY y eX = OK (igSubst MOD ys Y y X)" by simp
thus ?thesis using assms * by auto
qed
lemma eSubstAbs_invert:
assumes "eSubstAbs MOD ys eY y eA = OK Z"
shows
"\<exists> us s A Y. eY = OK Y \<and> eA = OK A \<and> igWlsAbs MOD (us,s) A \<and> igWls MOD (asSort ys) Y \<and>
Z = igSubstAbs MOD ys Y y A"
proof-
have 1: "eSubstAbs MOD ys eY y eA \<noteq> ERR" using assms by auto
then obtain us s A Y where *: "eA = OK A" "eY = OK Y"
"igWlsAbs MOD (us,s) A" "igWls MOD (asSort ys) Y"
using eSubstAbs_nchotomy[of eY _ _ eA] by fastforce
hence "eSubstAbs MOD ys eY y eA = OK (igSubstAbs MOD ys Y y A)" by simp
thus ?thesis using assms * by auto
qed
-subsubsection {* The error model is strongly well-sorted
-as a fresh-swap-subst and as a fresh-subst-swap model *}
-
-text{* That is, provided the original model is a well-sorted fresh-swap model. *}
-
-text{* The domains are disjoint: *}
+subsubsection \<open>The error model is strongly well-sorted
+as a fresh-swap-subst and as a fresh-subst-swap model\<close>
+
+text\<open>That is, provided the original model is a well-sorted fresh-swap model.\<close>
+
+text\<open>The domains are disjoint:\<close>
lemma errMOD_igWlsDisj:
assumes "igWlsDisj MOD"
shows "igWlsDisj (errMOD MOD)"
using assms unfolding errMOD_def igWlsDisj_def
apply clarify subgoal for _ _ X by(cases X) auto .
lemma errMOD_igWlsAbsDisj:
assumes "igWlsAbsDisj MOD"
shows "igWlsAbsDisj (errMOD MOD)"
using assms unfolding errMOD_def igWlsAbsDisj_def
apply clarify subgoal for _ _ _ _ A by(cases A) fastforce+ .
lemma errMOD_igWlsAllDisj:
assumes "igWlsAllDisj MOD"
shows "igWlsAllDisj (errMOD MOD)"
using assms unfolding igWlsAllDisj_def
using errMOD_igWlsDisj errMOD_igWlsAbsDisj by auto
-text{* Only ``bound arity" abstraction domains are inhabited: *}
+text\<open>Only ``bound arity" abstraction domains are inhabited:\<close>
lemma errMOD_igWlsAbsIsInBar:
assumes "igWlsAbsIsInBar MOD"
shows "igWlsAbsIsInBar (errMOD MOD)"
using assms eWlsAbs_invert unfolding igWlsAbsIsInBar_def by blast
-text{* The operators preserve the domains strongly: *}
+text\<open>The operators preserve the domains strongly:\<close>
lemma errMOD_igVarIPresIGWlsSTR:
assumes "igVarIPresIGWls MOD"
shows "igVarIPresIGWls (errMOD MOD)"
using assms unfolding errMOD_def igVarIPresIGWls_def by simp
lemma errMOD_igAbsIPresIGWlsSTR:
assumes *: "igAbsIPresIGWls MOD" and **: "igWlsAbsDisj MOD"
and ***: "igWlsAbsIsInBar MOD"
shows "igAbsIPresIGWlsSTR (errMOD MOD)"
using assms by (fastforce simp: errMOD_def igAbsIPresIGWls_def igAbsIPresIGWlsSTR_def
igWlsAbsIsInBar_def igWlsAbsDisj_def split: withERR.splits)
lemma errMOD_igOpIPresIGWlsSTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model"
assumes "igOpIPresIGWls MOD"
shows "igOpIPresIGWlsSTR (errMOD MOD)"
by (simp add: igOpIPresIGWlsSTR_def igOpIPresIGWls_def)
(smt assms eOp_nchotomy eOp_simp1 eWlsBinp_invert
eWlsBinp_simp1 eWlsInp_invert eWlsInp_simp1 eWls_simp1 eWls_simp2 igOpIPresIGWls_def)
lemma errMOD_igConsIPresIGWlsSTR:
assumes "igConsIPresIGWls MOD" and "igWlsAllDisj MOD"
and "igWlsAbsIsInBar MOD"
shows "igConsIPresIGWlsSTR (errMOD MOD)"
using assms unfolding igConsIPresIGWls_def igConsIPresIGWlsSTR_def igWlsAllDisj_def
using
errMOD_igVarIPresIGWlsSTR[of MOD]
errMOD_igAbsIPresIGWlsSTR[of MOD]
errMOD_igOpIPresIGWlsSTR[of MOD]
by auto
lemma errMOD_igSwapIPresIGWlsSTR:
assumes "igSwapIPresIGWls MOD" and "igWlsDisj MOD"
shows "igSwapIPresIGWlsSTR (errMOD MOD)"
-using `igSwapIPresIGWls MOD`
+using \<open>igSwapIPresIGWls MOD\<close>
using assms by (fastforce simp: errMOD_def igSwapIPresIGWls_def igSwapIPresIGWlsSTR_def
igWlsDisj_def split: withERR.splits)
lemma errMOD_igSwapAbsIPresIGWlsAbsSTR:
assumes *: "igSwapAbsIPresIGWlsAbs MOD" and **: "igWlsAbsDisj MOD"
and ***: "igWlsAbsIsInBar MOD"
shows "igSwapAbsIPresIGWlsAbsSTR (errMOD MOD)"
using assms by (simp add: errMOD_def igSwapAbsIPresIGWlsAbs_def igSwapAbsIPresIGWlsAbsSTR_def
igWlsAbsIsInBar_def igWlsAbsDisj_def split: withERR.splits) blast
lemma errMOD_igSwapAllIPresIGWlsAllSTR:
assumes "igSwapAllIPresIGWlsAll MOD" and "igWlsAllDisj MOD"
and "igWlsAbsIsInBar MOD"
shows "igSwapAllIPresIGWlsAllSTR (errMOD MOD)"
using assms
unfolding igSwapAllIPresIGWlsAll_def igSwapAllIPresIGWlsAllSTR_def igWlsAllDisj_def
using errMOD_igSwapIPresIGWlsSTR[of MOD] errMOD_igSwapIPresIGWlsSTR[of MOD]
errMOD_igSwapAbsIPresIGWlsAbsSTR[of MOD]
by auto
lemma errMOD_igSubstIPresIGWlsSTR:
assumes "igSubstIPresIGWls MOD" and "igWlsDisj MOD"
shows "igSubstIPresIGWlsSTR (errMOD MOD)"
-using `igSubstIPresIGWls MOD`
+using \<open>igSubstIPresIGWls MOD\<close>
using assms by (fastforce simp: errMOD_def igSubstIPresIGWls_def igSubstIPresIGWlsSTR_def
igWlsDisj_def split: withERR.splits)
lemma errMOD_igSubstAbsIPresIGWlsAbsSTR:
assumes *: "igSubstAbsIPresIGWlsAbs MOD" and **: "igWlsAbsDisj MOD"
and ***: "igWlsAbsIsInBar MOD"
shows "igSubstAbsIPresIGWlsAbsSTR (errMOD MOD)"
using assms by (simp add: errMOD_def igSubstAbsIPresIGWlsAbs_def igSubstAbsIPresIGWlsAbsSTR_def
igWlsAbsIsInBar_def igWlsAbsDisj_def split: withERR.splits) blast
lemma errMOD_igSubstAllIPresIGWlsAllSTR:
assumes "igSubstAllIPresIGWlsAll MOD" and "igWlsAllDisj MOD"
and "igWlsAbsIsInBar MOD"
shows "igSubstAllIPresIGWlsAllSTR (errMOD MOD)"
using assms
unfolding igSubstAllIPresIGWlsAll_def igSubstAllIPresIGWlsAllSTR_def igWlsAllDisj_def
using errMOD_igSubstIPresIGWlsSTR[of MOD] errMOD_igSubstIPresIGWlsSTR[of MOD]
errMOD_igSubstAbsIPresIGWlsAbsSTR[of MOD]
by auto
-text{* The strong ``fresh" clauses are satisfied: *}
+text\<open>The strong ``fresh" clauses are satisfied:\<close>
lemma errMOD_igFreshIGVarSTR:
assumes "igVarIPresIGWls MOD" and "igFreshIGVar MOD"
shows "igFreshIGVar (errMOD MOD)"
using assms eFresh_simp1
by(fastforce simp: igVarIPresIGWls_def igFreshIGVar_def)
lemma errMOD_igFreshIGAbs1STR:
assumes *: "igAbsIPresIGWls MOD" and **: "igFreshIGAbs1 MOD"
shows "igFreshIGAbs1STR (errMOD MOD)"
unfolding igFreshIGAbs1STR_def proof(clarify)
fix ys y eX
show "eFreshAbs MOD ys y (eAbs MOD ys y eX)"
proof(cases "eX \<noteq> ERR")
define X where "X \<equiv> check eX"
case True
hence eX: "eX = OK X" unfolding X_def using OK_check by auto
show ?thesis using assms eFreshAbs_simp1 unfolding eX
by (cases "\<exists> s. isInBar (ys,s) \<and> igWls MOD s X")
(fastforce simp: igAbsIPresIGWls_def igFreshIGAbs1_def)+
qed auto
qed
lemma errMOD_igFreshIGAbs2STR:
assumes "igAbsIPresIGWls MOD" and "igFreshIGAbs2 MOD"
shows "igFreshIGAbs2STR (errMOD MOD)"
unfolding igFreshIGAbs2STR_def proof(clarify)
fix ys y xs x eX
assume *: "eFresh MOD ys y eX"
define X where "X \<equiv> check eX"
show "eFreshAbs MOD ys y (eAbs MOD xs x eX)"
proof(cases "eX \<noteq> ERR")
case True
hence eX: "eX = OK X" unfolding X_def using OK_check by auto
show ?thesis unfolding eX
using assms * eFreshAbs_invert eX
by (cases "\<exists> s. isInBar (xs,s) \<and> igWls MOD s X")
(fastforce simp: igAbsIPresIGWls_def igFreshIGAbs2_def)+
qed auto
qed
(* HERE *)
lemma errMOD_igFreshIGOpSTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model"
assumes "igOpIPresIGWls MOD" and "igFreshIGOp MOD"
shows "igFreshIGOpSTR (errMOD MOD)"
unfolding igFreshIGOpSTR_def apply clarify
subgoal for ys y delta einp ebinp
apply(cases "liftAll (\<lambda>eX. eX \<noteq> ERR) einp \<and>
liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp")
using assms by (simp_all add: igOpIPresIGWls_def igFreshIGOp_def)
(metis eFreshBinp_simp eFreshInp_simp eFresh_invert eOp_invert)+ .
lemma errMOD_igFreshClsSTR:
assumes "igConsIPresIGWls MOD" and "igFreshCls MOD"
shows "igFreshClsSTR (errMOD MOD)"
using assms unfolding igConsIPresIGWls_def igFreshCls_def igFreshClsSTR_def
using
errMOD_igFreshIGVarSTR
errMOD_igFreshIGAbs1STR errMOD_igFreshIGAbs2STR
errMOD_igFreshIGOpSTR
by auto
-text{* The strong ``swap" clauses are satisfied: *}
+text\<open>The strong ``swap" clauses are satisfied:\<close>
lemma errMOD_igSwapIGVarSTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model"
assumes "igVarIPresIGWls MOD" and "igSwapIGVar MOD"
shows "igSwapIGVar (errMOD MOD)"
using assms by (simp add: igVarIPresIGWls_def igSwapIGVar_def) (metis eSwap_simp1)
lemma errMOD_igSwapIGAbsSTR:
assumes *: "igAbsIPresIGWls MOD" and **: "igWlsDisj MOD"
and ***: "igSwapIPresIGWls MOD" and ****: "igSwapIGAbs MOD"
shows "igSwapIGAbsSTR (errMOD MOD)"
unfolding igSwapIGAbsSTR_def apply(clarify)
subgoal for zs z1 z2 xs x eX
apply (cases eX)
subgoal by auto
subgoal for X
apply(cases "\<exists> s. isInBar (xs,s) \<and> igWls MOD s X")
subgoal using assms
using assms OK_check
by (simp_all add: igAbsIPresIGWls_def igSwapIPresIGWls_def igSwapIGAbs_def igWlsDisj_def)
(smt eAbs_simp1 eSwapAbs_simp1 eSwap_simp1 withERR.inject)
subgoal using assms
by(simp_all add: igAbsIPresIGWls_def igSwapIPresIGWls_def igSwapIGAbs_def igWlsDisj_def)
(metis check_OK eAbs_nchotomy eSwap_invert) . . .
lemma errMOD_igSwapIGOpSTR:
assumes "igWlsAbsIsInBar MOD" and "igOpIPresIGWls MOD"
and "igSwapIPresIGWls MOD" and "igSwapAbsIPresIGWlsAbs MOD"
and "igWlsDisj MOD" and "igWlsAbsDisj MOD"
and "igSwapIGOp MOD"
shows "igSwapIGOpSTR (errMOD MOD)"
unfolding igSwapIGOpSTR_def proof(clarify)
have 0: "igSwapInpIPresIGWlsInp MOD \<and> igSwapBinpIPresIGWlsBinp MOD"
- using `igSwapIPresIGWls MOD` `igSwapAbsIPresIGWlsAbs MOD`
+ using \<open>igSwapIPresIGWls MOD\<close> \<open>igSwapAbsIPresIGWlsAbs MOD\<close>
imp_igSwapInpIPresIGWlsInp imp_igSwapBinpIPresIGWlsBinp by auto
have "igSwapIPresIGWlsSTR (errMOD MOD) \<and>
igSwapAbsIPresIGWlsAbsSTR (errMOD MOD)"
using assms errMOD_igSwapIPresIGWlsSTR
errMOD_igSwapAbsIPresIGWlsAbsSTR by auto
hence 1: "igSwapInpIPresIGWlsInpSTR (errMOD MOD) \<and>
igSwapBinpIPresIGWlsBinpSTR (errMOD MOD)"
using imp_igSwapInpIPresIGWlsInpSTR
imp_igSwapBinpIPresIGWlsBinpSTR by auto
fix zs::'varSort and z1 z2 ::'var and delta einp ebinp
let ?Left = "eSwap MOD zs z1 z2 (eOp MOD delta einp ebinp)"
let ?einpsw = "eSwapInp MOD zs z1 z2 einp"
let ?ebinpsw = "eSwapBinp MOD zs z1 z2 ebinp"
let ?Right = "eOp MOD delta ?einpsw ?ebinpsw"
show "?Left = ?Right"
proof(cases "liftAll (\<lambda>eX. eX \<noteq> ERR) einp \<and>
liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp")
case True note t = True
moreover obtain inp and binp where
"inp = checkI einp" and "binp = checkI ebinp" by blast
ultimately have einp: "einp = OKI inp" "ebinp = OKI binp" by auto
show ?thesis
proof(cases "igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp")
case False
hence "?Left = ERR" unfolding einp by auto
have "\<not> (eWlsInp MOD delta einp \<and> eWlsBinp MOD delta ebinp)"
unfolding einp using False by auto
hence 2: "\<not> (eWlsInp MOD delta ?einpsw \<and> eWlsBinp MOD delta ?ebinpsw)"
using 1 unfolding igSwapInpIPresIGWlsInpSTR_def
igSwapBinpIPresIGWlsBinpSTR_def by auto
{fix X assume "?Right = OK X"
then obtain inpsw and binpsw
where "?einpsw = OKI inpsw" and "?ebinpsw = OKI binpsw"
and "igWlsInp MOD delta inpsw" and "igWlsBinp MOD delta binpsw"
and "X = igOp MOD delta inpsw binpsw"
using eOp_invert[of MOD delta ?einpsw ?ebinpsw X] by auto
hence False using 2 by auto
}
- with `?Left = ERR` show ?thesis by (cases ?Right) auto
+ with \<open>?Left = ERR\<close> show ?thesis by (cases ?Right) auto
next
case True
moreover have "igWls MOD (stOf delta) (igOp MOD delta inp binp)"
- using True `igOpIPresIGWls MOD` unfolding igOpIPresIGWls_def by simp
+ using True \<open>igOpIPresIGWls MOD\<close> unfolding igOpIPresIGWls_def by simp
moreover have "igWlsInp MOD delta (igSwapInp MOD zs z1 z2 inp) \<and>
igWlsBinp MOD delta (igSwapBinp MOD zs z1 z2 binp)"
using 0 unfolding igSwapInpIPresIGWlsInp_def igSwapBinpIPresIGWlsBinp_def
using True by simp
- ultimately show ?thesis using `igSwapIGOp MOD` unfolding einp igSwapIGOp_def by auto
+ ultimately show ?thesis using \<open>igSwapIGOp MOD\<close> unfolding einp igSwapIGOp_def by auto
qed
qed auto
qed
lemma errMOD_igSwapClsSTR:
assumes "igWlsAllDisj MOD" and "igWlsDisj MOD"
and "igWlsAbsIsInBar MOD" and "igConsIPresIGWls MOD"
and "igSwapAllIPresIGWlsAll MOD" and "igSwapCls MOD"
shows "igSwapClsSTR (errMOD MOD)"
using assms
unfolding igWlsAllDisj_def igConsIPresIGWls_def igSwapCls_def
igSwapAllIPresIGWlsAll_def igSwapClsSTR_def
using
errMOD_igSwapIGVarSTR[of MOD]
errMOD_igSwapIGAbsSTR[of MOD]
errMOD_igSwapIGOpSTR[of MOD]
by simp
-text{* The strong ``subst" clauses are satisfied: *}
+text\<open>The strong ``subst" clauses are satisfied:\<close>
lemma errMOD_igSubstIGVar1STR:
assumes "igVarIPresIGWls MOD" and "igSubstIGVar1 MOD"
shows "igSubstIGVar1STR (errMOD MOD)"
using assms
by (simp add: igSubstIGVar1STR_def igVarIPresIGWls_def igSubstIGVar1_def)
(metis eSubst_simp1 eWls_invert)
lemma errMOD_igSubstIGVar2STR:
assumes "igVarIPresIGWls MOD" and "igSubstIGVar2 MOD"
shows "igSubstIGVar2STR (errMOD MOD)"
using assms
by (simp add: igSubstIGVar2STR_def igVarIPresIGWls_def igSubstIGVar2_def)
(metis eSubst_simp1 eWls_invert)
lemma errMOD_igSubstIGAbsSTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model"
assumes *: "igAbsIPresIGWls MOD" and **: "igWlsDisj MOD"
and ***: "igSubstIPresIGWls MOD" and ****: "igSubstIGAbs MOD"
shows "igSubstIGAbsSTR (errMOD MOD)"
unfolding igSubstIGAbsSTR_def proof(clarify)
fix ys xs ::'varSort and y x ::'var and eX eY
assume diff: "xs \<noteq> ys \<or> x \<noteq> y"
and x_fresh_Y: "eFresh MOD xs x eY"
show "eSubstAbs MOD ys eY y (eAbs MOD xs x eX) =
eAbs MOD xs x (eSubst MOD ys eY y eX)"
proof(cases "eX \<noteq> ERR \<and> eY \<noteq> ERR")
case True
define X and Y where "X \<equiv> check eX" and "Y \<equiv> check eY"
hence eX: "eX = OK X" and eY: "eY = OK Y" unfolding X_def Y_def
using True OK_check by auto
show ?thesis
proof(cases "(\<exists> s. isInBar (xs,s) \<and> igWls MOD s X) \<and> igWls MOD (asSort ys) Y")
case True
then obtain s where xs_s: "isInBar (xs, s)" and X: "igWls MOD s X"
and Y: "igWls MOD (asSort ys) Y" by auto
hence "igWlsAbs MOD (xs,s) (igAbs MOD xs x X)"
using * unfolding igAbsIPresIGWls_def by simp
moreover have "igWls MOD s (igSubst MOD ys Y y X)"
using X Y *** unfolding igSubstIPresIGWls_def by simp
moreover have "igFresh MOD xs x Y"
using x_fresh_Y Y unfolding eY by simp
ultimately show ?thesis unfolding eX eY
using xs_s X Y apply simp
using x_fresh_Y diff **** unfolding igSubstIGAbs_def by fastforce
next
case False
show ?thesis
proof(cases "(EX s. igWls MOD s X) \<and> igWls MOD (asSort ys) Y")
case True
then obtain s where X: "igWls MOD s X" and Y: "igWls MOD (asSort ys) Y" by auto
hence 2: "~ isInBar (xs,s)" using False by (auto simp: eX eY)
let ?Xsb = "igSubst MOD ys Y y X"
have Xsb: "igWls MOD s ?Xsb"
using Y X *** unfolding igSubstIPresIGWls_def by auto
{fix s' assume 3: "isInBar (xs,s')" and "igWls MOD s' ?Xsb"
hence "s = s'" using Xsb ** unfolding igWlsDisj_def by auto
hence False using 2 3 by (simp add: eX eY)
}
thus ?thesis using False Y eAbs_simp2 X eX eY by fastforce
qed(auto simp add: eX eY)
qed
qed auto
qed
lemma errMOD_igSubstIGOpSTR:
assumes "igWlsAbsIsInBar MOD"
and "igVarIPresIGWls MOD" and "igOpIPresIGWls MOD"
and "igSubstIPresIGWls MOD" and "igSubstAbsIPresIGWlsAbs MOD"
and "igWlsDisj MOD" and "igWlsAbsDisj MOD"
and "igSubstIGOp MOD"
shows "igSubstIGOpSTR (errMOD MOD)"
proof-
have 0: "igSubstInpIPresIGWlsInp MOD \<and> igSubstBinpIPresIGWlsBinp MOD"
- using `igSubstIPresIGWls MOD` `igSubstAbsIPresIGWlsAbs MOD`
+ using \<open>igSubstIPresIGWls MOD\<close> \<open>igSubstAbsIPresIGWlsAbs MOD\<close>
imp_igSubstInpIPresIGWlsInp imp_igSubstBinpIPresIGWlsBinp by auto
have "igSubstIPresIGWlsSTR (errMOD MOD) \<and> igSubstAbsIPresIGWlsAbsSTR (errMOD MOD)"
using assms errMOD_igSubstIPresIGWlsSTR errMOD_igSubstAbsIPresIGWlsAbsSTR by auto
hence 1: "igSubstInpIPresIGWlsInpSTR (errMOD MOD) \<and>
igSubstBinpIPresIGWlsBinpSTR (errMOD MOD)"
using imp_igSubstInpIPresIGWlsInpSTR imp_igSubstBinpIPresIGWlsBinpSTR by auto
show ?thesis
unfolding igSubstIGOpSTR_def proof safe
fix ys::'varSort and y y1 ::'var and delta einp ebinp
let ?Left = "eSubst MOD ys (eVar MOD ys y1) y (eOp MOD delta einp ebinp)"
let ?einpsb = "eSubstInp MOD ys (eVar MOD ys y1) y einp"
let ?ebinpsb = "eSubstBinp MOD ys (eVar MOD ys y1) y ebinp"
let ?Right = "eOp MOD delta ?einpsb ?ebinpsb"
show "?Left = ?Right"
proof(cases "liftAll (\<lambda>eX. eX \<noteq> ERR) einp \<and> liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp")
case True
moreover obtain inp binp where
"inp = checkI einp" and "binp = checkI ebinp" by blast
ultimately have einp: "einp = OKI inp" "ebinp = OKI binp" by auto
have igWls_y1: "igWls MOD (asSort ys) (igVar MOD ys y1)"
- using `igVarIPresIGWls MOD` unfolding igVarIPresIGWls_def by simp
+ using \<open>igVarIPresIGWls MOD\<close> unfolding igVarIPresIGWls_def by simp
show ?thesis
proof(cases "igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp")
case False
hence "?Left = ERR" unfolding einp by auto
have "\<not> (eWlsInp MOD delta einp \<and> eWlsBinp MOD delta ebinp)"
unfolding einp using False by simp
hence 2: "\<not> (eWlsInp MOD delta ?einpsb \<and> eWlsBinp MOD delta ?ebinpsb)"
using igWls_y1 1
unfolding igSubstInpIPresIGWlsInpSTR_def igSubstBinpIPresIGWlsBinpSTR_def by simp
{fix X assume "?Right = OK X"
then obtain inpsb binpsb where
"?einpsb = OKI inpsb" and "?ebinpsb = OKI binpsb"
and "igWlsInp MOD delta inpsb" and "igWlsBinp MOD delta binpsb"
and "X = igOp MOD delta inpsb binpsb"
using eOp_invert[of MOD delta ?einpsb ?ebinpsb X] by auto
hence False using 2 by auto
}
hence "?Right = ERR" by (cases ?Right, auto)
- with `?Left = ERR` show ?thesis by simp
+ with \<open>?Left = ERR\<close> show ?thesis by simp
next
case True
moreover have "igWls MOD (stOf delta) (igOp MOD delta inp binp)"
- using True `igOpIPresIGWls MOD` unfolding igOpIPresIGWls_def by simp
+ using True \<open>igOpIPresIGWls MOD\<close> unfolding igOpIPresIGWls_def by simp
moreover
have "igWlsInp MOD delta (igSubstInp MOD ys (igVar MOD ys y1) y inp) \<and>
igWlsBinp MOD delta (igSubstBinp MOD ys (igVar MOD ys y1) y binp)"
using 0 unfolding igSubstInpIPresIGWlsInp_def igSubstBinpIPresIGWlsBinp_def
using True igWls_y1 by simp
ultimately show ?thesis
- using `igSubstIGOp MOD` igWls_y1 unfolding einp igSubstIGOp_def by auto
+ using \<open>igSubstIGOp MOD\<close> igWls_y1 unfolding einp igSubstIGOp_def by auto
qed
qed auto
next
fix ys::'varSort and y ::'var and eY delta einp ebinp
assume eY: "eWls MOD (asSort ys) eY"
let ?Left = "eSubst MOD ys eY y (eOp MOD delta einp ebinp)"
let ?einpsb = "eSubstInp MOD ys eY y einp"
let ?ebinpsb = "eSubstBinp MOD ys eY y ebinp"
let ?Right = "eOp MOD delta ?einpsb ?ebinpsb"
from eY obtain Y where eY_def: "eY = OK Y"
and Y: "igWls MOD (asSort ys) Y" using eWls_invert[of MOD "asSort ys" eY] by auto
show "?Left = ?Right"
proof(cases "liftAll (\<lambda>eX. eX \<noteq> ERR) einp \<and> liftAll (\<lambda>eA. eA \<noteq> ERR) ebinp")
case True
moreover obtain inp binp where
"inp = checkI einp" and "binp = checkI ebinp" by blast
ultimately have einp: "einp = OKI inp" "ebinp = OKI binp" by auto
show ?thesis
proof(cases "igWlsInp MOD delta inp \<and> igWlsBinp MOD delta binp")
case False
hence "?Left = ERR" unfolding einp by auto
have "\<not> (eWlsInp MOD delta einp \<and> eWlsBinp MOD delta ebinp)"
unfolding einp using False by simp
hence 2: "\<not> (eWlsInp MOD delta ?einpsb \<and> eWlsBinp MOD delta ?ebinpsb)"
unfolding eY_def using Y 1
unfolding igSubstInpIPresIGWlsInpSTR_def igSubstBinpIPresIGWlsBinpSTR_def by simp
{fix X assume "?Right = OK X"
then obtain inpsb binpsb
where "?einpsb = OKI inpsb" and "?ebinpsb = OKI binpsb"
and "igWlsInp MOD delta inpsb" and "igWlsBinp MOD delta binpsb"
and "X = igOp MOD delta inpsb binpsb"
using eOp_invert[of MOD delta ?einpsb ?ebinpsb X] by auto
hence False using 2 by auto
}
hence "?Right = ERR" by (cases ?Right, auto)
- with `?Left = ERR` show ?thesis by simp
+ with \<open>?Left = ERR\<close> show ?thesis by simp
next
case True
moreover have "igWls MOD (stOf delta) (igOp MOD delta inp binp)"
- using True `igOpIPresIGWls MOD` unfolding igOpIPresIGWls_def by simp
+ using True \<open>igOpIPresIGWls MOD\<close> unfolding igOpIPresIGWls_def by simp
moreover
have "igWlsInp MOD delta (igSubstInp MOD ys Y y inp) \<and>
igWlsBinp MOD delta (igSubstBinp MOD ys Y y binp)"
using 0 unfolding igSubstInpIPresIGWlsInp_def igSubstBinpIPresIGWlsBinp_def
using True Y by simp
ultimately show ?thesis unfolding einp eY_def
- using `igSubstIGOp MOD` Y unfolding igSubstIGOp_def by auto
+ using \<open>igSubstIGOp MOD\<close> Y unfolding igSubstIGOp_def by auto
qed
qed auto
qed
qed
lemma errMOD_igSubstClsSTR:
assumes "igWlsAllDisj MOD" and "igConsIPresIGWls MOD"
and "igWlsAbsIsInBar MOD"
and "igSubstAllIPresIGWlsAll MOD" and "igSubstCls MOD"
shows "igSubstClsSTR (errMOD MOD)"
using assms
unfolding igWlsAllDisj_def igConsIPresIGWls_def igSubstCls_def
igSubstAllIPresIGWlsAll_def igSubstClsSTR_def
using
errMOD_igSubstIGVar1STR[of MOD] errMOD_igSubstIGVar2STR[of MOD]
errMOD_igSubstIGAbsSTR[of MOD]
errMOD_igSubstIGOpSTR[of MOD]
by simp
-text{* Strong swap-based congruence for abstractions holds: *}
+text\<open>Strong swap-based congruence for abstractions holds:\<close>
lemma errMOD_igAbsCongSSTR:
assumes "igSwapIPresIGWls MOD" and "igWlsDisj MOD" and "igAbsCongS MOD"
shows "igAbsCongSSTR (errMOD MOD)"
unfolding igAbsCongSSTR_def proof(clarify)
fix xs ::'varSort and x x' y ::'var and eX eX'
assume *: "eFresh MOD xs y eX" and **: "eFresh MOD xs y eX'"
and ***: "eSwap MOD xs y x eX = eSwap MOD xs y x' eX'"
let ?phi = "\<lambda>eX. eX = ERR \<or> (\<exists> X. eX = OK X \<and> (\<forall> s. \<not> igWls MOD s X))"
have 1: "?phi eX = ?phi eX'"
proof
assume "?phi eX"
{fix X' s' assume "eX' = OK X' \<and> (\<exists> s. igWls MOD s X')"
- hence "ERR = OK (igSwap MOD xs y x' X')" using `?phi eX` *** by auto
+ hence "ERR = OK (igSwap MOD xs y x' X')" using \<open>?phi eX\<close> *** by auto
}
thus "?phi eX'" by(cases eX', auto)
next
assume "?phi eX'"
{fix X assume "eX = OK X \<and> (\<exists> s. igWls MOD s X)"
- hence "ERR = OK (igSwap MOD xs y x X)" using `?phi eX'` *** by auto
+ hence "ERR = OK (igSwap MOD xs y x X)" using \<open>?phi eX'\<close> *** by auto
}
thus "?phi eX" by(cases eX, auto)
qed
show "eAbs MOD xs x eX = eAbs MOD xs x' eX'"
proof(cases "?phi eX")
case True
thus ?thesis using 1 by auto
next
case False
then obtain s X where eX: "eX = OK X" and X_wls: "igWls MOD s X" by(cases eX, auto)
obtain s' X' where eX': "eX' = OK X'" and X'_wls: "igWls MOD s' X'"
- using `\<not> ?phi eX` 1 by(cases eX') auto
+ using \<open>\<not> ?phi eX\<close> 1 by(cases eX') auto
hence "igSwap MOD xs y x X = igSwap MOD xs y x' X'"
using eX X_wls *** by auto
moreover have "igWls MOD s (igSwap MOD xs y x X)"
- using X_wls `igSwapIPresIGWls MOD` unfolding igSwapIPresIGWls_def by simp
+ using X_wls \<open>igSwapIPresIGWls MOD\<close> unfolding igSwapIPresIGWls_def by simp
moreover have "igWls MOD s' (igSwap MOD xs y x' X')"
- using X'_wls `igSwapIPresIGWls MOD` unfolding igSwapIPresIGWls_def by simp
- ultimately have "s' = s" using `igWlsDisj MOD` unfolding igWlsDisj_def by auto
+ using X'_wls \<open>igSwapIPresIGWls MOD\<close> unfolding igSwapIPresIGWls_def by simp
+ ultimately have "s' = s" using \<open>igWlsDisj MOD\<close> unfolding igWlsDisj_def by auto
show ?thesis
proof (cases "isInBar (xs,s)")
case True
have "igFresh MOD xs y X" using * X_wls unfolding eX by simp
moreover have "igFresh MOD xs y X'" using ** X'_wls unfolding eX' by simp
moreover have "igSwap MOD xs y x X = igSwap MOD xs y x' X'"
using *** X_wls X'_wls unfolding eX eX' by simp
ultimately show ?thesis
unfolding eX eX'
- using X_wls X'_wls unfolding `s' = s`
- using `igAbsCongS MOD` True unfolding igAbsCongS_def
+ using X_wls X'_wls unfolding \<open>s' = s\<close>
+ using \<open>igAbsCongS MOD\<close> True unfolding igAbsCongS_def
by (metis FixSyn.eCons_simps(2) FixSyn_axioms)
next
case False
{fix s'' assume xs_s'': "isInBar (xs,s'')" and "igWls MOD s'' X"
- hence "s = s''" using X_wls `igWlsDisj MOD` unfolding igWlsDisj_def by auto
+ hence "s = s''" using X_wls \<open>igWlsDisj MOD\<close> unfolding igWlsDisj_def by auto
hence False using False xs_s'' by simp
}
moreover
{fix s'' assume xs_s'': "isInBar (xs,s'')" and "igWls MOD s'' X'"
- hence "s = s''" using X'_wls `igWlsDisj MOD` unfolding igWlsDisj_def `s' = s` by auto
+ hence "s = s''" using X'_wls \<open>igWlsDisj MOD\<close> unfolding igWlsDisj_def \<open>s' = s\<close> by auto
hence False using False xs_s'' by simp
}
ultimately show ?thesis
- using eX eX' X_wls X'_wls unfolding `s' = s` by fastforce
+ using eX eX' X_wls X'_wls unfolding \<open>s' = s\<close> by fastforce
qed
qed
qed
-text{* The renaming clause for abstractions holds: *}
+text\<open>The renaming clause for abstractions holds:\<close>
lemma errMOD_igAbsRenSTR:
assumes "igVarIPresIGWls MOD" and "igSubstIPresIGWls MOD"
and "igWlsDisj MOD" and "igAbsRen MOD"
shows "igAbsRenSTR (errMOD MOD)"
using assms unfolding igAbsRenSTR_def apply clarify
subgoal for xs y x eX
apply(cases eX)
subgoal by auto
subgoal for X
apply(cases "EX s. isInBar (xs,s) \<and> igWls MOD s X")
subgoal by (auto simp: igVarIPresIGWls_def igSubstIPresIGWls_def igAbsRen_def)
subgoal using assms by (simp add: igVarIPresIGWls_def igSubstIPresIGWls_def igAbsRen_def igWlsDisj_def)
(metis eAbs_simp2 eAbs_simp3 eSubst_simp1 eSubst_simp3) . . .
-text{* Strong subst-based congruence for abstractions holds: *}
+text\<open>Strong subst-based congruence for abstractions holds:\<close>
corollary errMOD_igAbsCongUSTR:
assumes "igVarIPresIGWls MOD" and "igSubstIPresIGWls MOD"
and "igWlsDisj MOD" and "igAbsRen MOD"
shows "igAbsCongUSTR (errMOD MOD)"
using assms errMOD_igAbsRenSTR igAbsRenSTR_imp_igAbsCongUSTR by auto
-text{* The error model is a strongly well-sorted fresh-swap model: *}
+text\<open>The error model is a strongly well-sorted fresh-swap model:\<close>
lemma errMOD_iwlsFSwSTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs) model"
assumes "iwlsFSw MOD"
shows "iwlsFSwSTR (errMOD MOD)"
using assms unfolding iwlsFSw_def iwlsFSwSTR_def
using errMOD_igWlsAllDisj[of MOD]
errMOD_igWlsAbsIsInBar[of MOD]
errMOD_igConsIPresIGWlsSTR[of MOD]
errMOD_igSwapAllIPresIGWlsAllSTR[of MOD]
errMOD_igFreshClsSTR[of MOD] errMOD_igSwapClsSTR[of MOD]
errMOD_igAbsCongSSTR[of MOD]
apply simp
unfolding igSwapAllIPresIGWlsAll_def igWlsAllDisj_defs by simp
-text{* The error model is a strongly well-sorted fresh-subst model: *}
+text\<open>The error model is a strongly well-sorted fresh-subst model:\<close>
lemma errMOD_iwlsFSbSwTR:
fixes MOD :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs) model"
assumes "iwlsFSb MOD"
shows "iwlsFSbSwTR (errMOD MOD)"
using assms unfolding iwlsFSb_def iwlsFSbSwTR_def
using errMOD_igWlsAllDisj[of MOD]
errMOD_igWlsAbsIsInBar[of MOD]
errMOD_igConsIPresIGWlsSTR[of MOD]
errMOD_igSubstAllIPresIGWlsAllSTR[of MOD]
errMOD_igFreshClsSTR[of MOD] errMOD_igSubstClsSTR[of MOD]
errMOD_igAbsRenSTR[of MOD]
by (simp add: igConsIPresIGWls_def igSubstAllIPresIGWlsAll_def igWlsAllDisj_defs)
-subsubsection {* The natural morhpism from an error model to its original model *}
-
-text{* This morphism is igiven by the ``check" functions. *}
-
-text{* Preservation of the domains: *}
+subsubsection \<open>The natural morhpism from an error model to its original model\<close>
+
+text\<open>This morphism is igiven by the ``check" functions.\<close>
+
+text\<open>Preservation of the domains:\<close>
lemma check_ipresIGWls:
"ipresIGWls check (errMOD MOD) MOD"
unfolding ipresIGWls_def apply clarify
subgoal for _ X by(cases X) auto .
lemma check_ipresIGWlsAbs:
"ipresIGWlsAbs check (errMOD MOD) MOD"
unfolding ipresIGWlsAbs_def apply clarify
subgoal for _ _ A by(cases A) auto .
lemma check_ipresIGWlsAll:
"ipresIGWlsAll check check (errMOD MOD) MOD"
unfolding ipresIGWlsAll_def
using check_ipresIGWls check_ipresIGWlsAbs by auto
-text{* Preservation of the operations: *}
+text\<open>Preservation of the operations:\<close>
lemma check_ipresIGVar:
"ipresIGVar check (errMOD MOD) MOD"
unfolding ipresIGVar_def by simp
lemma check_ipresIGAbs:
"ipresIGAbs check check (errMOD MOD) MOD"
unfolding ipresIGAbs_def apply clarify
subgoal for _ _ _ X by(cases X) auto .
lemma check_ipresIGOp:
"ipresIGOp check check (errMOD MOD) MOD"
unfolding ipresIGOp_def proof clarify
fix delta einp ebinp
assume "eWlsInp MOD delta einp" and "eWlsBinp MOD delta ebinp"
then obtain inp binp where
"igWlsInp MOD delta inp" and "igWlsBinp MOD delta binp"
and "einp = OKI inp" and "ebinp = OKI binp"
using eWlsInp_invert eWlsBinp_invert by blast
hence "check (eOp MOD delta einp ebinp) =
igOp MOD delta (checkI einp) (checkI ebinp)" by simp
thus "check (eOp MOD delta einp ebinp) =
igOp MOD delta (lift check einp) (lift check ebinp)"
unfolding checkI_def .
qed
lemma check_ipresIGCons:
"ipresIGCons check check (errMOD MOD) MOD"
unfolding ipresIGCons_def
using
check_ipresIGVar
check_ipresIGAbs
check_ipresIGOp
by auto
lemma check_ipresIGFresh:
"ipresIGFresh check (errMOD MOD) MOD"
unfolding ipresIGFresh_def apply clarify
subgoal for _ _ _ X by(cases X) auto .
lemma check_ipresIGFreshAbs:
"ipresIGFreshAbs check (errMOD MOD) MOD"
unfolding ipresIGFreshAbs_def apply clarify
subgoal for _ _ _ _ A by(cases A) auto .
lemma check_ipresIGFreshAll:
"ipresIGFreshAll check check (errMOD MOD) MOD"
unfolding ipresIGFreshAll_def
using check_ipresIGFresh check_ipresIGFreshAbs by auto
lemma check_ipresIGSwap:
"ipresIGSwap check (errMOD MOD) MOD"
unfolding ipresIGSwap_def apply clarify
subgoal for _ _ _ _ X by(cases X) auto .
lemma check_ipresIGSwapAbs:
"ipresIGSwapAbs check (errMOD MOD) MOD"
unfolding ipresIGSwapAbs_def apply clarify
subgoal for _ _ _ _ _ A by(cases A) auto .
lemma check_ipresIGSwapAll:
"ipresIGSwapAll check check (errMOD MOD) MOD"
unfolding ipresIGSwapAll_def
using check_ipresIGSwap check_ipresIGSwapAbs by auto
lemma check_ipresIGSubst:
"ipresIGSubst check (errMOD MOD) MOD"
unfolding ipresIGSubst_def apply clarify
subgoal for _ Y _ _ X by (cases X, simp, cases Y) auto .
lemma check_ipresIGSubstAbs:
"ipresIGSubstAbs check check (errMOD MOD) MOD"
unfolding ipresIGSubstAbs_def apply clarify
subgoal for _ Y _ _ _ A by (cases A, simp, cases Y) auto .
lemma check_ipresIGSubstAll:
"ipresIGSubstAll check check (errMOD MOD) MOD"
unfolding ipresIGSubstAll_def
using check_ipresIGSubst check_ipresIGSubstAbs by auto
-text{* ``check" is a fresh-swap morphism: *}
+text\<open>``check" is a fresh-swap morphism:\<close>
lemma check_FSwImorph:
"FSwImorph check check (errMOD MOD) MOD"
unfolding FSwImorph_def
using check_ipresIGWlsAll check_ipresIGCons
check_ipresIGFreshAll check_ipresIGSwapAll by auto
-text{* ``check" is a fresh-subst morphism: *}
+text\<open>``check" is a fresh-subst morphism:\<close>
lemma check_FSbImorph:
"FSbImorph check check (errMOD MOD) MOD"
unfolding FSbImorph_def
using check_ipresIGWlsAll check_ipresIGCons
check_ipresIGFreshAll check_ipresIGSubstAll by auto
-subsection {* Initiality of the models of terms *}
-
-text {* We show that terms form initial models in all the considered kinds.
+subsection \<open>Initiality of the models of terms\<close>
+
+text \<open>We show that terms form initial models in all the considered kinds.
The desired initial morphism will be the composition of ``check" with the
factorization of the standard (absolute-initial) function from quasi-terms, ``qInit",
to alpha-equivalence.
``qInit" preserving alpha-equivalence (in an unsorted fashion)
-was the main reason for introducing error models. *}
+was the main reason for introducing error models.\<close>
(* Here we need to switch back for a while to the quasi-term ``implementation" of terms: *)
declare qItem_simps[simp]
declare qItem_versus_item_simps[simp]
declare good_item_simps[simp]
-subsubsection {* The initial map from quasi-terms to a strong model *}
+subsubsection \<open>The initial map from quasi-terms to a strong model\<close>
(* The next is needed in the termination arigument for ``qInit": *)
definition
aux_qInit_ignoreFirst ::
"('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model *
('index,'bindex,'varSort,'var,'opSym)qTerm +
('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model *
('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTermItem"
where
"aux_qInit_ignoreFirst K =
(case K of Inl (MOD,qX) \<Rightarrow> termIn qX
|Inr (MOD,qA) \<Rightarrow> absIn qA)"
lemma qTermLess_ingoreFirst_wf:
"wf (inv_image qTermLess aux_qInit_ignoreFirst)"
using qTermLess_wf wf_inv_image by auto
function
qInit :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> 'gTerm"
and
qInitAbs :: "('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> 'gAbs"
where
"qInit MOD (qVar xs x) = igVar MOD xs x"
|
"qInit MOD (qOp delta qinp qbinp) =
igOp MOD delta (lift (qInit MOD) qinp) (lift (qInitAbs MOD) qbinp)"
|
"qInitAbs MOD (qAbs xs x qX) = igAbs MOD xs x (qInit MOD qX)"
by(pat_completeness) auto
termination
apply(relation "inv_image qTermLess aux_qInit_ignoreFirst")
apply(simp add: qTermLess_ingoreFirst_wf)
by(auto simp: qTermLess_def aux_qInit_ignoreFirst_def)
lemma qFreshAll_igFreshAll_qInitAll:
assumes "igFreshClsSTR MOD"
shows
"(qFresh ys y qX \<longrightarrow> igFresh MOD ys y (qInit MOD qX)) \<and>
(qFreshAbs ys y qA \<longrightarrow> igFreshAbs MOD ys y (qInitAbs MOD qA))"
apply(induct rule: qTerm_rawInduct)
using assms
by (auto simp: igFreshClsSTR_def igFreshIGVar_def qFreshInp_def qFreshBinp_def liftAll_lift_comp
liftAll_def igFreshInp_def igFreshBinp_def lift_def igFreshIGAbs1STR_def igFreshIGAbs2STR_def igFreshIGOpSTR_def
split: option.splits)
corollary iwlsFSwSTR_qFreshAll_igFreshAll_qInitAll:
assumes "iwlsFSwSTR MOD"
shows
"(qFresh ys y qX \<longrightarrow> igFresh MOD ys y (qInit MOD qX)) \<and>
(qFreshAbs ys y qA \<longrightarrow> igFreshAbs MOD ys y (qInitAbs MOD qA))"
using assms unfolding iwlsFSwSTR_def by(simp add: qFreshAll_igFreshAll_qInitAll)
corollary iwlsFSbSwTR_qFreshAll_igFreshAll_qInitAll:
assumes "iwlsFSbSwTR MOD"
shows
"(qFresh ys y qX \<longrightarrow> igFresh MOD ys y (qInit MOD qX)) \<and>
(qFreshAbs ys y qA \<longrightarrow> igFreshAbs MOD ys y (qInitAbs MOD qA))"
using assms unfolding iwlsFSbSwTR_def by(simp add: qFreshAll_igFreshAll_qInitAll)
lemma qSwapAll_igSwapAll_qInitAll:
assumes "igSwapClsSTR MOD"
shows
"qInit MOD (qX #[[ z1 \<and> z2]]_zs) = igSwap MOD zs z1 z2 (qInit MOD qX) \<and>
qInitAbs MOD (qA $[[z1 \<and> z2]]_zs) = igSwapAbs MOD zs z1 z2 (qInitAbs MOD qA)"
proof(induction rule: qTerm_rawInduct)
case (Var xs x)
then show ?case using assms unfolding igSwapClsSTR_def igSwapIGVar_def by simp
next
case (Op delta qinp qbinp)
hence "lift (qInit MOD) (qSwapInp zs z1 z2 qinp) =
igSwapInp MOD zs z1 z2 (lift (qInit MOD) qinp) \<and>
lift (qInitAbs MOD) (qSwapBinp zs z1 z2 qbinp) =
igSwapBinp MOD zs z1 z2 (lift (qInitAbs MOD) qbinp)"
using Op.IH by (auto simp: qSwapInp_def qSwapBinp_def igSwapInp_def lift_def liftAll_def
igSwapBinp_def iwlsFSwSTR_def igSwapClsSTR_def igSwapIGOpSTR_def
split: option.splits)
thus ?case
using assms unfolding iwlsFSwSTR_def igSwapClsSTR_def igSwapIGOpSTR_def by simp
next
case (Abs xs x X)
then show ?case using assms unfolding igSwapClsSTR_def igSwapIGAbsSTR_def by simp
qed
corollary iwlsFSwSTR_qSwapAll_igSwapAll_qInitAll:
assumes wls: "iwlsFSwSTR MOD"
shows
"qInit MOD (qX #[[ z1 \<and> z2]]_zs) = igSwap MOD zs z1 z2 (qInit MOD qX) \<and>
qInitAbs MOD (qA $[[z1 \<and> z2]]_zs) = igSwapAbs MOD zs z1 z2 (qInitAbs MOD qA)"
using assms unfolding iwlsFSwSTR_def by(simp add: qSwapAll_igSwapAll_qInitAll)
lemma qSwapAll_igSubstAll_qInitAll:
fixes qX::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
qA::"('index,'bindex,'varSort,'var,'opSym)qAbs"
assumes *: "igSubstClsSTR MOD" and "igFreshClsSTR MOD"
and "igAbsRenSTR MOD"
shows
"(qGood qX \<longrightarrow>
(\<forall> ys y1 y.
qAFresh ys y1 qX \<longrightarrow>
qInit MOD (qX #[[y1 \<and> y]]_ys) = igSubst MOD ys (igVar MOD ys y1) y (qInit MOD qX)))
\<and>
(qGoodAbs qA \<longrightarrow>
(\<forall> ys y1 y.
qAFreshAbs ys y1 qA \<longrightarrow>
qInitAbs MOD (qA $[[y1 \<and> y]]_ys) = igSubstAbs MOD ys (igVar MOD ys y1) y (qInitAbs MOD qA)))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
show ?case apply safe
subgoal for ys y1 y using *
by (cases "ys = xs \<and> y = x")
(auto simp: igSubstClsSTR_defs igSubstIGVar2STR_def igSubstClsSTR_defs igSubstIGVar1STR_def).
next
let ?h = "qInit MOD" let ?hA = "qInitAbs MOD"
case (Op delta qinp qbinp)
then show ?case proof safe
fix ys y1 y
assume ***: "qAFresh ys y1 (qOp delta qinp qbinp)"
have "lift ?h (qSwapInp ys y1 y qinp) =
igSubstInp MOD ys (igVar MOD ys y1) y (lift ?h qinp) \<and>
lift ?hA (qSwapBinp ys y1 y qbinp) =
igSubstBinp MOD ys (igVar MOD ys y1) y (lift ?hA qbinp)"
using Op.IH ***
by (auto simp: qSwapInp_def igSubstInp_def qSwapBinp_def igSubstBinp_def
lift_def liftAll_def split: option.splits)
thus "qInit MOD (qOp delta qinp qbinp #[[y1 \<and> y]]_ys) =
igSubst MOD ys (igVar MOD ys y1) y (qInit MOD (qOp delta qinp qbinp))"
using assms unfolding iwlsFSwSTR_def igSubstClsSTR_defs igSubstIGOpSTR_def by simp
qed
next
let ?h = "qInit MOD" let ?hA = "qInitAbs MOD"
case (Abs xs x qX)
show ?case proof safe
fix ys y1 y
let ?xy1y = "x @xs[y1 \<and> y]_ys" let ?y1 = "igVar MOD ys y1"
assume "qAFreshAbs ys y1 (qAbs xs x qX)"
hence y1_fresh: "ys = xs \<longrightarrow> y1 \<noteq> x" "qAFresh ys y1 qX" by auto
hence 1: "qFresh ys y1 qX" using qAFresh_imp_qFresh by auto
hence y1_fresh_qX: "igFresh MOD ys y1 (?h qX)"
using assms unfolding igSubstClsSTR_def
by(simp add: qFreshAll_igFreshAll_qInitAll)
(* *)
obtain x1 where x1_fresh: "x1 \<notin> {y,y1}" "qFresh xs x1 qX" "qAFresh xs x1 qX"
using obtain_qFresh[of "{y,y1}" "{qX}"] using Abs by blast
hence [simp]: "igFresh MOD xs x1 (?h qX)"
using assms by(simp add: qFreshAll_igFreshAll_qInitAll)
let ?x1 = "igVar MOD xs x1" let ?x1y1y = "x1 @xs[y1 \<and> y]_ys"
let ?qX_x1x = "qX #[[x1 \<and> x]]_xs" let ?qX_x1x_y1y = "?qX_x1x #[[y1 \<and> y]]_ys"
let ?qX_y1y = "qX #[[y1 \<and> y]]_ys" let ?qX_y1y_x1_xy1y = "?qX_y1y #[[x1 \<and> ?xy1y]]_xs"
let ?qX_y1y_x1y1y_xy1y = "?qX_y1y #[[?x1y1y \<and> ?xy1y]]_xs"
have [simp]: "qAFresh ys y1 ?qX_x1x"
using y1_fresh x1_fresh by(auto simp add: qSwap_preserves_qAFresh_distinct)
have [simp]: "qAFresh xs x1 ?qX_y1y"
using y1_fresh x1_fresh by(auto simp add: qSwap_preserves_qAFresh_distinct)
hence "qFresh xs x1 ?qX_y1y" by (simp add: qAFresh_imp_qFresh)
hence [simp]: "igFresh MOD xs x1 (?h ?qX_y1y)"
using assms by(simp add: qFreshAll_igFreshAll_qInitAll)
have [simp]: "igFresh MOD xs x1 ?y1"
using x1_fresh assms unfolding igFreshClsSTR_def igFreshIGVar_def by simp
have x1_def: "x1 = ?x1y1y" using x1_fresh by simp
(* *)
have "?hA ((qAbs xs x qX) $[[y1 \<and> y]]_ys) = igAbs MOD xs ?xy1y (?h ?qX_y1y)" by simp
also have "\<dots> = igAbs MOD xs x1 (igSubst MOD xs ?x1 ?xy1y (?h ?qX_y1y))"
using assms unfolding igAbsRenSTR_def by simp
also have "igSubst MOD xs ?x1 ?xy1y (?h ?qX_y1y) = ?h (?qX_y1y_x1_xy1y)"
using y1_fresh Abs.IH[of "?qX_y1y"] by(simp add: qSwap_qSwapped)
also have "?qX_y1y_x1_xy1y = ?qX_y1y_x1y1y_xy1y" using x1_def by simp
also have "\<dots> = ?qX_x1x_y1y" apply(rule sym) by(rule qSwap_compose)
also have "?h ?qX_x1x_y1y = igSubst MOD ys ?y1 y (?h ?qX_x1x)"
using Abs.IH[of "?qX_x1x"] by(simp add: qSwap_qSwapped)
also have
"igAbs MOD xs x1 (igSubst MOD ys ?y1 y (?h ?qX_x1x)) =
igSubstAbs MOD ys ?y1 y (igAbs MOD xs x1 (?h (?qX_x1x)))"
using assms unfolding igSubstClsSTR_def igSubstIGAbsSTR_def
using x1_fresh y1_fresh by simp
also have "?h (?qX_x1x) = igSubst MOD xs ?x1 x (?h qX)"
using Abs.IH[of "qX"] x1_fresh by(simp add: qSwapped.Refl)
also have
"igAbs MOD xs x1 (igSubst MOD xs ?x1 x (?h qX)) =
igAbs MOD xs x (?h qX)"
using assms unfolding igAbsRenSTR_def by simp
also have "igAbs MOD xs x (?h qX) = ?hA (qAbs xs x qX)"
using assms by simp
finally show "?hA ((qAbs xs x qX) $[[y1 \<and> y]]_ys) =
igSubstAbs MOD ys ?y1 y (?hA (qAbs xs x qX))" .
qed
qed
lemma iwlsFSbSwTR_qSwapAll_igSubstAll_qInitAll:
assumes wls: "iwlsFSbSwTR MOD"
shows
"(qGood qX \<longrightarrow>
qAFresh ys y1 qX \<longrightarrow>
qInit MOD (qX #[[y1 \<and> y]]_ys) = igSubst MOD ys (igVar MOD ys y1) y (qInit MOD qX))
\<and>
(qGoodAbs qA \<longrightarrow>
qAFreshAbs ys y1 qA \<longrightarrow>
qInitAbs MOD (qA $[[y1 \<and> y]]_ys) = igSubstAbs MOD ys (igVar MOD ys y1) y (qInitAbs MOD qA))"
using assms unfolding iwlsFSbSwTR_def by(simp add: qSwapAll_igSubstAll_qInitAll)
lemma iwlsFSwSTR_alphaAll_qInitAll:
assumes "iwlsFSwSTR MOD"
shows
"(\<forall> qX'. qX #= qX' \<longrightarrow> qInit MOD qX = qInit MOD qX') \<and>
(\<forall> qA'. qA $= qA' \<longrightarrow> qInitAbs MOD qA = qInitAbs MOD qA')"
proof(induction rule: qTerm_induct)
case (Var xs x)
then show ?case by(simp add: qVar_alpha_iff)
next
case (Op delta qinp qbinp)
show ?case proof safe
fix qX'
assume "qOp delta qinp qbinp #= qX'"
then obtain qinp' qbinp' where qX': "qX' = qOp delta qinp' qbinp'"
and *: "sameDom qinp qinp' \<and> sameDom qbinp qbinp'"
and **: "liftAll2 (\<lambda>qX qX'. qX #= qX') qinp qinp' \<and>
liftAll2 (\<lambda>qA qA'. qA $= qA') qbinp qbinp'"
using qOp_alpha_iff[of delta qinp qbinp qX'] by auto
hence "lift (qInit MOD) qinp = lift (qInit MOD) qinp'"
by (smt Op.IH(1) liftAll2_def liftAll2_lift_ext liftAll_def)
moreover have "lift (qInitAbs MOD) qbinp = lift (qInitAbs MOD) qbinp'"
by (smt * ** Op.IH(2) liftAll2_def liftAll2_lift_ext liftAll_def)
ultimately
show "qInit MOD (qOp delta qinp qbinp) = qInit MOD qX'" unfolding qX' by simp
qed
next
case (Abs xs x qX)
show ?case proof safe
fix qA'
assume "qAbs xs x qX $= qA'"
then obtain x' y qX' where qA': "qA' = qAbs xs x' qX'"
and y_not: "y \<notin> {x, x'}" and "qAFresh xs y qX" "qAFresh xs y qX'"
and alpha: "(qX #[[y \<and> x]]_xs) #= (qX' #[[y \<and> x']]_xs)"
using qAbs_alphaAbs_iff[of xs x qX qA'] by auto
hence y_fresh: "qFresh xs y qX \<and> qFresh xs y qX'" using qAFresh_imp_qFresh by auto
have "(qX, qX #[[y \<and> x]]_xs) \<in> qSwapped" using qSwap_qSwapped by fastforce
hence "qInit MOD (qX #[[y \<and> x]]_xs) = qInit MOD (qX' #[[y \<and> x']]_xs)"
using Abs.IH alpha by simp
hence "igSwap MOD xs y x (qInit MOD qX) = igSwap MOD xs y x' (qInit MOD qX')"
using assms by(auto simp: iwlsFSwSTR_qSwapAll_igSwapAll_qInitAll)
moreover have "igFresh MOD xs y (qInit MOD qX) \<and> igFresh MOD xs y (qInit MOD qX')"
using y_fresh assms by(auto simp add: iwlsFSwSTR_qFreshAll_igFreshAll_qInitAll)
ultimately have "igAbs MOD xs x (qInit MOD qX) = igAbs MOD xs x' (qInit MOD qX')"
using y_not assms unfolding iwlsFSwSTR_def igAbsCongSSTR_def
apply clarify by (erule allE[of _ xs], erule allE[of _ x]) blast
thus "qInitAbs MOD (qAbs xs x qX) = qInitAbs MOD qA'" unfolding qA' by simp
qed
qed
corollary iwlsFSwSTR_qInit_respectsP_alpha:
assumes "iwlsFSwSTR MOD" shows "(qInit MOD) respectsP alpha"
unfolding congruentP_def using assms iwlsFSwSTR_alphaAll_qInitAll by blast
corollary iwlsFSwSTR_qInitAbs_respectsP_alphaAbs:
assumes "iwlsFSwSTR MOD" shows "(qInitAbs MOD) respectsP alphaAbs"
unfolding congruentP_def using assms iwlsFSwSTR_alphaAll_qInitAll by blast
lemma iwlsFSbSwTR_alphaAll_qInitAll:
fixes qX::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
qA::"('index,'bindex,'varSort,'var,'opSym)qAbs"
assumes "iwlsFSbSwTR MOD"
shows
"(qGood qX \<longrightarrow> (\<forall> qX'. qX #= qX' \<longrightarrow> qInit MOD qX = qInit MOD qX')) \<and>
(qGoodAbs qA \<longrightarrow> (\<forall> qA'. qA $= qA' \<longrightarrow> qInitAbs MOD qA = qInitAbs MOD qA'))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
then show ?case by(simp add: qVar_alpha_iff)
next
case (Op delta qinp qbinp)
show ?case proof safe
fix qX'
assume "qOp delta qinp qbinp #= qX'"
then obtain qinp' qbinp' where qX': "qX' = qOp delta qinp' qbinp'"
and *: "sameDom qinp qinp' \<and> sameDom qbinp qbinp'"
and **: "liftAll2 (\<lambda>qX qX'. qX #= qX') qinp qinp' \<and>
liftAll2 (\<lambda>qA qA'. qA $= qA') qbinp qbinp'"
using qOp_alpha_iff[of delta qinp qbinp qX'] by auto
have "lift (qInit MOD) qinp = lift (qInit MOD) qinp'"
using "*" "**" Op.IH(1) by (simp add: lift_def liftAll2_def liftAll_def
sameDom_def fun_eq_iff split: option.splits) (metis option.exhaust)
moreover
have "lift (qInitAbs MOD) qbinp = lift (qInitAbs MOD) qbinp'"
using "*" "**" Op.IH(2) by (simp add: lift_def liftAll2_def liftAll_def
sameDom_def fun_eq_iff split: option.splits) (metis option.exhaust)
ultimately
show "qInit MOD (qOp delta qinp qbinp) = qInit MOD qX'"
unfolding qX' by simp
qed
next
case (Abs xs x qX)
show ?case proof safe
fix qA'
assume "qAbs xs x qX $= qA'"
then obtain x' y qX' where qA': "qA' = qAbs xs x' qX'"
and y_not: "y \<notin> {x, x'}" and y_afresh: "qAFresh xs y qX" "qAFresh xs y qX'"
and alpha: "(qX #[[y \<and> x]]_xs) #= (qX' #[[y \<and> x']]_xs)"
using qAbs_alphaAbs_iff[of xs x qX qA'] by auto
hence y_fresh: "qFresh xs y qX \<and> qFresh xs y qX'" using qAFresh_imp_qFresh by auto
have qX': "qGood qX'" using alpha Abs by(simp add: alpha_qSwap_preserves_qGood1)
have "(qX, qX #[[y \<and> x]]_xs) \<in> qSwapped" using qSwap_qSwapped by fastforce
hence "qInit MOD (qX #[[y \<and> x]]_xs) = qInit MOD (qX' #[[y \<and> x']]_xs)"
using Abs.IH alpha by simp
moreover have "qInit MOD (qX #[[y \<and> x]]_xs) = igSubst MOD xs (igVar MOD xs y) x (qInit MOD qX)"
using Abs y_afresh assms by(simp add: iwlsFSbSwTR_qSwapAll_igSubstAll_qInitAll)
moreover have "qInit MOD (qX' #[[y \<and> x']]_xs) = igSubst MOD xs (igVar MOD xs y) x' (qInit MOD qX')"
using qX' y_afresh assms by(simp add: iwlsFSbSwTR_qSwapAll_igSubstAll_qInitAll)
ultimately
have "igSubst MOD xs (igVar MOD xs y) x (qInit MOD qX) =
igSubst MOD xs (igVar MOD xs y) x' (qInit MOD qX')"
by simp
moreover have "igFresh MOD xs y (qInit MOD qX) \<and> igFresh MOD xs y (qInit MOD qX')"
using y_fresh assms by(auto simp add: iwlsFSbSwTR_qFreshAll_igFreshAll_qInitAll)
moreover have "igAbsCongUSTR MOD"
using assms unfolding iwlsFSbSwTR_def using igAbsRenSTR_imp_igAbsCongUSTR by auto
ultimately have "igAbs MOD xs x (qInit MOD qX) = igAbs MOD xs x' (qInit MOD qX')"
using y_not unfolding igAbsCongUSTR_def apply clarify
by (erule allE[of _ xs], erule allE[of _ x]) blast
thus "qInitAbs MOD (qAbs xs x qX) = qInitAbs MOD qA'" unfolding qA' by simp
qed
qed
corollary iwlsFSbSwTR_qInit_respectsP_alphaGood:
assumes "iwlsFSbSwTR MOD"
shows "(qInit MOD) respectsP alphaGood"
unfolding congruentP_def alphaGood_def
using assms iwlsFSbSwTR_alphaAll_qInitAll by fastforce
corollary iwlsFSbSwTR_qInitAbs_respectsP_alphaAbsGood:
assumes "iwlsFSbSwTR MOD"
shows "(qInitAbs MOD) respectsP alphaAbsGood"
unfolding congruentP_def alphaAbsGood_def
using assms iwlsFSbSwTR_alphaAll_qInitAll by auto
-subsubsection {* The initial morphism (iteration map) from the term model to any strong model *}
-
-text {* This morphism has the same definition for fresh-swap and fresh-subst strong models *}
+subsubsection \<open>The initial morphism (iteration map) from the term model to any strong model\<close>
+
+text \<open>This morphism has the same definition for fresh-swap and fresh-subst strong models\<close>
definition iterSTR where
"iterSTR MOD == univ (qInit MOD)"
definition iterAbsSTR where
"iterAbsSTR MOD == univ (qInitAbs MOD)"
lemma iwlsFSwSTR_iterSTR_ipresVar:
assumes "iwlsFSwSTR MOD"
shows "ipresVar (iterSTR MOD) MOD"
using assms by(simp add: ipresVar_def Var_def iterSTR_def iwlsFSwSTR_qInit_respectsP_alpha)
lemma iwlsFSbSwTR_iterSTR_ipresVar:
assumes "iwlsFSbSwTR MOD"
shows "ipresVar (iterSTR MOD) MOD"
using assms by (simp add: ipresVar_def Var_def iterSTR_def iwlsFSbSwTR_qInit_respectsP_alphaGood)
lemma iwlsFSwSTR_iterSTR_ipresAbs:
assumes "iwlsFSwSTR MOD"
shows "ipresAbs (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresAbs_def proof clarify
fix xs x s X assume X: "wls s X"
hence "qGood (pick X)" by(simp add: good_imp_qGood_pick)
hence 1: "qGoodAbs (qAbs xs x (pick X))" by simp
have "iterAbsSTR MOD (Abs xs x X) = univ (qInitAbs MOD) (asAbs (qAbs xs x (pick X)))"
using X unfolding Abs_def iterAbsSTR_def by simp
also have "\<dots> = qInitAbs MOD (qAbs xs x (pick X))"
using assms 1 by(simp add: iwlsFSwSTR_qInitAbs_respectsP_alphaAbs)
also have "\<dots> = igAbs MOD xs x (qInit MOD (pick X))" by simp
also have "\<dots> = igAbs MOD xs x (iterSTR MOD X)" unfolding iterSTR_def
unfolding univ_def pick_def ..
finally show "iterAbsSTR MOD (Abs xs x X) = igAbs MOD xs x (iterSTR MOD X)" .
qed
lemma iwlsFSbSwTR_iterSTR_ipresAbs:
assumes "iwlsFSbSwTR MOD"
shows "ipresAbs (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresAbs_def proof clarify
fix xs x s X assume X: "wls s X"
hence "qGood (pick X)" by(simp add: good_imp_qGood_pick)
hence 1: "qGoodAbs (qAbs xs x (pick X))" by simp
have "iterAbsSTR MOD (Abs xs x X) = univ (qInitAbs MOD) (asAbs (qAbs xs x (pick X)))"
using X unfolding Abs_def iterAbsSTR_def by simp
also have "\<dots> = qInitAbs MOD (qAbs xs x (pick X))"
using assms 1 by(simp add: iwlsFSbSwTR_qInitAbs_respectsP_alphaAbsGood)
also have "\<dots> = igAbs MOD xs x (qInit MOD (pick X))" by simp
also have "\<dots> = igAbs MOD xs x (iterSTR MOD X)" unfolding iterSTR_def univ_def
unfolding univ_def pick_def ..
finally show "iterAbsSTR MOD (Abs xs x X) = igAbs MOD xs x (iterSTR MOD X)" .
qed
lemma iwlsFSwSTR_iterSTR_ipresOp:
assumes "iwlsFSwSTR MOD"
shows "ipresOp (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresOp_def proof clarify
fix delta inp binp
assume inp: "wlsInp delta inp" "wlsBinp delta binp"
hence "qGoodInp (pickInp inp) \<and> qGoodBinp (pickBinp binp)"
by(simp add: goodInp_imp_qGoodInp_pickInp goodBinp_imp_qGoodBinp_pickBinp)
hence 1: "qGood (qOp delta (pickInp inp) (pickBinp binp))" by simp
have "iterSTR MOD (Op delta inp binp) =
univ (qInit MOD) (asTerm (qOp delta (pickInp inp) (pickBinp binp)))"
using inp unfolding Op_def iterSTR_def by simp
moreover have "\<dots> = qInit MOD (qOp delta (pickInp inp) (pickBinp binp))"
using assms 1 by(simp add: iwlsFSwSTR_qInit_respectsP_alpha)
moreover have "\<dots> = igOp MOD delta (lift (qInit MOD) (pickInp inp))
(lift (qInitAbs MOD) (pickBinp binp))" by auto
moreover
have "lift (qInit MOD) (pickInp inp) = lift (iterSTR MOD) inp \<and>
lift (qInitAbs MOD) (pickBinp binp) = lift (iterAbsSTR MOD) binp"
unfolding pickInp_def pickBinp_def iterSTR_def iterAbsSTR_def
lift_comp univ_def[abs_def] comp_def
unfolding univ_def pick_def by simp
ultimately
show "iterSTR MOD (Op delta inp binp) =
igOp MOD delta (lift (iterSTR MOD) inp) (lift (iterAbsSTR MOD) binp)" by simp
qed
lemma iwlsFSbSwTR_iterSTR_ipresOp:
assumes "iwlsFSbSwTR MOD"
shows "ipresOp (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresOp_def proof clarify
fix delta inp binp
assume inp: "wlsInp delta inp" "wlsBinp delta binp"
hence "qGoodInp (pickInp inp) \<and> qGoodBinp (pickBinp binp)"
by(simp add: goodInp_imp_qGoodInp_pickInp goodBinp_imp_qGoodBinp_pickBinp)
hence 1: "qGood (qOp delta (pickInp inp) (pickBinp binp))" by simp
have "iterSTR MOD (Op delta inp binp) =
univ (qInit MOD) (asTerm (qOp delta (pickInp inp) (pickBinp binp)))"
using inp unfolding Op_def iterSTR_def by simp
moreover have "\<dots> = qInit MOD (qOp delta (pickInp inp) (pickBinp binp))"
using assms 1 by(simp add: iwlsFSbSwTR_qInit_respectsP_alphaGood)
moreover have "\<dots> = igOp MOD delta (lift (qInit MOD) (pickInp inp))
(lift (qInitAbs MOD) (pickBinp binp))" by simp
moreover have "lift (qInit MOD) (pickInp inp) = lift (iterSTR MOD) inp \<and>
lift (qInitAbs MOD) (pickBinp binp) = lift (iterAbsSTR MOD) binp"
unfolding pickInp_def pickBinp_def iterSTR_def iterAbsSTR_def
lift_comp univ_def[abs_def] comp_def
unfolding univ_def pick_def by simp
ultimately
show "iterSTR MOD (Op delta inp binp) =
igOp MOD delta (lift (iterSTR MOD) inp) (lift (iterAbsSTR MOD) binp)" by simp
qed
lemma iwlsFSwSTR_iterSTR_ipresCons:
assumes "iwlsFSwSTR MOD"
shows "ipresCons (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresCons_def using assms
iwlsFSwSTR_iterSTR_ipresVar
iwlsFSwSTR_iterSTR_ipresAbs
iwlsFSwSTR_iterSTR_ipresOp by auto
lemma iwlsFSbSwTR_iterSTR_ipresCons:
assumes "iwlsFSbSwTR MOD"
shows "ipresCons (iterSTR MOD) (iterAbsSTR MOD) MOD"
unfolding ipresCons_def using assms
iwlsFSbSwTR_iterSTR_ipresVar
iwlsFSbSwTR_iterSTR_ipresAbs
iwlsFSbSwTR_iterSTR_ipresOp by auto
lemma iwlsFSwSTR_iterSTR_termFSwImorph:
assumes "iwlsFSwSTR MOD"
shows "termFSwImorph (iterSTR MOD) (iterAbsSTR MOD) MOD"
using assms by (auto simp: iwlsFSwSTR_termFSwImorph_iff intro: iwlsFSwSTR_iterSTR_ipresCons)
corollary iterSTR_termFSwImorph_errMOD:
assumes "iwlsFSw MOD"
shows
"termFSwImorph (iterSTR (errMOD MOD))
(iterAbsSTR (errMOD MOD))
(errMOD MOD)"
using assms errMOD_iwlsFSwSTR iwlsFSwSTR_iterSTR_termFSwImorph by auto
lemma iwlsFSbSwTR_iterSTR_termFSbImorph:
assumes "iwlsFSbSwTR MOD"
shows "termFSbImorph (iterSTR MOD) (iterAbsSTR MOD) MOD"
using assms by (auto simp: iwlsFSbSwTR_termFSbImorph_iff intro: iwlsFSbSwTR_iterSTR_ipresCons)
corollary iterSTR_termFSbImorph_errMOD:
assumes "iwlsFSb MOD"
shows
"termFSbImorph (iterSTR (errMOD MOD))
(iterAbsSTR (errMOD MOD))
(errMOD MOD)"
using assms errMOD_iwlsFSbSwTR iwlsFSbSwTR_iterSTR_termFSbImorph by auto
(* We are done with the use of quesi-term ``implementations". *)
declare qItem_simps[simp del]
declare qItem_versus_item_simps[simp del]
declare good_item_simps[simp del]
-subsubsection {* The initial morhpism (iteration map) from the term model to any model *}
-
-text {* Again, this morphism has the same definition for fresh-swap and fresh-subst models,
+subsubsection \<open>The initial morhpism (iteration map) from the term model to any model\<close>
+
+text \<open>Again, this morphism has the same definition for fresh-swap and fresh-subst models,
as well as (of course) for fresh-swap-subst and fresh-subst-swap models. (Remember that
there is no such thing as ``fresh-subst-swap" morhpism -- we use the notion of
-``fresh-swap-subst" morphism.) *}
-
-text {* Existence of the morphism: *}
+``fresh-swap-subst" morphism.)\<close>
+
+text \<open>Existence of the morphism:\<close>
definition iter where
"iter MOD == check o (iterSTR (errMOD MOD))"
definition iterAbs where
"iterAbs MOD == check o (iterAbsSTR (errMOD MOD))"
theorem iwlsFSw_iterAll_termFSwImorph:
"iwlsFSw MOD \<Longrightarrow> termFSwImorph (iter MOD) (iterAbs MOD) MOD"
using iterSTR_termFSwImorph_errMOD check_FSwImorph
by (auto simp: iter_def iterAbs_def intro: comp_termFSwImorph)
theorem iwlsFSb_iterAll_termFSbImorph:
"iwlsFSb MOD \<Longrightarrow> termFSbImorph (iter MOD) (iterAbs MOD) MOD"
using iterSTR_termFSbImorph_errMOD check_FSbImorph
by (auto simp: iter_def iterAbs_def intro: comp_termFSbImorph)
theorem iwlsFSwSb_iterAll_termFSwSbImorph:
"iwlsFSwSb MOD \<Longrightarrow> termFSwSbImorph (iter MOD) (iterAbs MOD) MOD"
using iwlsFSw_iterAll_termFSwImorph
by (auto simp: iwlsFSwSb_termFSwSbImorph_iff iwlsFSwSb_def termFSwImorph_def)
theorem iwlsFSbSw_iterAll_termFSwSbImorph:
"iwlsFSbSw MOD \<Longrightarrow> termFSwSbImorph (iter MOD) (iterAbs MOD) MOD"
using iwlsFSb_iterAll_termFSbImorph
by (auto simp: iwlsFSbSw_termFSwSbImorph_iff iwlsFSbSw_def termFSbImorph_def)
-text {* Uniqueness of the morphism *}
-
-text {* In fact, already a presumptive construct-preserving map has to be unique: *}
+text \<open>Uniqueness of the morphism\<close>
+
+text \<open>In fact, already a presumptive construct-preserving map has to be unique:\<close>
lemma ipresCons_unique:
assumes "ipresCons f fA MOD" and "ipresCons ig igA MOD"
shows
"(wls s X \<longrightarrow> f X = ig X) \<and>
(wlsAbs (us,s') A \<longrightarrow> fA A = igA A)"
proof(induction rule: wls_rawInduct)
case (Var xs x)
then show ?case using assms unfolding ipresCons_def ipresVar_def by simp
next
case (Op delta inp binp)
hence "lift f inp = lift ig inp \<and> lift fA binp = lift igA binp"
using assms
apply(simp add: lift_def liftAll2_def sameDom_def fun_eq_iff wlsInp_iff wlsBinp_iff split: option.splits)
using not_None_eq by (metis surj_pair)
thus "f (Op delta inp binp) = ig (Op delta inp binp)"
using assms unfolding ipresCons_def ipresOp_def by (simp add: Op.IH)
next
case (Abs s xs x X)
then show ?case using assms unfolding ipresCons_def ipresAbs_def apply clarify
by (erule allE[of _ xs], erule allE[of _ x]) fastforce
qed
theorem iwlsFSw_iterAll_unique_ipresCons:
assumes "iwlsFSw MOD" and "ipresCons h hA MOD"
shows
"(wls s X \<longrightarrow> h X = iter MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = iterAbs MOD A)"
using assms iwlsFSw_iterAll_termFSwImorph
by (auto simp: termFSwImorph_def intro!: ipresCons_unique)
theorem iwlsFSb_iterAll_unique_ipresCons:
assumes "iwlsFSb MOD" and "ipresCons h hA MOD"
shows
"(wls s X \<longrightarrow> h X = iter MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = iterAbs MOD A)"
using assms iwlsFSb_iterAll_termFSbImorph
by (auto simp: termFSbImorph_def intro!: ipresCons_unique)
theorem iwlsFSwSb_iterAll_unique_ipresCons:
assumes "iwlsFSwSb MOD" and "ipresCons h hA MOD"
shows
"(wls s X \<longrightarrow> h X = iter MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = iterAbs MOD A)"
using assms unfolding iwlsFSwSb_def
using iwlsFSw_iterAll_unique_ipresCons by blast
theorem iwlsFSbSw_iterAll_unique_ipresCons:
assumes *: "iwlsFSbSw MOD" and **: "ipresCons h hA MOD"
shows
"(wls s X \<longrightarrow> h X = iter MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = iterAbs MOD A)"
using assms unfolding iwlsFSbSw_def
using iwlsFSb_iterAll_unique_ipresCons by blast
(**************************************)
lemmas iteration_simps =
input_igSwap_igSubst_None
termMOD_simps
error_model_simps
declare iteration_simps [simp del]
end (* context FixSyn *)
(************************************************)
end
diff --git a/thys/Binding_Syntax_Theory/Preliminaries.thy b/thys/Binding_Syntax_Theory/Preliminaries.thy
--- a/thys/Binding_Syntax_Theory/Preliminaries.thy
+++ b/thys/Binding_Syntax_Theory/Preliminaries.thy
@@ -1,482 +1,482 @@
-section {* Preliminaries *}
+section \<open>Preliminaries\<close>
theory Preliminaries imports "HOL-Cardinals.Cardinals"
begin
-text {* This section discusses preliminaries on families of items (technically,
+text \<open>This section discusses preliminaries on families of items (technically,
partial functions from a type of {\em indexes})
that we call {\em inputs} because they will be inputs to the binding operations.
For inputs, we define some monad-like lifting operators.
We also define simple infinitely branching trees (with no info attached
to the nodes and with branching given by partial functions from
indexes) -- these will be used as ``skeletons'' for terms, giving a size
on which we can induct.
- *}
+\<close>
abbreviation regular where "regular \<equiv> stable"
lemmas regular_UNION = stable_UNION
-subsection {* Trivia *}
+subsection \<open>Trivia\<close>
type_synonym 'a pair = "'a * 'a"
type_synonym 'a triple = "'a * 'a *'a"
type_synonym 'a rel = "'a pair set"
(* Selectors for triples *)
definition fst3 where "fst3 == fst"
definition snd3 where "snd3 == fst o snd"
definition trd3 where "trd3 == snd o snd"
lemma fst3_simp[simp]: "fst3 (a,b,c) = a"
unfolding fst3_def by simp
lemma snd3_simp[simp]: "snd3 (a,b,c) = b"
unfolding snd3_def by simp
lemma trd3_simp[simp]: "trd3 (a,b,c) = c"
unfolding trd3_def by simp
lemma fst3_snd3_trd3: "abc = (fst3 abc, snd3 abc, trd3 abc)"
unfolding fst3_def snd3_def trd3_def by auto
lemma fst3_snd3_trd3_rev[simp]:
"(fst3 abc, snd3 abc, trd3 abc) = abc"
using fst3_snd3_trd3[of abc] by simp
lemma map_id[simp]: "map id l = l"
unfolding id_def by simp
abbreviation max3 where
"max3 x y z == max (max x y) z"
lemmas map_id_cong = map_idI
lemma ext2:
"(f \<noteq> g) = (\<exists> x. f x \<noteq> g x)"
using ext by auto
lemma not_equals_and_not_equals_not_in:
"(y \<noteq> x \<and> y \<noteq> x' \<and> phi) =
(y \<notin> {x,x'} \<and> phi)"
by simp
lemma mp2:
assumes "!! x y. phi x y \<Longrightarrow> chi x y" and "phi x y"
shows "chi x y"
using assms by simp
lemma mp3:
assumes "!! x y z. phi x y z \<Longrightarrow> chi x y z" and "phi x y z"
shows "chi x y z"
using assms by simp
lemma all_lt_Suc:
"(\<forall> i < Suc n. phi i) = ((\<forall> i < n. phi i) \<and> phi n)"
using less_Suc_eq by auto
declare hd_map[simp]
lemmas tl_map[simp] = list.map_sel
declare last_map[simp]
lemma tl_last[simp]:
assumes "tl L \<noteq> []"
shows "last (tl L) = last L"
using assms apply - by(induct L, auto)
lemma tl_last_hd:
assumes "L \<noteq> []" and "tl L = []"
shows "last L = hd L"
using assms apply - by(induct L, auto)
-subsection {* Lexicographic induction *}
+subsection \<open>Lexicographic induction\<close>
definition lt2 where
"lt2 == less_than <*lex*> less_than"
definition lt3 where
"lt3 == less_than <*lex*> lt2"
lemma wf_lt2:
"wf lt2"
unfolding lt2_def by auto
lemma wf_lt3:
"wf lt3"
unfolding lt3_def by (auto simp add: wf_lt2)
lemma lt2[intro]:
"!! k1 k2 j1 j2. k1 < j1 \<Longrightarrow> ((k1,k2),(j1,j2)) \<in> lt2"
"!! k1 k2 j1 j2. \<lbrakk>k1 \<le> j1; k2 < j2\<rbrakk> \<Longrightarrow> ((k1,k2),(j1,j2)) \<in> lt2"
unfolding lt2_def by auto
lemma lt3[intro]:
"!! k1 k2 k3 j1 j2 j3. k1 < j1 \<Longrightarrow> ((k1,k2,k3),(j1,j2,j3)) \<in> lt3"
"!! k1 k2 k3 j1 j2 j3. \<lbrakk>k1 \<le> j1; k2 < j2\<rbrakk> \<Longrightarrow> ((k1,k2,k3),(j1,j2,j3)) \<in> lt3"
"!! k1 k2 k3 j1 j2 j3. \<lbrakk>k1 \<le> j1; k2 \<le> j2; k3 < j3\<rbrakk> \<Longrightarrow> ((k1,k2,k3),(j1,j2,j3)) \<in> lt3"
unfolding lt3_def by auto
lemma measure_lex2_induct:
fixes h1 :: "'a1 \<Rightarrow> nat" and h2 :: "'a2 \<Rightarrow> nat"
assumes
"!! x1 x2.
\<lbrakk>(!! y1 y2. h1 y1 < h1 x1 \<Longrightarrow> phi y1 y2);
(!! y1 y2. \<lbrakk>h1 y1 \<le> h1 x1; h2 y2 < h2 x2\<rbrakk> \<Longrightarrow> phi y1 y2)\<rbrakk>
\<Longrightarrow> phi x1 x2"
shows "phi x1 x2"
proof-
let ?chi = "%(n1,n2). ALL x1 x2. h1 x1 = n1 \<and> h2 x2 = n2 \<longrightarrow> phi x1 x2"
{fix n1 n2
have "?chi (n1,n2)"
apply(rule wf_induct[of lt2 ?chi]) using wf_lt2 assms by blast+
}
thus ?thesis by blast
qed
lemma measure_lex3_induct:
fixes h1 :: "'a1 \<Rightarrow> nat" and h2 :: "'a2 \<Rightarrow> nat" and h3 :: "'a3 \<Rightarrow> nat"
assumes
"!! x1 x2 x3.
\<lbrakk>(!! y1 y2 y3. h1 y1 < h1 x1 \<Longrightarrow> phi y1 y2 y3);
(!! y1 y2 y3. \<lbrakk>h1 y1 \<le> h1 x1; h2 y2 < h2 x2\<rbrakk> \<Longrightarrow> phi y1 y2 y3);
(!! y1 y2 y3. \<lbrakk>h1 y1 \<le> h1 x1; h2 y2 \<le> h2 x2; h3 y3 < h3 x3\<rbrakk> \<Longrightarrow> phi y1 y2 y3)\<rbrakk>
\<Longrightarrow> phi x1 x2 x3"
shows "phi x1 x2 x3"
proof-
let ?chi = "%(n1,n2,n3). ALL x1 x2 x3. h1 x1 = n1 \<and> h2 x2 = n2 \<and> h3 x3 = n3 \<longrightarrow> phi x1 x2 x3"
{fix n1 n2 n3
have "?chi (n1,n2,n3)"
apply(rule wf_induct[of lt3 ?chi]) using wf_lt3 assms by blast+
}
thus ?thesis by blast
qed
-subsection {* Inputs and lifting operators *}
+subsection \<open>Inputs and lifting operators\<close>
type_synonym ('index,'val)input = "'index \<Rightarrow> 'val option"
definition
lift :: "('val1 \<Rightarrow> 'val2) \<Rightarrow> ('index,'val1)input \<Rightarrow> ('index,'val2)input"
where
"lift h inp == \<lambda>i. case inp i of None \<Rightarrow> None
|Some v \<Rightarrow> Some (h v)"
definition
liftAll :: "('val \<Rightarrow> bool) \<Rightarrow> ('index,'val)input \<Rightarrow> bool"
where
"liftAll phi inp == \<forall> i v. inp i = Some v \<longrightarrow> phi v"
definition
lift2 ::
"('val1' \<Rightarrow> 'val1 \<Rightarrow> 'val2) \<Rightarrow> ('index,'val1')input \<Rightarrow> ('index,'val1)input \<Rightarrow> ('index,'val2)input"
where
"lift2 h inp' inp ==
\<lambda>i. case (inp' i, inp i) of
(Some v',Some v) \<Rightarrow> Some (h v' v)
|_ \<Rightarrow> None"
definition
sameDom :: "('index,'val1)input \<Rightarrow> ('index,'val2)input \<Rightarrow> bool"
where "sameDom inp1 inp2 == \<forall> i. (inp1 i = None) = (inp2 i = None)"
definition
liftAll2 :: "('val1 \<Rightarrow> 'val2 \<Rightarrow> bool) \<Rightarrow> ('index,'val1)input \<Rightarrow> ('index,'val2)input \<Rightarrow> bool"
where
"liftAll2 phi inp1 inp2 == (\<forall> i v1 v2. inp1 i = Some v1 \<and> inp2 i = Some v2 \<longrightarrow> phi v1 v2)"
lemma lift_None: "(lift h inp i = None) = (inp i = None)"
unfolding lift_def by (cases "inp i", auto)
lemma lift_Some:
"(\<exists> v. lift h inp i = Some v) = (\<exists> v'. inp i = Some v')"
using lift_None[of h inp i] by force
lemma lift_cong[fundef_cong]:
assumes "\<And> i v. inp i = Some v \<Longrightarrow> h v = h' v"
shows "lift h inp = lift h' inp"
unfolding lift_def apply(rule ext)+
using assms by (case_tac "inp i", auto)
lemma lift_preserves_inj:
assumes "inj f"
shows "inj (lift f)"
unfolding inj_on_def apply auto proof(rule ext)
fix inp inp' i assume *: "lift f inp = lift f inp'"
show "inp i = inp' i"
proof(cases "inp i")
assume inp: "inp i = None"
hence "lift f inp i = None" unfolding lift_def by simp
hence "lift f inp' i = None" using * by simp
hence "inp' i = None" by(auto simp add: lift_None)
thus ?thesis using inp by simp
next
fix v assume inp: "inp i = Some v"
hence "lift f inp i = Some (f v)" unfolding lift_def by simp
hence "lift f inp' i = Some (f v)" using * by simp
then obtain v' where inp': "inp' i = Some v'" and "f v' = f v"
unfolding lift_def by(case_tac "inp' i", auto)
hence "v = v'" using assms unfolding inj_on_def by simp
thus ?thesis using inp inp' by simp
qed
qed
lemma liftAll_cong[fundef_cong]:
assumes "\<And> i v. inp i = Some v \<Longrightarrow> phi v = phi' v"
shows "liftAll phi inp = liftAll phi' inp"
unfolding liftAll_def apply((rule iff_allI)+) using assms by simp
lemma liftAll2_cong[fundef_cong]:
assumes "\<And> i v1 v2. \<lbrakk>inp1 i = Some v1; inp2 i = Some v2\<rbrakk> \<Longrightarrow> phi v1 v2 = phi' v1 v2"
shows "liftAll2 phi inp1 inp2 = liftAll2 phi' inp1 inp2"
unfolding liftAll2_def apply((rule iff_allI)+) using assms by blast
lemma lift_ident: "lift (\<lambda>v. v) inp = inp"
by(unfold lift_def, rule ext, case_tac "inp i", auto)
lemma lift_id[simp]:
"lift id inp = inp"
unfolding lift_def apply (rule ext) by(case_tac "inp i", auto)
lemma lift_comp: "lift g (lift f inp) = lift (g o f) inp"
by(unfold lift_def o_def, rule ext, case_tac "inp i", auto)
lemma liftAll_mono:
assumes "\<And> v. phi v \<Longrightarrow> chi v" and "liftAll phi inp"
shows "liftAll chi inp"
using assms unfolding liftAll_def by blast
lemma liftAll_True: "liftAll (\<lambda>v. True) inp"
unfolding liftAll_def by auto
lemma liftAll_lift_comp: "liftAll phi (lift f inp) = liftAll (phi o f) inp"
unfolding liftAll_def o_def
by (metis (mono_tags, lifting) lift_Some lift_def option.inject option.simps(5))
lemma liftAll_lift_ext:
"liftAll (\<lambda> x. f x = g x) inp = (lift f inp = lift g inp)"
unfolding lift_def liftAll_def
by (auto simp: fun_eq_iff option.case_eq_if)
lemma liftAll_and:
"liftAll (\<lambda> x. phi x \<and> chi x) inp = (liftAll phi inp \<and> liftAll chi inp)"
unfolding liftAll_def by blast
lemma liftAll_mp:
assumes "liftAll (\<lambda> v. phi v \<longrightarrow> chi v) inp" and "liftAll phi inp"
shows "liftAll chi inp"
using assms unfolding liftAll_def by auto
lemma sameDom_refl: "sameDom inp inp"
unfolding sameDom_def by auto
lemma sameDom_sym:
"sameDom inp inp' = sameDom inp' inp"
unfolding sameDom_def by auto
lemma sameDom_trans:
"\<lbrakk>sameDom inp inp'; sameDom inp' inp''\<rbrakk> \<Longrightarrow> sameDom inp inp''"
unfolding sameDom_def by auto
lemma sameDom_lift1:
"sameDom inp (lift f inp)"
unfolding sameDom_def lift_def
by (auto simp: option.case_eq_if)
lemma sameDom_lift2:
"sameDom (lift f inp) inp"
unfolding sameDom_def lift_def
by (auto simp: option.case_eq_if)
lemma sameDom_lift_simp1[simp]:
"sameDom inp (lift f inp') = sameDom inp inp'"
unfolding sameDom_def lift_def by (force simp: fun_eq_iff option.case_eq_if)
lemma sameDom_lift_simp2[simp]:
"sameDom (lift f inp) inp' = sameDom inp inp'"
unfolding sameDom_def lift_def by (force simp: fun_eq_iff option.case_eq_if)
lemma lift_preserves_sameDom:
assumes "sameDom inp inp'"
shows "sameDom (lift f inp) (lift g inp')"
using assms unfolding sameDom_def lift_def
by (force simp: fun_eq_iff option.case_eq_if)
definition comp2 ::
"('b1 \<Rightarrow> 'b2 \<Rightarrow> 'c) \<Rightarrow> ('a1 \<Rightarrow> 'b1) \<Rightarrow> ('a2 \<Rightarrow> 'b2) \<Rightarrow> ('a1 \<Rightarrow> 'a2 \<Rightarrow> 'c)"
("_ o2 '(_,_')" 55)
where "h o2 (f,g) == \<lambda> x y. h (f x) (g y)"
lemma comp2_simp[simp]:
"(h o2 (f,g)) x y = h (f x) (g y)"
unfolding comp2_def by simp
lemma comp2_comp:
"((h o2 (f,g)) o2 (f',g')) = (h o2 (f o f', g o g'))"
unfolding comp_def[abs_def] comp2_def[abs_def] by auto
lemma liftAll_imp_liftAll2:
assumes "liftAll (\<lambda>v. \<forall> v'. phi v v') inp"
shows "liftAll2 phi inp inp'"
using assms unfolding liftAll_def liftAll2_def by auto
lemma liftAll2_mono:
assumes "\<And> v v'. phi v v' \<Longrightarrow> chi v v'" and "liftAll2 phi inp inp'"
shows "liftAll2 chi inp inp'"
using assms unfolding liftAll2_def by blast
lemma liftAll2_True: "liftAll2 (\<lambda> v v'. True) inp inp'"
unfolding liftAll2_def by auto
lemma liftAll2_lift_comp2:
"liftAll2 phi (lift f1 inp1) (lift f2 inp2) =
liftAll2 (phi o2 (f1,f2)) inp1 inp2"
unfolding liftAll2_def lift_def
by (auto simp: fun_eq_iff option.case_eq_if)
lemma lift_imp_sameDom:
"lift f inp = lift f inp' \<Longrightarrow> sameDom inp inp'"
unfolding lift_def sameDom_def
by (force simp: fun_eq_iff option.case_eq_if split: if_splits)
lemma lift_lift2:
"lift f (lift2 g inp' inp) =
lift2 (\<lambda> v' v. f (g v' v)) inp' inp"
unfolding lift_def lift2_def
by (force simp: option.case_eq_if split: if_splits)
lemma lift2_left[simp]:
assumes "sameDom inp' inp"
shows "lift2 (\<lambda> v' v. v') inp' inp = inp'"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if) metis
lemma lift2_right[simp]:
assumes "sameDom inp' inp"
shows "lift2 (\<lambda> v' v. v) inp' inp = inp"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if)
lemma lift2_preserves_sameDom:
assumes "sameDom inp' inp1'" and "sameDom inp inp1"
shows "sameDom (lift2 f inp' inp) (lift2 g inp1' inp1)"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if)
lemma sameDom_lift2_1:
assumes "sameDom inp' inp"
shows
"sameDom inp' (lift2 f inp' inp) \<and>
sameDom inp (lift2 f inp' inp)"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if)
lemma sameDom_lift2_2:
assumes "sameDom inp' inp"
shows
"sameDom (lift2 f inp' inp) inp' \<and>
sameDom (lift2 f inp' inp) inp"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if)
lemma sameDom_lift2_simp1[simp]:
assumes "sameDom inp1' inp1"
shows "sameDom inp (lift2 f inp1' inp1) = sameDom inp inp1'"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if) (metis not_Some_eq)
lemma sameDom_lift2_simp2[simp]:
assumes "sameDom inp' inp"
shows "sameDom (lift2 f inp' inp) inp1 = sameDom inp' inp1"
using assms unfolding sameDom_def lift2_def
by (simp add: fun_eq_iff option.case_eq_if) (metis not_Some_eq)
lemma liftAll2_lift_ext:
"(sameDom inp inp' \<and> liftAll2 (\<lambda> v v'. f v = f v') inp inp') =
(lift f inp = lift f inp')"
unfolding sameDom_def lift_def liftAll2_def
by (force simp add: fun_eq_iff option.case_eq_if)
lemma liftAll2_and:
"liftAll2 (\<lambda> v v'. phi v v' \<and> chi v v') inp inp' =
(liftAll2 phi inp inp' \<and> liftAll2 chi inp inp')"
unfolding liftAll2_def by force
lemma liftAll2_mp:
assumes "liftAll2 (\<lambda> v v'. phi v v' \<longrightarrow> chi v v') inp inp'" and "liftAll2 phi inp inp'"
shows "liftAll2 chi inp inp'"
using assms unfolding liftAll2_def by auto
lemma sameDom_and_liftAll2_iff:
"(sameDom inp inp' \<and> liftAll2 phi inp inp') =
(\<forall> i. (inp i = None \<and> inp' i = None) \<or>
(\<exists> v v'. inp i = Some v \<and> inp' i = Some v' \<and> phi v v'))"
unfolding sameDom_def liftAll2_def
apply (auto simp add: fun_eq_iff option.case_eq_if)
using option.sel by fastforce
-subsection {* Doubly infinitely-branching trees *}
+subsection \<open>Doubly infinitely-branching trees\<close>
text "These simple infinitary trees shall be used for measuring the sizes
of possibly infinitary terms."
datatype ('index,'bindex)tree =
Branch "('index,('index,'bindex)tree) input" "('bindex,('index,'bindex)tree) input"
(* The natural induction principle for (infinitary) trees: *)
lemma tree_induct:
fixes phi::"('index,'bindex)tree \<Rightarrow> bool" and T::"('index,'bindex)tree"
assumes
"\<And> inp binp. \<lbrakk>liftAll phi inp; liftAll phi binp\<rbrakk> \<Longrightarrow> phi (Branch inp binp)"
shows "phi T"
using assms unfolding liftAll_def
by (induct T) (simp, metis rangeI)
definition treeLess :: "('index,'bindex)tree rel"
where
"treeLess ==
{(T,T'). \<exists> inp binp i j. T' = Branch inp binp \<and> (inp i = Some T \<or> binp j = Some T)}"
lemma treeLess_induct:
fixes phi::"('index,'bindex)tree \<Rightarrow> bool" and
T::"('index,'bindex)tree"
assumes "\<And> T'. (\<And> T. (T,T') \<in> treeLess \<Longrightarrow> phi T) \<Longrightarrow> phi T'"
shows "phi T"
apply(induct rule: tree_induct)
using assms unfolding treeLess_def liftAll_def
by simp (metis tree.inject)
lemma treeLess_wf: "wf treeLess"
unfolding wf_def using treeLess_induct by blast
-subsection {* Ordering *}
+subsection \<open>Ordering\<close>
lemma Least_Max:
assumes phi: "phi (n::nat)" and fin: "finite {n. phi n}"
shows "(LEAST m. \<forall> n. phi n \<longrightarrow> n \<le> m) =
Max {n. phi n}"
using assms Max_in by (intro Least_equality) auto
end
diff --git a/thys/Binding_Syntax_Theory/QuasiTerms_Environments_Substitution.thy b/thys/Binding_Syntax_Theory/QuasiTerms_Environments_Substitution.thy
--- a/thys/Binding_Syntax_Theory/QuasiTerms_Environments_Substitution.thy
+++ b/thys/Binding_Syntax_Theory/QuasiTerms_Environments_Substitution.thy
@@ -1,830 +1,830 @@
-section {* Environments and Substitution for Quasi-Terms *}
+section \<open>Environments and Substitution for Quasi-Terms\<close>
theory QuasiTerms_Environments_Substitution
imports QuasiTerms_PickFresh_Alpha
begin
-text{* Inside this theory, since anyway all the interesting properties hold only
-modulo alpha, we forget completely about qAFresh and use only qFresh. *}
+text\<open>Inside this theory, since anyway all the interesting properties hold only
+modulo alpha, we forget completely about qAFresh and use only qFresh.\<close>
-text{* In this section we define, for quasi-terms, parallel substitution according to
+text\<open>In this section we define, for quasi-terms, parallel substitution according to
{\em environments}.
This is the most general kind of substitution -- an environment, i.e., a partial
map from variables
to quasi-terms, indicates which quasi-term (if any) to be substituted for each
variable; substitution
is then applied to a subject quasi-term and an environment. In order to ``keep up"
with the notion
of good quasi-term, we define good environments and show that substitution
preserves goodness. Since,
unlike swapping, substitution does not behave well w.r.t. quasi-terms
(but only w.r.t. terms, i.e., to alpha-equivalence classes),
here we prove the minimum amount of properties required for properly lifting
parallel substitution to terms. Then compositionality properties
of parallel substitution will be proved directly for terms.
-*}
+\<close>
-subsection {* Environments *}
+subsection \<open>Environments\<close>
type_synonym ('index,'bindex,'varSort,'var,'opSym)qEnv =
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTerm option"
(*********************************************)
context FixVars (* scope all throughout the file *)
begin
definition qGoodEnv :: "('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow> bool"
where
"qGoodEnv rho ==
(\<forall> xs. liftAll qGood (rho xs)) \<and>
(\<forall> ys. |{y. rho ys y \<noteq> None}| <o |UNIV :: 'var set| )"
definition qFreshEnv where
"qFreshEnv zs z rho ==
rho zs z = None \<and> (\<forall> xs. liftAll (qFresh zs z) (rho xs))"
definition alphaEnv where
"alphaEnv =
{(rho,rho'). \<forall> xs. sameDom (rho xs) (rho' xs) \<and>
liftAll2 (\<lambda>X X'. X #= X') (rho xs) (rho' xs)}"
abbreviation alphaEnv_abbrev ::
"('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow> bool" (infix "&=" 50)
where
"rho &= rho' == (rho,rho') \<in> alphaEnv"
definition pickQFreshEnv
where
"pickQFreshEnv xs V XS Rho ==
pickQFresh xs (V \<union> (\<Union> rho \<in> Rho. {x. rho xs x \<noteq> None}))
(XS \<union> (\<Union> rho \<in> Rho. {X. \<exists> ys y. rho ys y = Some X}))"
lemma qGoodEnv_imp_card_of_qTerm:
assumes "qGoodEnv rho"
shows "|{X. \<exists> y. rho ys y = Some X}| <o |UNIV :: 'var set|"
proof-
let ?rel = "{(y,X). rho ys y = Some X}"
let ?Left = "{X. \<exists> y. rho ys y = Some X}"
let ?Left' = "{y. \<exists> X. rho ys y = Some X}"
have "\<And> y X X'. (y,X) \<in> ?rel \<and> (y,X') \<in> ?rel \<longrightarrow> X = X'" by force
hence "|?Left| \<le>o |?Left'|" using card_of_inj_rel[of ?rel] by auto
moreover have "|?Left'| <o |UNIV :: 'var set|" using assms unfolding qGoodEnv_def by auto
ultimately show ?thesis using ordLeq_ordLess_trans by blast
qed
lemma qGoodEnv_imp_card_of_qTerm2:
assumes "qGoodEnv rho"
shows "|{X. \<exists> ys y. rho ys y = Some X}| <o |UNIV :: 'var set|"
proof-
let ?Left = "{X. \<exists> ys y. rho ys y = Some X}"
let ?F = "\<lambda> ys. {X. \<exists> y. rho ys y = Some X}"
have "?Left = (\<Union> ys. ?F ys)" by auto
moreover have "\<forall> ys. |?F ys| <o |UNIV :: 'var set|"
using assms qGoodEnv_imp_card_of_qTerm by auto
ultimately show ?thesis
using var_regular_INNER varSort_lt_var_INNER by(force simp add: regular_UNION)
qed
lemma qGoodEnv_iff:
"qGoodEnv rho =
((\<forall> xs. liftAll qGood (rho xs)) \<and>
(\<forall> ys. |{y. rho ys y \<noteq> None}| <o |UNIV :: 'var set| ) \<and>
|{X. \<exists> ys y. rho ys y = Some X}| <o |UNIV :: 'var set| )"
unfolding qGoodEnv_def apply auto
apply(rule qGoodEnv_imp_card_of_qTerm2) unfolding qGoodEnv_def by simp
lemma alphaEnv_refl:
"qGoodEnv rho \<Longrightarrow> rho &= rho"
using alpha_refl
unfolding alphaEnv_def qGoodEnv_def liftAll_def liftAll2_def sameDom_def by fastforce
lemma alphaEnv_sym:
"rho &= rho' \<Longrightarrow> rho' &= rho"
using alpha_sym unfolding alphaEnv_def liftAll2_def sameDom_def by fastforce
lemma alphaEnv_trans:
assumes good: "qGoodEnv rho" and
alpha1: "rho &= rho'" and alpha2: "rho' &= rho''"
shows "rho &= rho''"
using assms unfolding alphaEnv_def
apply(auto)
using sameDom_trans apply blast
unfolding liftAll2_def proof(auto)
fix xs x X X''
assume rho: "rho xs x = Some X" and rho'': "rho'' xs x = Some X''"
moreover have "(rho xs x = None) = (rho' xs x = None)"
using alpha1 unfolding alphaEnv_def sameDom_def by auto
ultimately obtain X' where rho': "rho' xs x = Some X'" by auto
hence "X #= X'" using alpha1 rho unfolding alphaEnv_def liftAll2_def by auto
moreover have "X' #= X''"
using alpha2 rho' rho'' unfolding alphaEnv_def liftAll2_def by auto
moreover have "qGood X" using good rho unfolding qGoodEnv_def liftAll_def by auto
ultimately show "X #= X''" using alpha_trans by blast
qed
lemma pickQFreshEnv_card_of:
assumes Vvar: "|V| <o |UNIV :: 'var set|" and XSvar: "|XS| <o |UNIV :: 'var set|" and
good: "\<forall> X \<in> XS. qGood X" and
Rhovar: "|Rho| <o |UNIV :: 'var set|" and RhoGood: "\<forall> rho \<in> Rho. qGoodEnv rho"
shows
"pickQFreshEnv xs V XS Rho \<notin> V \<and>
(\<forall> X \<in> XS. qFresh xs (pickQFreshEnv xs V XS Rho) X) \<and>
(\<forall> rho \<in> Rho. qFreshEnv xs (pickQFreshEnv xs V XS Rho) rho)"
proof-
let ?z =" pickQFreshEnv xs V XS Rho"
let ?V2 = "\<Union> rho \<in> Rho. {x. rho xs x \<noteq> None}" let ?W = "V \<union> ?V2"
let ?XS2 = "\<Union> rho \<in> Rho. {X. \<exists> ys y. rho ys y = Some X}" let ?YS = "XS \<union> ?XS2"
have "|?W| <o |UNIV :: 'var set|"
proof-
have "\<forall> rho \<in> Rho. |{x. rho xs x \<noteq> None}| <o |UNIV :: 'var set|"
using RhoGood unfolding qGoodEnv_iff using qGoodEnv_iff by auto
hence "|?V2| <o |UNIV :: 'var set|"
using var_regular_INNER Rhovar by (auto simp add: regular_UNION)
thus ?thesis using var_infinite_INNER Vvar card_of_Un_ordLess_infinite by auto
qed
moreover have "|?YS| <o |UNIV :: 'var set|"
proof-
have "\<forall> rho \<in> Rho. |{X. \<exists> ys y. rho ys y = Some X}| <o |UNIV :: 'var set|"
using RhoGood unfolding qGoodEnv_iff by auto
hence "|?XS2| <o |UNIV :: 'var set|"
using var_regular_INNER Rhovar by (auto simp add: regular_UNION)
thus ?thesis using var_infinite_INNER XSvar card_of_Un_ordLess_infinite by auto
qed
moreover have "\<forall> Y \<in> ?YS. qGood Y"
using good RhoGood unfolding qGoodEnv_iff liftAll_def by blast
ultimately
have "?z \<notin> ?W \<and> (\<forall> Y \<in> ?YS. qFresh xs ?z Y)"
unfolding pickQFreshEnv_def using pickQFresh_card_of[of ?W ?YS] by auto
thus ?thesis unfolding qFreshEnv_def liftAll_def by(auto)
qed
lemma pickQFreshEnv:
assumes Vvar: "|V| <o |UNIV :: 'var set| \<or> finite V"
and XSvar: "|XS| <o |UNIV :: 'var set| \<or> finite XS"
and good: "\<forall> X \<in> XS. qGood X"
and Rhovar: "|Rho| <o |UNIV :: 'var set| \<or> finite Rho"
and RhoGood: "\<forall> rho \<in> Rho. qGoodEnv rho"
shows
"pickQFreshEnv xs V XS Rho \<notin> V \<and>
(\<forall> X \<in> XS. qFresh xs (pickQFreshEnv xs V XS Rho) X) \<and>
(\<forall> rho \<in> Rho. qFreshEnv xs (pickQFreshEnv xs V XS Rho) rho)"
proof-
have 1: "|V| <o |UNIV :: 'var set| \<and> |XS| <o |UNIV :: 'var set| \<and> |Rho| <o |UNIV :: 'var set|"
using assms var_infinite_INNER by(auto simp add: finite_ordLess_infinite2)
show ?thesis
apply(rule pickQFreshEnv_card_of)
using assms 1 by auto
qed
corollary obtain_qFreshEnv:
fixes XS::"('index,'bindex,'varSort,'var,'opSym)qTerm set" and
Rho::"('index,'bindex,'varSort,'var,'opSym)qEnv set" and rho
assumes Vvar: "|V| <o |UNIV :: 'var set| \<or> finite V"
and XSvar: "|XS| <o |UNIV :: 'var set| \<or> finite XS"
and good: "\<forall> X \<in> XS. qGood X"
and Rhovar: "|Rho| <o |UNIV :: 'var set| \<or> finite Rho"
and RhoGood: "\<forall> rho \<in> Rho. qGoodEnv rho"
shows
"\<exists> z. z \<notin> V \<and>
(\<forall> X \<in> XS. qFresh xs z X) \<and> (\<forall> rho \<in> Rho. qFreshEnv xs z rho)"
apply(rule exI[of _ "pickQFreshEnv xs V XS Rho"])
using assms by(rule pickQFreshEnv)
-subsection {* Parallel substitution *}
+subsection \<open>Parallel substitution\<close>
(* I shall prove only a *minimal* collection of facts for quasi-
[parallel substitution], just enough
to show that substitution preserves alpha. The other properties shall be proved
for alpha-equivalence directly. *)
definition aux_qPsubst_ignoreFirst ::
"('index,'bindex,'varSort,'var,'opSym)qEnv * ('index,'bindex,'varSort,'var,'opSym)qTerm +
('index,'bindex,'varSort,'var,'opSym)qEnv * ('index,'bindex,'varSort,'var,'opSym)qAbs
\<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTermItem"
where
"aux_qPsubst_ignoreFirst K ==
case K of Inl (rho,X) \<Rightarrow> termIn X
|Inr (rho,A) \<Rightarrow> absIn A"
lemma aux_qPsubst_ignoreFirst_qTermLessQSwapped_wf:
"wf(inv_image qTermQSwappedLess aux_qPsubst_ignoreFirst)"
using qTermQSwappedLess_wf wf_inv_image by auto
function
qPsubst ::
"('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTerm"
and
qPsubstAbs ::
"('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qAbs"
where
"qPsubst rho (qVar xs x) = (case rho xs x of None \<Rightarrow> qVar xs x| Some X \<Rightarrow> X)"
|
"qPsubst rho (qOp delta inp binp) =
qOp delta (lift (qPsubst rho) inp) (lift (qPsubstAbs rho) binp)"
|
"qPsubstAbs rho (qAbs xs x X) =
(let x' = pickQFreshEnv xs {x} {X} {rho} in qAbs xs x' (qPsubst rho (X #[[x' \<and> x]]_xs)))"
by(pat_completeness, auto)
termination
apply(relation "inv_image qTermQSwappedLess aux_qPsubst_ignoreFirst")
apply(simp add: aux_qPsubst_ignoreFirst_qTermLessQSwapped_wf)
by(auto simp add: qTermQSwappedLess_def qTermLess_modulo_def
aux_qPsubst_ignoreFirst_def qSwap_qSwapped)
abbreviation qPsubst_abbrev ::
"('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTerm" ("_ #[[_]]")
where "X #[[rho]] == qPsubst rho X"
abbreviation qPsubstAbs_abbrev ::
"('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qEnv \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qAbs" ("_ $[[_]]")
where "A $[[rho]] == qPsubstAbs rho A"
lemma qPsubstAll_preserves_qGoodAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and rho
assumes GOOD_ENV: "qGoodEnv rho"
shows
"(qGood X \<longrightarrow> qGood (X #[[rho]])) \<and> (qGoodAbs A \<longrightarrow> qGoodAbs (A $[[rho]]))"
proof(induction rule: qTerm_induct[of _ _ X A])
case (Var xs x)
show ?case
using GOOD_ENV unfolding qGoodEnv_iff liftAll_def
by(cases "rho xs x", auto)
next
case (Op delta inp binp)
show ?case proof safe
assume g: "qGood (qOp delta inp binp)"
hence 0: "liftAll qGood (lift (qPsubst rho) inp) \<and>
liftAll qGoodAbs (lift (qPsubstAbs rho) binp)"
using Op unfolding liftAll_lift_comp comp_def
by (simp_all add: Let_def liftAll_mp)
have "{i. lift (qPsubst rho) inp i \<noteq> None} = {i. inp i \<noteq> None} \<and>
{i. lift (qPsubstAbs rho) binp i \<noteq> None} = {i. binp i \<noteq> None}"
by simp (meson lift_Some)
hence "|{i. \<exists>y. lift (qPsubst rho) inp i = Some y}| <o |UNIV:: 'var set|"
and "|{i. \<exists>y. lift (qPsubstAbs rho) binp i = Some y}| <o |UNIV:: 'var set|"
using g by (auto simp: liftAll_def)
thus "qGood qOp delta inp binp #[[rho]]" using 0 by simp
qed
next
case (Abs xs x X)
show ?case proof safe
assume g: "qGoodAbs (qAbs xs x X)"
let ?x' = "pickQFreshEnv xs {x} {X} {rho}" let ?X' = "X #[[?x' \<and> x]]_xs"
have "qGood ?X'" using g qSwap_preserves_qGood by auto
moreover have "(X,?X') \<in> qSwapped" using qSwap_qSwapped by fastforce
ultimately have "qGood (qPsubst rho ?X')" using Abs.IH by simp
thus "qGoodAbs ((qAbs xs x X) $[[rho]])" by (simp add: Let_def)
qed
qed
corollary qPsubst_preserves_qGood:
"\<lbrakk>qGoodEnv rho; qGood X\<rbrakk> \<Longrightarrow> qGood (X #[[rho]])"
using qPsubstAll_preserves_qGoodAll by auto
corollary qPsubstAbs_preserves_qGoodAbs:
"\<lbrakk>qGoodEnv rho; qGoodAbs A\<rbrakk> \<Longrightarrow> qGoodAbs (A $[[rho]])"
using qPsubstAll_preserves_qGoodAll by auto
lemma qPsubstAll_preserves_qFreshAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and rho
assumes GOOD_ENV: "qGoodEnv rho"
shows
"(qFresh zs z X \<longrightarrow>
(qGood X \<and> qFreshEnv zs z rho \<longrightarrow> qFresh zs z (X #[[rho]]))) \<and>
(qFreshAbs zs z A \<longrightarrow>
(qGoodAbs A \<and> qFreshEnv zs z rho \<longrightarrow> qFreshAbs zs z (A $[[rho]])))"
proof(induction rule: qTerm_induct[of _ _ X A])
case (Var xs x)
then show ?case
unfolding qFreshEnv_def liftAll_def by (cases "rho xs x") auto
next
case (Op delta inp binp)
thus ?case
by (auto simp add: lift_def liftAll_def qFreshEnv_def split: option.splits)
next
case (Abs xs x X)
show ?case proof safe
assume q: "qFreshAbs zs z (qAbs xs x X)"
"qGoodAbs (qAbs xs x X)" "qFreshEnv zs z rho"
let ?x' = "pickQFreshEnv xs {x} {X} {rho}" let ?X' = "X #[[?x' \<and> x]]_xs"
have x': "qFresh xs ?x' X \<and> qFreshEnv xs ?x' rho"
using q GOOD_ENV by(auto simp add: pickQFreshEnv)
hence goodX': "qGood ?X'" using q qSwap_preserves_qGood by auto
have XX': "(X,?X') \<in> qSwapped" using qSwap_qSwapped by fastforce
have "(zs = xs \<and> z = ?x') \<or> qFresh zs z (qPsubst rho ?X')"
by (meson qSwap_preserves_qFresh_distinct
Abs.IH(1) XX' goodX' q qAbs_alphaAbs_qSwap_qFresh qFreshAbs.simps
qFreshAbs_preserves_alphaAbs1 qSwap_preserves_qGood2 x')
thus "qFreshAbs zs z ((qAbs xs x X) $[[rho]])"
by simp (meson qFreshAbs.simps)+
qed
qed
lemma qPsubst_preserves_qFresh:
"\<lbrakk>qGood X; qGoodEnv rho; qFresh zs z X; qFreshEnv zs z rho\<rbrakk>
\<Longrightarrow> qFresh zs z (X #[[rho]])"
by(simp add: qPsubstAll_preserves_qFreshAll)
lemma qPsubstAbs_preserves_qFreshAbs:
"\<lbrakk>qGoodAbs A; qGoodEnv rho; qFreshAbs zs z A; qFreshEnv zs z rho\<rbrakk>
\<Longrightarrow> qFreshAbs zs z (A $[[rho]])"
by(simp add: qPsubstAll_preserves_qFreshAll)
-text{* While in general we try to avoid proving facts in parallel,
+text\<open>While in general we try to avoid proving facts in parallel,
here we seem to have no choice -- it is the first time we must use mutual
-induction: *}
+induction:\<close>
lemma qPsubstAll_preserves_alphaAll_qSwapAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and
rho::"('index,'bindex,'varSort,'var,'opSym)qEnv"
assumes goodRho: "qGoodEnv rho"
shows
"(qGood X \<longrightarrow>
(\<forall> Y. X #= Y \<longrightarrow> (X #[[rho]]) #= (Y #[[rho]])) \<and>
(\<forall> xs z1 z2. qFreshEnv xs z1 rho \<and> qFreshEnv xs z2 rho \<longrightarrow>
((X #[[z1 \<and> z2]]_xs) #[[rho]]) #= ((X #[[rho]]) #[[z1 \<and> z2]]_xs))) \<and>
(qGoodAbs A \<longrightarrow>
(\<forall> B. A $= B \<longrightarrow> (A $[[rho]]) $= (B $[[rho]])) \<and>
(\<forall> xs z1 z2. qFreshEnv xs z1 rho \<and> qFreshEnv xs z2 rho \<longrightarrow>
((A $[[z1 \<and> z2]]_xs) $[[rho]]) $= ((A $[[rho]]) $[[z1 \<and> z2]]_xs)))"
proof(induction rule: qGood_qTerm_induct_mutual)
case (Var1 xs x)
then show ?case
by (metis alpha_refl goodRho qGood.simps(1) qPsubst_preserves_qGood qVar_alpha_iff)
next
case (Var2 xs x)
show ?case proof safe
fix s::'sort and zs z1 z2
assume FreshEnv: "qFreshEnv zs z1 rho" "qFreshEnv zs z2 rho"
hence n: "rho zs z1 = None \<and> rho zs z2 = None" unfolding qFreshEnv_def by simp
let ?Left = "qPsubst rho ((qVar xs x) #[[z1 \<and> z2]]_zs)"
let ?Right = "(qPsubst rho (qVar xs x)) #[[z1 \<and> z2]]_zs"
have "qGood (qVar xs x)" by simp
hence "qGood ((qVar xs x) #[[z1 \<and> z2]]_zs)"
using qSwap_preserves_qGood by blast
hence goodLeft: "qGood ?Left" using goodRho qPsubst_preserves_qGood by blast
show "?Left #= ?Right"
proof(cases "rho xs x")
case None
hence "rho xs (x @xs[z1 \<and> z2]_zs) = None"
using n unfolding sw_def by auto
thus ?thesis using None by simp
next
case (Some X)
hence "xs \<noteq> zs \<or> x \<notin> {z1,z2}" using n by auto
hence "(x @xs[z1 \<and> z2]_zs) = x" unfolding sw_def by auto
moreover
{have "qFresh zs z1 X \<and> qFresh zs z2 X"
using Some FreshEnv unfolding qFreshEnv_def liftAll_def by auto
moreover have "qGood X" using Some goodRho unfolding qGoodEnv_def liftAll_def by auto
ultimately have "X #= (X #[[z1 \<and> z2]]_zs)"
by(auto simp: alpha_qFresh_qSwap_id alpha_sym)
}
ultimately show ?thesis using Some by simp
qed
qed
next
case (Op1 delta inp binp)
show ?case proof safe
fix Y assume q: "qOp delta inp binp #= Y"
then obtain inp' binp' where Y: "Y = qOp delta inp' binp'" and
*: "(\<forall>i. (inp i = None) = (inp' i = None)) \<and>
(\<forall>i. (binp i = None) = (binp' i = None))" and
**: "(\<forall>i X X'. inp i = Some X \<and> inp' i = Some X' \<longrightarrow> X #= X') \<and>
(\<forall>i A A'. binp i = Some A \<and> binp' i = Some A' \<longrightarrow> A $= A')"
unfolding qOp_alpha_iff sameDom_def liftAll2_def by auto
show "(qOp delta inp binp) #[[rho]] #= (Y #[[rho]])"
using Op1 **
by (simp add: Y sameDom_def liftAll2_def)
(fastforce simp add: * lift_None lift_Some
liftAll_def lift_def split: option.splits)
qed
next
case (Op2 delta inp binp)
thus ?case
by (auto simp: sameDom_def liftAll2_def lift_None lift_def liftAll_def split: option.splits)
next
case (Abs1 xs x X)
show ?case proof safe
fix B
assume alpha_xXB: "qAbs xs x X $= B"
then obtain y Y where B: "B = qAbs xs y Y" unfolding qAbs_alphaAbs_iff by auto
- have "qGoodAbs B" using `qGood X` alpha_xXB alphaAbs_preserves_qGoodAbs by force
+ have "qGoodAbs B" using \<open>qGood X\<close> alpha_xXB alphaAbs_preserves_qGoodAbs by force
hence goodY: "qGood Y" unfolding B by simp
let ?x' = "pickQFreshEnv xs {x} {X} {rho}"
let ?y' = "pickQFreshEnv xs {y} {Y} {rho}"
obtain x' and y' where x'y'_def: "x' = ?x'" "y' = ?y'" and
x'y'_rev: "?x' = x'" "?y' = y'" by blast
have x'y'_freshXY: "qFresh xs x' X \<and> qFresh xs y' Y"
- unfolding x'y'_def using `qGood X` goodY goodRho by (auto simp add: pickQFreshEnv)
+ unfolding x'y'_def using \<open>qGood X\<close> goodY goodRho by (auto simp add: pickQFreshEnv)
have x'y'_fresh_rho: "qFreshEnv xs x' rho \<and> qFreshEnv xs y' rho"
- unfolding x'y'_def using `qGood X` goodY goodRho by (auto simp add: pickQFreshEnv)
+ unfolding x'y'_def using \<open>qGood X\<close> goodY goodRho by (auto simp add: pickQFreshEnv)
have x'y'_not_xy: "x' \<noteq> x \<and> y' \<noteq> y"
- unfolding x'y'_def using `qGood X` goodY goodRho
+ unfolding x'y'_def using \<open>qGood X\<close> goodY goodRho
using pickQFreshEnv[of "{x}" "{X}"] pickQFreshEnv[of "{y}" "{Y}"] by force
- have goodXx'x: "qGood (X #[[x' \<and> x]]_xs)" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXx'x: "qGood (X #[[x' \<and> x]]_xs)" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence good: "qGood(qPsubst rho (X #[[x' \<and> x]]_xs))"
using goodRho qPsubst_preserves_qGood by auto
have goodYy'y: "qGood (Y #[[y' \<and> y]]_xs)" using goodY qSwap_preserves_qGood by auto
obtain z where z_not: "z \<notin> {x,y,x',y'}" and
z_fresh_XY: "qFresh xs z X \<and> qFresh xs z Y"
- and z_fresh_rho: "qFreshEnv xs z rho" using `qGood X` goodY goodRho
+ and z_fresh_rho: "qFreshEnv xs z rho" using \<open>qGood X\<close> goodY goodRho
using obtain_qFreshEnv[of "{x,y,x',y'}" "{X,Y}" "{rho}"] by auto
(* Notations: *)
let ?Xx'x = "X #[[x' \<and> x]]_xs" let ?Yy'y = "Y #[[y' \<and> y]]_xs"
let ?Xx'xzx' = "?Xx'x #[[z \<and> x']]_xs" let ?Yy'yzy' = "?Yy'y #[[z \<and> y']]_xs"
let ?Xzx = "X #[[z \<and> x]]_xs" let ?Yzy = "Y #[[z \<and> y]]_xs"
(* Preliminary facts: *)
- have goodXx'x: "qGood ?Xx'x" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXx'x: "qGood ?Xx'x" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence goodXx'xzx': "qGood ?Xx'xzx'" using qSwap_preserves_qGood by auto
have "qGood (?Xx'x #[[rho]])" using goodXx'x goodRho qPsubst_preserves_qGood by auto
hence goodXx'x_rho_zx': "qGood ((?Xx'x #[[rho]]) #[[z \<and> x']]_xs)"
using qSwap_preserves_qGood by auto
have goodYy'y: "qGood ?Yy'y" using goodY qSwap_preserves_qGood by auto
(* *)
have skelXx'x: "qSkel ?Xx'x = qSkel X" using qSkel_qSwap by fastforce
hence skelXx'xzx': "qSkel ?Xx'xzx' = qSkel X" by (auto simp add: qSkel_qSwap)
have "qSkelAbs B = qSkelAbs (qAbs xs x X)"
using alpha_xXB alphaAll_qSkelAll by fastforce
hence "qSkel Y = qSkel X" unfolding B by(auto simp add: fun_eq_iff)
hence skelYy'y: "qSkel ?Yy'y = qSkel X" by(auto simp add: qSkel_qSwap)
(* Main proof: *)
have "((?Xx'x #[[rho]]) #[[z \<and> x']]_xs) #= (?Xx'xzx' #[[rho]])"
using skelXx'x goodXx'x z_fresh_rho x'y'_fresh_rho
Abs1.IH(2)[of "?Xx'x"] by (auto simp add: alpha_sym)
moreover
{have "?Xx'xzx' #= ?Xzx"
- using `qGood X` x'y'_freshXY z_fresh_XY alpha_qFresh_qSwap_compose by fastforce
+ using \<open>qGood X\<close> x'y'_freshXY z_fresh_XY alpha_qFresh_qSwap_compose by fastforce
moreover have "?Xzx #= ?Yzy" using alpha_xXB unfolding B
- using z_fresh_XY `qGood X` goodY
+ using z_fresh_XY \<open>qGood X\<close> goodY
by (simp only: alphaAbs_qAbs_iff_all_qFresh)
moreover have "?Yzy #= ?Yy'yzy'" using goodY x'y'_freshXY z_fresh_XY
by(auto simp add: alpha_qFresh_qSwap_compose alpha_sym)
ultimately have "?Xx'xzx' #= ?Yy'yzy'" using goodXx'xzx' alpha_trans by blast
hence "(?Xx'xzx' #[[rho]]) #= (?Yy'yzy' #[[rho]])"
using goodXx'xzx' skelXx'xzx' Abs1.IH(1) by auto
}
moreover have "(?Yy'yzy' #[[rho]]) #= ((?Yy'y #[[rho]]) #[[z \<and> y']]_xs)"
using skelYy'y goodYy'y z_fresh_rho x'y'_fresh_rho
Abs1.IH(2)[of "?Yy'y"] alpha_sym by fastforce
ultimately
have "((?Xx'x #[[rho]]) #[[z \<and> x']]_xs) #= ((?Yy'y #[[rho]]) #[[z \<and> y']]_xs)"
using goodXx'x_rho_zx' alpha_trans by blast
thus "(qAbs xs x X) $[[rho]] $= (B $[[rho]])"
unfolding B apply simp unfolding Let_def
unfolding x'y'_rev
using good z_not apply(simp only: alphaAbs_qAbs_iff_ex_qFresh)
by (auto intro!: exI[of _ z]
simp: alphaAbs_qAbs_iff_ex_qFresh goodRho goodXx'x qPsubstAll_preserves_qFreshAll
qSwap_preserves_qFresh_distinct z_fresh_XY goodYy'y qPsubst_preserves_qFresh z_fresh_rho)
qed
next
case (Abs2 xs x X)
show ?case proof safe
fix zs z1 z2
assume z1z2_fresh_rho: "qFreshEnv zs z1 rho" "qFreshEnv zs z2 rho"
let ?x' = "pickQFreshEnv xs {x @xs[z1 \<and> z2]_zs} {X #[[z1 \<and> z2]]_zs} {rho}"
let ?x'' = "pickQFreshEnv xs {x} {X} {rho}"
obtain x' x'' where x'x''_def: "x' = ?x'" "x'' = ?x''" and
x'x''_rev: "?x' = x'" "?x'' = x''" by blast
let ?xa = "x @xs[z1 \<and> z2]_zs" let ?xa'' = "x'' @xs[z1 \<and> z2]_zs"
obtain u where "u \<notin> {x,x',x'',z1,z2}" and
u_fresh_X: "qFresh xs u X" and u_fresh_rho: "qFreshEnv xs u rho"
- using `qGood X` goodRho using obtain_qFreshEnv[of "{x,x',x'',z1,z2}" "{X}" "{rho}"] by auto
+ using \<open>qGood X\<close> goodRho using obtain_qFreshEnv[of "{x,x',x'',z1,z2}" "{X}" "{rho}"] by auto
hence u_not: "u \<notin> {x,x',x'',z1,z2,?xa,?xa''}" unfolding sw_def by auto
let ?ua = "u @xs [z1 \<and> z2]_zs"
let ?Xz1z2 = "X #[[z1 \<and> z2]]_zs"
let ?Xz1z2x'xa = "?Xz1z2 #[[x' \<and> ?xa]]_xs"
let ?Xz1z2x'xa_rho = "?Xz1z2x'xa #[[rho]]"
let ?Xz1z2x'xa_rho_ux' = "?Xz1z2x'xa_rho #[[u \<and> x']]_xs"
let ?Xz1z2x'xaux' = "?Xz1z2x'xa #[[u \<and> x']]_xs"
let ?Xz1z2x'xaux'_rho = "?Xz1z2x'xaux' #[[rho]]"
let ?Xz1z2uxa = "?Xz1z2 #[[u \<and> ?xa]]_xs"
let ?Xz1z2uaxa = "?Xz1z2 #[[?ua \<and> ?xa]]_xs"
let ?Xux = "X #[[u \<and> x]]_xs"
let ?Xuxz1z2 = "?Xux #[[z1 \<and> z2]]_zs"
let ?Xx''x = "X #[[x'' \<and> x]]_xs"
let ?Xx''xux'' = "?Xx''x #[[u \<and> x'']]_xs"
let ?Xx''xux''z1z2 = "?Xx''xux'' #[[z1 \<and> z2]]_zs"
let ?Xx''xz1z2 = "?Xx''x #[[z1 \<and> z2]]_zs"
let ?Xx''xz1z2uaxa'' = "?Xx''xz1z2 #[[?ua \<and> ?xa'']]_xs"
let ?Xx''xz1z2uaxa''_rho = "?Xx''xz1z2uaxa'' #[[rho]]"
let ?Xx''xz1z2uxa'' = "?Xx''xz1z2 #[[u \<and> ?xa'']]_xs"
let ?Xx''xz1z2uxa''_rho = "?Xx''xz1z2uxa'' #[[rho]]"
let ?Xx''xz1z2_rho = "?Xx''xz1z2 #[[rho]]"
let ?Xx''xz1z2_rho_uxa'' = "?Xx''xz1z2_rho #[[u \<and> ?xa'']]_xs"
let ?Xx''x_rho = "?Xx''x #[[rho]]"
let ?Xx''x_rho_z1z2 = "?Xx''x_rho #[[z1 \<and> z2]]_zs"
let ?Xx''x_rho_z1z2uxa'' = "?Xx''x_rho_z1z2 #[[u \<and> ?xa'']]_xs"
(* Facts about x', x'', ?xa, ?ua, ?xa'': *)
- have goodXz1z2: "qGood ?Xz1z2" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXz1z2: "qGood ?Xz1z2" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
have x'x''_fresh_Xz1z2: "qFresh xs x' ?Xz1z2 \<and> qFresh xs x'' X"
- unfolding x'x''_def using `qGood X` goodXz1z2 goodRho by (auto simp add: pickQFreshEnv)
+ unfolding x'x''_def using \<open>qGood X\<close> goodXz1z2 goodRho by (auto simp add: pickQFreshEnv)
have x'x''_fresh_rho: "qFreshEnv xs x' rho \<and> qFreshEnv xs x'' rho"
- unfolding x'x''_def using `qGood X` goodXz1z2 goodRho by (auto simp add: pickQFreshEnv)
+ unfolding x'x''_def using \<open>qGood X\<close> goodXz1z2 goodRho by (auto simp add: pickQFreshEnv)
have ua_eq_u: "?ua = u" using u_not unfolding sw_def by auto
(* Good: *)
have goodXz1z2x'xa: "qGood ?Xz1z2x'xa" using goodXz1z2 qSwap_preserves_qGood by auto
- have goodXux: "qGood ?Xux" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXux: "qGood ?Xux" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence goodXuxz1z2: "qGood ?Xuxz1z2" using qSwap_preserves_qGood by auto
- have goodXx''x: "qGood ?Xx''x" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXx''x: "qGood ?Xx''x" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence goodXx''xz1z2: "qGood ?Xx''xz1z2" using qSwap_preserves_qGood by auto
hence "qGood ?Xx''xz1z2_rho" using goodRho qPsubst_preserves_qGood by auto
hence goodXx''xz1z2_rho: "qGood ?Xx''xz1z2_rho"
using goodRho qPsubst_preserves_qGood by auto
have goodXz1z2x'xaux': "qGood ?Xz1z2x'xaux'"
using goodXz1z2x'xa qSwap_preserves_qGood by auto
have goodXz1z2x'xa_rho: "qGood ?Xz1z2x'xa_rho"
using goodXz1z2x'xa goodRho qPsubst_preserves_qGood by auto
hence goodXz1z2x'xa_rho_ux': "qGood ?Xz1z2x'xa_rho_ux'"
using qSwap_preserves_qGood by auto
(* Fresh: *)
have xa''_fresh_rho: "qFreshEnv xs ?xa'' rho"
using x'x''_fresh_rho z1z2_fresh_rho unfolding sw_def by auto
have u_fresh_Xz1z2: "qFresh xs u ?Xz1z2"
using u_fresh_X u_not by(auto simp add: qSwap_preserves_qFresh_distinct)
hence "qFresh xs u ?Xz1z2x'xa" using u_not by(auto simp add: qSwap_preserves_qFresh_distinct)
hence u_fresh_Xz1z2x'xa_rho: "qFresh xs u ?Xz1z2x'xa_rho"
using u_fresh_rho u_fresh_X goodRho goodXz1z2x'xa qPsubst_preserves_qFresh by auto
have "qFresh xs u ?Xx''x"
using u_fresh_X u_not by(auto simp add: qSwap_preserves_qFresh_distinct)
hence "qFresh xs u ?Xx''x_rho" using goodRho goodXx''x u_fresh_rho
by(auto simp add: qPsubst_preserves_qFresh)
hence u_fresh_Xx''x_rho_z1z2: "qFresh xs u ?Xx''x_rho_z1z2"
using u_not by(auto simp add: qSwap_preserves_qFresh_distinct)
(* Skeleton: *)
have skel_Xz1z2x'xa: "qSkel ?Xz1z2x'xa = qSkel X" by(auto simp add: qSkel_qSwap)
hence skel_Xz1z2x'xaux': "qSkel ?Xz1z2x'xaux' = qSkel X" by(auto simp add: qSkel_qSwap)
have skel_Xx''x: "qSkel ?Xx''x = qSkel X" by(auto simp add: qSkel_qSwap)
hence skel_Xx''xz1z2: "qSkel ?Xx''xz1z2 = qSkel X" by(auto simp add: qSkel_qSwap)
(* Main proof: *)
have "?Xz1z2x'xaux'_rho #= ?Xz1z2x'xa_rho_ux'"
using x'x''_fresh_rho u_fresh_rho skel_Xz1z2x'xa goodXz1z2x'xa
using Abs2.IH(2)[of ?Xz1z2x'xa] by auto
hence "?Xz1z2x'xa_rho_ux' #= ?Xz1z2x'xaux'_rho" using alpha_sym by auto
moreover
{have "?Xz1z2x'xaux' #= ?Xz1z2uxa"
using goodXz1z2 u_fresh_Xz1z2 x'x''_fresh_Xz1z2
using alpha_qFresh_qSwap_compose by fastforce
moreover have "?Xz1z2uxa = ?Xuxz1z2"
using ua_eq_u qSwap_compose[of zs z1 z2 xs x u X] by(auto simp: qSwap_sym)
moreover
{have "?Xux #= ?Xx''xux''"
- using `qGood X` u_fresh_X x'x''_fresh_Xz1z2
+ using \<open>qGood X\<close> u_fresh_X x'x''_fresh_Xz1z2
by(auto simp: alpha_qFresh_qSwap_compose alpha_sym)
hence "?Xuxz1z2 #= ?Xx''xux''z1z2"
using goodXux by (auto simp add: qSwap_preserves_alpha)
}
moreover have "?Xx''xux''z1z2 = ?Xx''xz1z2uxa''"
using ua_eq_u qSwap_compose[of zs z1 z2 _ _ _ ?Xx''x] by auto
ultimately have "?Xz1z2x'xaux' #= ?Xx''xz1z2uxa''"
using goodXz1z2x'xaux' alpha_trans by auto
hence "?Xz1z2x'xaux'_rho #= ?Xx''xz1z2uxa''_rho"
using goodXz1z2x'xaux' skel_Xz1z2x'xaux' Abs2.IH(1) by auto
}
moreover have "?Xx''xz1z2uxa''_rho #= ?Xx''xz1z2_rho_uxa''"
using xa''_fresh_rho u_fresh_rho skel_Xx''xz1z2 goodXx''xz1z2
using Abs2.IH(2)[of ?Xx''xz1z2] by auto
moreover
{have "?Xx''xz1z2_rho #= ?Xx''x_rho_z1z2"
using z1z2_fresh_rho skel_Xx''x goodXx''x
using Abs2.IH(2)[of ?Xx''x] by auto
hence "?Xx''xz1z2_rho_uxa'' #= ?Xx''x_rho_z1z2uxa''"
using goodXx''xz1z2_rho by(auto simp add: qSwap_preserves_alpha)
}
ultimately have "?Xz1z2x'xa_rho_ux' #= ?Xx''x_rho_z1z2uxa''"
using goodXz1z2x'xa_rho_ux' alpha_trans by blast
thus "((qAbs xs x X) $[[z1 \<and> z2]]_zs) $[[rho]] $=
(((qAbs xs x X) $[[rho]]) $[[z1 \<and> z2]]_zs)"
using goodXz1z2x'xa_rho
goodXz1z2x'xa u_not u_fresh_Xz1z2x'xa_rho u_fresh_Xx''x_rho_z1z2
apply(simp add: Let_def x'x''_rev del: alpha.simps alphaAbs.simps )
by (auto simp only: Let_def alphaAbs_qAbs_iff_ex_qFresh)
qed
qed
corollary qPsubst_preserves_alpha1:
assumes "qGoodEnv rho" and "qGood X \<or> qGood Y" and "X #= Y"
shows "(X #[[rho]]) #= (Y #[[rho]])"
using alpha_preserves_qGood assms qPsubstAll_preserves_alphaAll_qSwapAll by blast
corollary qPsubstAbs_preserves_alphaAbs1:
assumes "qGoodEnv rho" and "qGoodAbs A \<or> qGoodAbs B" and "A $= B"
shows "(A $[[rho]]) $= (B $[[rho]])"
using alphaAbs_preserves_qGoodAbs assms qPsubstAll_preserves_alphaAll_qSwapAll by blast
corollary alpha_qFreshEnv_qSwap_qPsubst_commute:
"\<lbrakk>qGoodEnv rho; qGood X; qFreshEnv zs z1 rho; qFreshEnv zs z2 rho\<rbrakk> \<Longrightarrow>
((X #[[z1 \<and> z2]]_zs) #[[rho]]) #= ((X #[[rho]]) #[[z1 \<and> z2]]_zs)"
by(simp add: qPsubstAll_preserves_alphaAll_qSwapAll)
corollary alphaAbs_qFreshEnv_qSwapAbs_qPsubstAbs_commute:
"\<lbrakk>qGoodEnv rho; qGoodAbs A;
qFreshEnv zs z1 rho; qFreshEnv zs z2 rho\<rbrakk> \<Longrightarrow>
((A $[[z1 \<and> z2]]_zs) $[[rho]]) $= ((A $[[rho]]) $[[z1 \<and> z2]]_zs)"
by(simp add: qPsubstAll_preserves_alphaAll_qSwapAll)
lemma qPsubstAll_preserves_alphaAll2:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and
rho'::"('index,'bindex,'varSort,'var,'opSym)qEnv" and rho''
assumes rho'_alpha_rho'': "rho' &= rho''" and
goodRho': "qGoodEnv rho'" and goodRho'': "qGoodEnv rho''"
shows
"(qGood X \<longrightarrow> (X #[[rho']]) #= (X #[[rho'']])) \<and>
(qGoodAbs A \<longrightarrow> (A $[[rho']]) $= (A $[[rho'']]))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
then show ?case
proof (cases "rho' xs x")
case None
hence "rho'' xs x = None" using rho'_alpha_rho'' unfolding alphaEnv_def sameDom_def by auto
thus ?thesis using None by simp
next
case (Some X')
then obtain X'' where rho'': "rho'' xs x = Some X''"
using assms unfolding alphaEnv_def sameDom_def by force
hence "X' #= X''" using Some rho'_alpha_rho''
unfolding alphaEnv_def liftAll2_def by auto
thus ?thesis using Some rho'' by simp
qed
next
case (Op delta inp binp)
then show ?case
by (auto simp: lift_def liftAll_def liftAll2_def sameDom_def Let_def
split: option.splits)
next
case (Abs xs x X)
let ?x' = "pickQFreshEnv xs {x} {X} {rho'}"
let ?x'' = "pickQFreshEnv xs {x} {X} {rho''}"
obtain x' x'' where x'x''_def: "x' = ?x'" "x'' = ?x''" and
x'x''_rev: "?x' = x'" "?x'' = x''" by blast
have x'x''_fresh_X: "qFresh xs x' X \<and> qFresh xs x'' X"
- unfolding x'x''_def using `qGood X` goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
+ unfolding x'x''_def using \<open>qGood X\<close> goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
have x'_fresh_rho': "qFreshEnv xs x' rho'"
- unfolding x'x''_def using `qGood X` goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
+ unfolding x'x''_def using \<open>qGood X\<close> goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
have x''_fresh_rho'': "qFreshEnv xs x'' rho''"
- unfolding x'x''_def using `qGood X` goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
+ unfolding x'x''_def using \<open>qGood X\<close> goodRho' goodRho'' by (auto simp add: pickQFreshEnv)
obtain u where u_not: "u \<notin> {x,x',x''}" and
u_fresh_X: "qFresh xs u X" and
u_fresh_rho': "qFreshEnv xs u rho'" and u_fresh_rho'': "qFreshEnv xs u rho''"
- using `qGood X` goodRho' goodRho''
+ using \<open>qGood X\<close> goodRho' goodRho''
using obtain_qFreshEnv[of "{x,x',x''}" "{X}" "{rho',rho''}"] by auto
(* Preliminary facts and notations: *)
let ?Xx'x = "X #[[x' \<and> x]]_xs"
let ?Xx'x_rho' = "?Xx'x #[[rho']]"
let ?Xx'x_rho'_ux' = "?Xx'x_rho' #[[u \<and> x']]_xs"
let ?Xx'xux' = "?Xx'x #[[u \<and> x']]_xs"
let ?Xx'xux'_rho' = "?Xx'xux' #[[rho']]"
let ?Xux = "X #[[u \<and> x]]_xs"
let ?Xux_rho' = "?Xux #[[rho']]"
let ?Xux_rho'' = "?Xux #[[rho'']]"
let ?Xx''x = "X #[[x'' \<and> x]]_xs"
let ?Xx''xux'' = "?Xx''x #[[u \<and> x'']]_xs"
let ?Xx''xux''_rho'' = "?Xx''xux'' #[[rho'']]"
let ?Xx''x_rho'' = "?Xx''x #[[rho'']]"
let ?Xx''x_rho''_ux'' = "?Xx''x_rho'' #[[u \<and> x'']]_xs"
(* Good: *)
- have goodXx'x: "qGood ?Xx'x" using `qGood X` qSwap_preserves_qGood by auto
- hence goodXx'x_rho': "qGood ?Xx'x_rho'" using `qGood X` goodRho' qPsubst_preserves_qGood by auto
+ have goodXx'x: "qGood ?Xx'x" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
+ hence goodXx'x_rho': "qGood ?Xx'x_rho'" using \<open>qGood X\<close> goodRho' qPsubst_preserves_qGood by auto
hence goodXx'x_rho'_ux': "qGood ?Xx'x_rho'_ux'"
- using `qGood X` qSwap_preserves_qGood by auto
+ using \<open>qGood X\<close> qSwap_preserves_qGood by auto
have goodXx'xux': "qGood ?Xx'xux'" using goodXx'x qSwap_preserves_qGood by auto
- have goodXux: "qGood ?Xux" using `qGood X` qSwap_preserves_qGood by auto
- have goodXx''x: "qGood ?Xx''x" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXux: "qGood ?Xux" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
+ have goodXx''x: "qGood ?Xx''x" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence goodXx''x_rho'': "qGood ?Xx''x_rho''"
- using `qGood X` goodRho'' qPsubst_preserves_qGood by auto
+ using \<open>qGood X\<close> goodRho'' qPsubst_preserves_qGood by auto
(* Fresh: *)
have "qFresh xs u ?Xx'x" using u_not u_fresh_X
by(auto simp add: qSwap_preserves_qFresh_distinct)
hence fresh_Xx'x_rho': "qFresh xs u ?Xx'x_rho'"
using u_fresh_rho' goodXx'x goodRho' by(auto simp add: qPsubst_preserves_qFresh)
have "qFresh xs u ?Xx''x" using u_not u_fresh_X
by(auto simp add: qSwap_preserves_qFresh_distinct)
hence fresh_Xx''x_rho'': "qFresh xs u ?Xx''x_rho''"
using u_fresh_rho'' goodXx''x goodRho'' by(auto simp add: qPsubst_preserves_qFresh)
(* qSwapped: *)
have Xux: "(X,?Xux) :qSwapped" by(simp add: qSwap_qSwapped)
(* Main proof: *)
have "?Xx'x_rho'_ux' #= ?Xx'xux'_rho'"
using goodRho' goodXx'x u_fresh_rho' x'_fresh_rho'
by(auto simp: alpha_qFreshEnv_qSwap_qPsubst_commute alpha_sym)
moreover
- {have "?Xx'xux' #= ?Xux" using `qGood X` u_fresh_X x'x''_fresh_X
+ {have "?Xx'xux' #= ?Xux" using \<open>qGood X\<close> u_fresh_X x'x''_fresh_X
using alpha_qFresh_qSwap_compose by fastforce
hence "?Xx'xux'_rho' #= ?Xux_rho'" using goodXx'xux' goodRho'
using qPsubst_preserves_alpha1 by auto
}
moreover have "?Xux_rho' #= ?Xux_rho''" using Xux Abs.IH by auto
moreover
- {have "?Xux #= ?Xx''xux''" using `qGood X` u_fresh_X x'x''_fresh_X
+ {have "?Xux #= ?Xx''xux''" using \<open>qGood X\<close> u_fresh_X x'x''_fresh_X
by(auto simp add: alpha_qFresh_qSwap_compose alpha_sym)
hence "?Xux_rho'' #= ?Xx''xux''_rho''" using goodXux goodRho''
using qPsubst_preserves_alpha1 by auto
}
moreover have "?Xx''xux''_rho'' #= ?Xx''x_rho''_ux''"
using goodRho'' goodXx''x u_fresh_rho'' x''_fresh_rho''
by(auto simp: alpha_qFreshEnv_qSwap_qPsubst_commute)
ultimately have "?Xx'x_rho'_ux' #= ?Xx''x_rho''_ux''"
using goodXx'x_rho'_ux' alpha_trans by blast
hence "qAbs xs ?x' (qPsubst rho' (X #[[?x' \<and> x]]_xs)) $=
qAbs xs ?x''(qPsubst rho''(X #[[?x''\<and> x]]_xs))"
unfolding x'x''_rev using goodXx'x_rho' fresh_Xx'x_rho' fresh_Xx''x_rho''
by (auto simp only: alphaAbs_qAbs_iff_ex_qFresh)
thus ?case by (metis qPsubstAbs.simps)
qed
corollary qPsubst_preserves_alpha2:
"\<lbrakk>qGood X; qGoodEnv rho'; qGoodEnv rho''; rho' &= rho''\<rbrakk>
\<Longrightarrow> (X #[[rho']]) #= (X #[[rho'']])"
by(simp add: qPsubstAll_preserves_alphaAll2)
corollary qPsubstAbs_preserves_alphaAbs2:
"\<lbrakk>qGoodAbs A; qGoodEnv rho'; qGoodEnv rho''; rho' &= rho''\<rbrakk>
\<Longrightarrow> (A $[[rho']]) $= (A $[[rho'']])"
by(simp add: qPsubstAll_preserves_alphaAll2)
lemma qPsubst_preserves_alpha:
assumes "qGood X \<or> qGood X'" and "qGoodEnv rho" and "qGoodEnv rho'"
and "X #= X'" and "rho &= rho'"
shows "(X #[[rho]]) #= (X' #[[rho']])"
by (metis (no_types, lifting) assms alpha_trans qPsubst_preserves_alpha1
qPsubst_preserves_alpha2 qPsubst_preserves_qGood)
lemma qPsubstAbs_preserves_alphaAbs:
assumes "qGoodAbs A \<or> qGoodAbs A'" and "qGoodEnv rho" and "qGoodEnv rho'"
and "A $= A'" and "rho &= rho'"
shows "(A $[[rho]]) $= (A' $[[rho']])"
using assms
by (meson alphaAbs_trans qPsubstAbs_preserves_alphaAbs1
qPsubstAbs_preserves_qGoodAbs qPsubstAll_preserves_alphaAll2)
lemma qFresh_qPsubst_commute_qAbs:
assumes good_X: "qGood X" and good_rho: "qGoodEnv rho" and
x_fresh_rho: "qFreshEnv xs x rho"
shows "((qAbs xs x X) $[[rho]]) $= qAbs xs x (X #[[rho]])"
proof-
(* Preliminary facts and notations: *)
let ?x' = "pickQFreshEnv xs {x} {X} {rho}"
obtain x' where x'_def: "x' = ?x'" and x'_rev: "?x' = x'" by blast
have x'_not: "x' \<noteq> x" unfolding x'_def
using assms pickQFreshEnv[of "{x}" "{X}"] by auto
have x'_fresh_X: "qFresh xs x' X" unfolding x'_def
using assms pickQFreshEnv[of "{x}" "{X}"] by auto
have x'_fresh_rho: "qFreshEnv xs x' rho" unfolding x'_def
using assms pickQFreshEnv[of "{x}" "{X}"] by auto
obtain u where u_not: "u \<notin> {x,x'}" and
u_fresh_X: "qFresh xs u X" and u_fresh_rho: "qFreshEnv xs u rho"
using good_X good_rho obtain_qFreshEnv[of "{x,x'}" "{X}" "{rho}"] by auto
let ?Xx'x = "X #[[x' \<and> x]]_xs"
let ?Xx'x_rho = "?Xx'x #[[rho]]"
let ?Xx'x_rho_ux' = "?Xx'x_rho #[[u \<and> x']]_xs"
let ?Xx'xux' = "?Xx'x #[[u \<and> x']]_xs"
let ?Xx'xux'_rho = "?Xx'xux' #[[rho]]"
let ?Xux = "X #[[u \<and> x]]_xs"
let ?Xux_rho = "?Xux #[[rho]]"
let ?Xrho = "X #[[rho]]"
let ?Xrho_ux = "?Xrho #[[u \<and> x]]_xs"
(* Good: *)
have good_Xx'x: "qGood ?Xx'x" using good_X qSwap_preserves_qGood by auto
hence good_Xx'x_rho: "qGood ?Xx'x_rho" using good_rho qPsubst_preserves_qGood by auto
hence good_Xx'x_rho_ux': "qGood ?Xx'x_rho_ux'" using qSwap_preserves_qGood by auto
have good_Xx'xux': "qGood ?Xx'xux'" using good_Xx'x qSwap_preserves_qGood by auto
(* Fresh: *)
have u_fresh_Xx'x: "qFresh xs u ?Xx'x"
using u_fresh_X u_not by(auto simp add: qSwap_preserves_qFresh_distinct)
hence u_fresh_Xx'x_rho: "qFresh xs u ?Xx'x_rho"
using good_rho good_Xx'x u_fresh_rho by(auto simp add: qPsubst_preserves_qFresh)
have u_fresh_Xrho: "qFresh xs u ?Xrho"
using good_rho good_X u_fresh_X u_fresh_rho by(auto simp add: qPsubst_preserves_qFresh)
(* Main proof: *) -
have "?Xx'x_rho_ux' #= ?Xx'xux'_rho"
using good_Xx'x good_rho u_fresh_rho x'_fresh_rho
using alpha_qFreshEnv_qSwap_qPsubst_commute alpha_sym by blast
moreover
{have "?Xx'xux' #= ?Xux"
using good_X u_fresh_X x'_fresh_X by (auto simp add: alpha_qFresh_qSwap_compose)
hence "?Xx'xux'_rho #= ?Xux_rho"
using good_Xx'xux' good_rho qPsubst_preserves_alpha1 by auto
}
moreover have "?Xux_rho #= ?Xrho_ux"
using good_X good_rho u_fresh_rho x_fresh_rho
using alpha_qFreshEnv_qSwap_qPsubst_commute by blast
ultimately have "?Xx'x_rho_ux' #= ?Xrho_ux"
using good_Xx'x_rho_ux' alpha_trans by blast
thus ?thesis apply (simp add: Let_def del: alpha.simps alphaAbs.simps)
unfolding x'_rev using good_Xx'x_rho
using u_fresh_Xx'x_rho u_fresh_Xrho by (auto simp only: alphaAbs_qAbs_iff_ex_qFresh)
qed
end (* context FixVars *)
end
diff --git a/thys/Binding_Syntax_Theory/QuasiTerms_PickFresh_Alpha.thy b/thys/Binding_Syntax_Theory/QuasiTerms_PickFresh_Alpha.thy
--- a/thys/Binding_Syntax_Theory/QuasiTerms_PickFresh_Alpha.thy
+++ b/thys/Binding_Syntax_Theory/QuasiTerms_PickFresh_Alpha.thy
@@ -1,2119 +1,2119 @@
-section {* Availability of Fresh Variables and Alpha-Equivalence *}
+section \<open>Availability of Fresh Variables and Alpha-Equivalence\<close>
theory QuasiTerms_PickFresh_Alpha
imports QuasiTerms_Swap_Fresh
begin
-text{* Here we define good quasi-terms and alpha-equivalence on quasi-terms,
+text\<open>Here we define good quasi-terms and alpha-equivalence on quasi-terms,
and prove relevant properties
such as the ability to pick fresh variables for good
quasi-terms and the fact that alpha is indeed an equivalence
and is compatible with all the operators.
We do most of the work on freshness and alpha-equivalence
unsortedly, for raw quasi-terms. (And we do it in such a way that
it then applies immediately to sorted quasi-terms.)
We do need sortedness of variables (as well as a cardinality
assumption), however, for alpha-equivalence to have the desired properties.
-Therefore we work in a locale. *}
+Therefore we work in a locale.\<close>
-subsection {* The FixVars locale *}
+subsection \<open>The FixVars locale\<close>
definition var_infinite where
"var_infinite (_ :: 'var) ==
infinite (UNIV :: 'var set)"
definition var_regular where
"var_regular (_ :: 'var) ==
regular |UNIV :: 'var set|"
definition varSort_lt_var where
"varSort_lt_var (_ :: 'varSort) (_ :: 'var) ==
|UNIV :: 'varSort set| <o |UNIV :: 'var set|"
locale FixVars =
fixes dummyV :: 'var and dummyVS :: 'varSort
assumes var_infinite: "var_infinite (undefined :: 'var)"
and var_regular: "var_regular (undefined :: 'var)"
and varSort_lt_var: "varSort_lt_var (undefined :: 'varSort) (undefined :: 'var)"
(*********************************************)
context FixVars
begin
lemma varSort_lt_var_INNER:
"|UNIV :: 'varSort set| <o |UNIV :: 'var set|"
using varSort_lt_var
unfolding varSort_lt_var_def by simp
lemma varSort_le_Var:
"|UNIV :: 'varSort set| \<le>o |UNIV :: 'var set|"
using varSort_lt_var_INNER ordLess_imp_ordLeq by auto
theorem var_infinite_INNER: "infinite (UNIV :: 'var set)"
using var_infinite unfolding var_infinite_def by simp
theorem var_regular_INNER: "regular |UNIV :: 'var set|"
using var_regular unfolding var_regular_def by simp
theorem infinite_var_regular_INNER:
"infinite (UNIV :: 'var set) \<and> regular |UNIV :: 'var set|"
by (simp add: var_infinite_INNER var_regular_INNER)
(* Below and elsewhere: We want both full generality and usefulness in one single
theorem, and therefore we include as a disjunction the general condition w.r.t. the variable cardinality
and the stronger (most often needed) one of finiteness. This way, we need not
invoke variable infiniteness each time we want to use the finiteness. *)
theorem finite_ordLess_var:
"( |S| <o |UNIV :: 'var set| \<or> finite S) = ( |S| <o |UNIV :: 'var set| )"
by (auto simp add: var_infinite_INNER finite_ordLess_infinite2)
-subsection {* Good quasi-terms *}
+subsection \<open>Good quasi-terms\<close>
-text {* Essentially, good quasi-term items
+text \<open>Essentially, good quasi-term items
will be those with meaningful binders and
not too many variables. Good quasi-terms are a concept intermediate
between (raw) quasi-terms and sorted quasi-terms. This concept was chosen to be strong
enough to facilitate proofs of most of the desired properties of alpha-equivalence, avoiding,
{\em for most of the hard part of the work},
the overhead of sortedness. Since we later prove that quasi-terms
are good,
- all the results are then immediately transported to a sorted setting. *}
+ all the results are then immediately transported to a sorted setting.\<close>
function
qGood :: "('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> bool"
and
qGoodAbs :: "('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> bool"
where
"qGood (qVar xs x) = True"
|
"qGood (qOp delta inp binp) =
(liftAll qGood inp \<and> liftAll qGoodAbs binp \<and>
|{i. inp i \<noteq> None}| <o |UNIV :: 'var set| \<and>
|{i. binp i \<noteq> None}| <o |UNIV :: 'var set| )"
|
"qGoodAbs (qAbs xs x X) = qGood X"
by (pat_completeness, auto)
termination
apply(relation qTermLess)
apply(simp_all add: qTermLess_wf)
by(auto simp add: qTermLess_def)
fun qGoodItem :: "('index,'bindex,'varSort,'var,'opSym)qTermItem \<Rightarrow> bool" where
"qGoodItem (Inl qX) = qGood qX"
|
"qGoodItem (Inr qA) = qGoodAbs qA"
lemma qSwapAll_preserves_qGoodAll1:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs x y
shows
"(qGood X \<longrightarrow> qGood (X #[[x \<and> y]]_zs)) \<and>
(qGoodAbs A \<longrightarrow> qGoodAbs (A $[[x \<and> y]]_zs))"
apply(rule qTerm_induct[of _ _ X A])
apply(simp_all add: sw_def)
unfolding lift_def liftAll_def apply auto
apply(case_tac "inp i", auto)
apply(case_tac "binp i", auto)
proof-
fix inp::"('index,('index,'bindex,'varSort,'var,'opSym)qTerm)input" and zs xs x y
let ?K1 = "{i. \<exists>X. inp i = Some X}"
let ?K2 = "{i. \<exists>X. (case inp i of None \<Rightarrow> None | Some X \<Rightarrow> Some (X #[[x \<and> y]]_zs))
= Some X}"
assume "|?K1| <o |UNIV :: 'var set|"
moreover have "?K1 = ?K2" by(auto, case_tac "inp x", auto)
ultimately show "|?K2| <o |UNIV :: 'var set|" by simp
next
fix binp::"('bindex,('index,'bindex,'varSort,'var,'opSym)qAbs)input" and zs xs x y
let ?K1 = "{i. \<exists>A. binp i = Some A}"
let ?K2 = "{i. \<exists>A. (case binp i of None \<Rightarrow> None | Some A \<Rightarrow> Some (A $[[x \<and> y]]_zs))
= Some A}"
assume "|?K1| <o |UNIV :: 'var set|"
moreover have "?K1 = ?K2" by(auto, case_tac "binp x", auto)
ultimately show "|?K2| <o |UNIV :: 'var set|" by simp
qed
corollary qSwap_preserves_qGood1:
"qGood X \<Longrightarrow> qGood (X #[[x \<and> y]]_zs)"
by(simp add: qSwapAll_preserves_qGoodAll1)
corollary qSwapAbs_preserves_qGoodAbs1:
"qGoodAbs A \<Longrightarrow> qGoodAbs (A $[[x \<and> y]]_zs)"
by(simp add: qSwapAll_preserves_qGoodAll1)
lemma qSwap_preserves_qGood2:
assumes "qGood(X #[[x \<and> y]]_zs)"
shows "qGood X"
by (metis assms qSwap_involutive qSwap_preserves_qGood1)
lemma qSwapAbs_preserves_qGoodAbs2:
assumes "qGoodAbs(A $[[x \<and> y]]_zs)"
shows "qGoodAbs A"
by (metis assms qSwapAbs_preserves_qGoodAbs1 qSwapAll_involutive)
lemma qSwap_preserves_qGood: "(qGood (X #[[x \<and> y]]_zs)) = (qGood X)"
using qSwap_preserves_qGood1 qSwap_preserves_qGood2 by blast
lemma qSwapAbs_preserves_qGoodAbs:
"(qGoodAbs (A $[[x \<and> y]]_zs)) = (qGoodAbs A)"
using qSwapAbs_preserves_qGoodAbs1 qSwapAbs_preserves_qGoodAbs2 by blast
lemma qSwap_twice_preserves_qGood:
"(qGood ((X #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs')) = (qGood X)"
by (simp add: qSwap_preserves_qGood)
lemma qSwapped_preserves_qGood:
"(X,Y) \<in> qSwapped \<Longrightarrow> qGood Y = qGood X"
apply (induct rule: qSwapped.induct)
using qSwap_preserves_qGood by auto
lemma qGood_qTerm_templateInduct[case_names Rel Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm"
and A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and phi phiAbs rel
assumes
REL: "\<And> X Y. \<lbrakk>qGood X; (X,Y) \<in> rel\<rbrakk> \<Longrightarrow> qGood Y \<and> qSkel Y = qSkel X" and
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>|{i. inp i \<noteq> None}| <o |UNIV :: 'var set|;
|{i. binp i \<noteq> None}| <o |UNIV :: 'var set|;
liftAll (\<lambda>X. qGood X \<and> phi X) inp;
liftAll (\<lambda>A. qGoodAbs A \<and> phiAbs A) binp\<rbrakk>
\<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>qGood X; \<And> Y. (X,Y) \<in> rel \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (qAbs xs x X)"
shows
"(qGood X \<longrightarrow> phi X) \<and> (qGoodAbs A \<longrightarrow> phiAbs A)"
apply(induct rule: qTerm_templateInduct[of "{(X,Y). qGood X \<and> (X,Y) \<in> rel}"])
using assms by (simp_all add: liftAll_def)
lemma qGood_qTerm_rawInduct[case_names Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>|{i. inp i \<noteq> None}| <o |UNIV :: 'var set|;
|{i. binp i \<noteq> None}| <o |UNIV :: 'var set|;
liftAll (\<lambda> X. qGood X \<and> phi X) inp;
liftAll (\<lambda> A. qGoodAbs A \<and> phiAbs A) binp\<rbrakk>
\<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>qGood X; phi X\<rbrakk> \<Longrightarrow> phiAbs (qAbs xs x X)"
shows "(qGood X \<longrightarrow> phi X) \<and> (qGoodAbs A \<longrightarrow> phiAbs A)"
apply(induct rule: qGood_qTerm_templateInduct [of Id])
by(simp_all add: assms)
lemma qGood_qTerm_induct[case_names Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>|{i. inp i \<noteq> None}| <o |UNIV :: 'var set|;
|{i. binp i \<noteq> None}| <o |UNIV :: 'var set|;
liftAll (\<lambda> X. qGood X \<and> phi X) inp;
liftAll (\<lambda> A. qGoodAbs A \<and> phiAbs A) binp\<rbrakk>
\<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>qGood X;
\<And> Y. qGood Y \<and> qSkel Y = qSkel X \<Longrightarrow> phi Y;
\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (qAbs xs x X)"
shows
"(qGood X \<longrightarrow> phi X) \<and> (qGoodAbs A \<longrightarrow> phiAbs A)"
apply(induct rule: qGood_qTerm_templateInduct
[of "qSwapped \<union> {(X,Y). qGood Y \<and> qSkel Y = qSkel X}"])
using qSwapped_qSkel qSwapped_preserves_qGood
by(auto simp add: assms)
text "A form specialized for mutual induction
(this time, without the cardinality hypotheses):"
lemma qGood_qTerm_induct_mutual[case_names Var1 Var2 Op1 Op2 Abs1 Abs2]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and phi1 phi2 phiAbs1 phiAbs2
assumes
Var1: "\<And> xs x. phi1 (qVar xs x)" and
Var2: "\<And> xs x. phi2 (qVar xs x)" and
Op1: "\<And> delta inp binp. \<lbrakk>liftAll (\<lambda> X. qGood X \<and> phi1 X) inp;
liftAll (\<lambda> A. qGoodAbs A \<and> phiAbs1 A) binp\<rbrakk>
\<Longrightarrow> phi1 (qOp delta inp binp)" and
Op2: "\<And> delta inp binp. \<lbrakk>liftAll (\<lambda> X. qGood X \<and> phi2 X) inp;
liftAll (\<lambda> A. qGoodAbs A \<and> phiAbs2 A) binp\<rbrakk>
\<Longrightarrow> phi2 (qOp delta inp binp)" and
Abs1: "\<And> xs x X. \<lbrakk>qGood X;
\<And> Y. qGood Y \<and> qSkel Y = qSkel X \<Longrightarrow> phi1 Y;
\<And> Y. qGood Y \<and> qSkel Y = qSkel X \<Longrightarrow> phi2 Y;
\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi1 Y;
\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi2 Y\<rbrakk>
\<Longrightarrow> phiAbs1 (qAbs xs x X)" and
Abs2: "\<And> xs x X. \<lbrakk>qGood X;
\<And> Y. qGood Y \<and> qSkel Y = qSkel X \<Longrightarrow> phi1 Y;
\<And> Y. qGood Y \<and> qSkel Y = qSkel X \<Longrightarrow> phi2 Y;
\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi1 Y;
\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi2 Y;
phiAbs1 (qAbs xs x X)\<rbrakk>
\<Longrightarrow> phiAbs2 (qAbs xs x X)"
shows
"(qGood X \<longrightarrow> (phi1 X \<and> phi2 X)) \<and>
(qGoodAbs A \<longrightarrow> (phiAbs1 A \<and> phiAbs2 A))"
apply(induct rule: qGood_qTerm_induct[of _ _ X A])
by(auto simp add: assms liftAll_and)
-subsection {* The ability to pick fresh variables *}
+subsection \<open>The ability to pick fresh variables\<close>
lemma single_non_qAFreshAll_ordLess_var:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(qGood X \<longrightarrow> |{x. \<not> qAFresh xs x X}| <o |UNIV :: 'var set| ) \<and>
(qGoodAbs A \<longrightarrow> |{x. \<not> qAFreshAbs xs x A}| <o |UNIV :: 'var set| )"
proof(induct rule: qGood_qTerm_rawInduct)
case (Var xs x)
then show ?case using infinite_var_regular_INNER by simp
next
case (Op delta inp binp)
let ?Left = "{x. \<not> qAFresh xs x (qOp delta inp binp)}"
obtain J where J_def: "J = {i. \<exists> X. inp i = Some X}" by blast
let ?S = "\<Union> i \<in> J. {x. \<exists> X. inp i = Some X \<and> \<not> qAFresh xs x X}"
{fix i
obtain K where K_def: "K = {X. inp i = Some X}" by blast
have "finite K" unfolding K_def by (cases "inp i", auto)
hence "|K| <o |UNIV :: 'var set|" using var_infinite_INNER finite_ordLess_infinite2 by auto
moreover have "\<forall> X \<in> K. |{x. \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
unfolding K_def using Op unfolding liftAll_def by simp
ultimately have "|\<Union> X \<in> K. {x. \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
using var_regular_INNER by (simp add: regular_UNION)
moreover
have "{x. \<exists>X. inp i = Some X \<and> \<not> qAFresh xs x X} =
(\<Union> X \<in> K. {x. \<not> qAFresh xs x X})" unfolding K_def by blast
ultimately
have "|{x. \<exists>X. inp i = Some X \<and> \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
by simp
}
moreover have "|J| <o |UNIV :: 'var set|" unfolding J_def
using Op unfolding liftAll_def by simp
ultimately
have 1: "|?S| <o |UNIV :: 'var set|"
using var_regular_INNER by (simp add: regular_UNION)
(* *)
obtain Ja where Ja_def: "Ja = {i. \<exists> A. binp i = Some A}" by blast
let ?Sa = "\<Union> i \<in> Ja. {x. \<exists> A. binp i = Some A \<and> \<not> qAFreshAbs xs x A}"
{fix i
obtain K where K_def: "K = {A. binp i = Some A}" by blast
have "finite K" unfolding K_def by (cases "binp i", auto)
hence "|K| <o |UNIV :: 'var set|" using var_infinite_INNER finite_ordLess_infinite2 by auto
moreover have "\<forall> A \<in> K. |{x. \<not> qAFreshAbs xs x A}| <o |UNIV :: 'var set|"
unfolding K_def using Op unfolding liftAll_def by simp
ultimately have "|\<Union> A \<in> K. {x. \<not> qAFreshAbs xs x A}| <o |UNIV :: 'var set|"
using var_regular_INNER by (simp add: regular_UNION)
moreover
have "{x. \<exists>A. binp i = Some A \<and> \<not> qAFreshAbs xs x A} =
(\<Union> A \<in> K. {x. \<not> qAFreshAbs xs x A})" unfolding K_def by blast
ultimately
have "|{x. \<exists>A. binp i = Some A \<and> \<not> qAFreshAbs xs x A}| <o |UNIV :: 'var set|"
by simp
}
moreover have "|Ja| <o |UNIV :: 'var set|"
unfolding Ja_def using Op unfolding liftAll_def by simp
ultimately have "|?Sa| <o |UNIV :: 'var set|"
using var_regular_INNER by (simp add: regular_UNION)
with 1 have "|?S Un ?Sa| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_ordLess_infinite by auto
moreover have "?Left = ?S Un ?Sa"
by (auto simp: J_def Ja_def liftAll_def )
ultimately show ?case by simp
next
case (Abs xsa x X)
let ?Left = "{xa. xs = xsa \<and> xa = x \<or> \<not> qAFresh xs xa X}"
have "|{x}| <o |UNIV :: 'var set|" by (auto simp add: var_infinite_INNER)
hence "|{x} \<union> {x. \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
using Abs var_infinite_INNER card_of_Un_ordLess_infinite by blast
moreover
{have "?Left \<subseteq> {x} \<union> {x. \<not> qAFresh xs x X}" by blast
hence "|?Left| \<le>o |{x} \<union> {x. \<not> qAFresh xs x X}|" using card_of_mono1 by auto
}
ultimately show ?case using ordLeq_ordLess_trans by auto
qed
corollary single_non_qAFresh_ordLess_var:
"qGood X \<Longrightarrow> |{x. \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
by(simp add: single_non_qAFreshAll_ordLess_var)
corollary single_non_qAFreshAbs_ordLess_var:
"qGoodAbs A \<Longrightarrow> |{x. \<not> qAFreshAbs xs x A}| <o |UNIV :: 'var set|"
by(simp add: single_non_qAFreshAll_ordLess_var)
lemma single_non_qFresh_ordLess_var:
assumes "qGood X"
shows "|{x. \<not> qFresh xs x X}| <o |UNIV :: 'var set|"
using qAFresh_imp_qFresh card_of_mono1 single_non_qAFresh_ordLess_var
ordLeq_ordLess_trans by (metis Collect_mono assms)
lemma single_non_qFreshAbs_ordLess_var:
assumes "qGoodAbs A"
shows "|{x. \<not> qFreshAbs xs x A}| <o |UNIV :: 'var set|"
using qAFreshAll_imp_qFreshAll card_of_mono1 single_non_qAFreshAbs_ordLess_var
ordLeq_ordLess_trans by (metis Collect_mono assms)
lemma non_qAFresh_ordLess_var:
assumes GOOD: "\<forall> X \<in> XS. qGood X" and Var: "|XS| <o |UNIV :: 'var set|"
shows "|{x| x X. X \<in> XS \<and> \<not> qAFresh xs x X}| <o |UNIV :: 'var set|"
proof-
define K and F where "K \<equiv> {x| x X. X \<in> XS \<and> \<not> qAFresh xs x X}"
and "F \<equiv> (\<lambda> X. {x. X \<in> XS \<and> \<not> qAFresh xs x X})"
have "K = (\<Union> X \<in> XS. F X)" unfolding K_def F_def by auto
moreover have "\<forall> X \<in> XS. |F X| <o |UNIV :: 'var set|"
unfolding F_def using GOOD single_non_qAFresh_ordLess_var by auto
ultimately have "|K| <o |UNIV :: 'var set|" using var_regular_INNER Var
by(auto simp add: regular_UNION)
thus ?thesis unfolding K_def .
qed
lemma non_qAFresh_or_in_ordLess_var:
assumes Var: "|V| <o |UNIV :: 'var set|" and "|XS| <o |UNIV :: 'var set|" and "\<forall> X \<in> XS. qGood X"
shows "|{x| x X. (x \<in> V \<or> (X \<in> XS \<and> \<not> qAFresh xs x X))}| <o |UNIV :: 'var set|"
proof-
define J and K where "J \<equiv> {x| x X. (x \<in> V \<or> (X \<in> XS \<and> \<not> qAFresh xs x X))}"
and "K \<equiv> {x| x X. X \<in> XS \<and> \<not> qAFresh xs x X}"
have "J \<subseteq> K \<union> V" unfolding J_def K_def by auto
hence "|J| \<le>o |K \<union> V|" using card_of_mono1 by auto
moreover
{have "|K| <o |UNIV :: 'var set|" unfolding K_def using assms non_qAFresh_ordLess_var by auto
hence "|K \<union> V| <o |UNIV :: 'var set|" using Var var_infinite_INNER card_of_Un_ordLess_infinite by auto
}
ultimately have "|J| <o |UNIV :: 'var set|" using ordLeq_ordLess_trans by blast
thus ?thesis unfolding J_def .
qed
lemma obtain_set_qFresh_card_of:
assumes "|V| <o |UNIV :: 'var set|" and "|XS| <o |UNIV :: 'var set|" and "\<forall> X \<in> XS. qGood X"
shows "\<exists> W. infinite W \<and> W Int V = {} \<and>
(\<forall> x \<in> W. \<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
proof-
define J where "J \<equiv> {x| x X. (x \<in> V \<or> (X \<in> XS \<and> \<not> qAFresh xs x X))}"
let ?W = "UNIV - J"
have "|J| <o |UNIV :: 'var set|"
unfolding J_def using assms non_qAFresh_or_in_ordLess_var by auto
hence "infinite ?W" using var_infinite_INNER subset_ordLeq_diff_infinite[of _ J] by auto
moreover
have "?W \<inter> V = {} \<and> (\<forall> x \<in> ?W. \<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
unfolding J_def using qAFresh_imp_qFresh by fastforce
ultimately show ?thesis by blast
qed
lemma obtain_set_qFresh:
assumes "finite V \<or> |V| <o |UNIV :: 'var set|" and "finite XS \<or> |XS| <o |UNIV :: 'var set|" and
"\<forall> X \<in> XS. qGood X"
shows "\<exists> W. infinite W \<and> W Int V = {} \<and>
(\<forall> x \<in> W. \<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
using assms
by(fastforce simp add: var_infinite_INNER obtain_set_qFresh_card_of)
lemma obtain_qFresh_card_of:
assumes "|V| <o |UNIV :: 'var set|" and "|XS| <o |UNIV :: 'var set|" and "\<forall> X \<in> XS. qGood X"
shows "\<exists> x. x \<notin> V \<and> (\<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
proof-
obtain W where "infinite W" and
*: "W \<inter> V = {} \<and> (\<forall> x \<in> W. \<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
using assms obtain_set_qFresh_card_of by blast
then obtain x where "x \<in> W" using infinite_imp_nonempty by fastforce
thus ?thesis using * by auto
qed
lemma obtain_qFresh:
assumes "finite V \<or> |V| <o |UNIV :: 'var set|" and "finite XS \<or> |XS| <o |UNIV :: 'var set|" and
"\<forall> X \<in> XS. qGood X"
shows "\<exists> x. x \<notin> V \<and> (\<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
using assms
by(fastforce simp add: var_infinite_INNER obtain_qFresh_card_of)
definition pickQFresh where
"pickQFresh xs V XS ==
SOME x. x \<notin> V \<and> (\<forall> X \<in> XS. qAFresh xs x X \<and> qFresh xs x X)"
lemma pickQFresh_card_of:
assumes "|V| <o |UNIV :: 'var set|" and "|XS| <o |UNIV :: 'var set|" and "\<forall> X \<in> XS. qGood X"
shows "pickQFresh xs V XS \<notin> V \<and>
(\<forall> X \<in> XS. qAFresh xs (pickQFresh xs V XS) X \<and> qFresh xs (pickQFresh xs V XS) X)"
unfolding pickQFresh_def apply(rule someI_ex)
using assms obtain_qFresh_card_of by blast
lemma pickQFresh:
assumes "finite V \<or> |V| <o |UNIV :: 'var set|" and "finite XS \<or> |XS| <o |UNIV :: 'var set|" and
"\<forall> X \<in> XS. qGood X"
shows "pickQFresh xs V XS \<notin> V \<and>
(\<forall> X \<in> XS. qAFresh xs (pickQFresh xs V XS) X \<and> qFresh xs (pickQFresh xs V XS) X)"
unfolding pickQFresh_def apply(rule someI_ex)
using assms by(auto simp add: obtain_qFresh)
end (* context FixVars *)
(*****************************************)
-subsection {* Alpha-equivalence *}
+subsection \<open>Alpha-equivalence\<close>
-subsubsection {* Definition *}
+subsubsection \<open>Definition\<close>
definition aux_alpha_ignoreSecond ::
"('index,'bindex,'varSort,'var,'opSym)qTerm * ('index,'bindex,'varSort,'var,'opSym)qTerm +
('index,'bindex,'varSort,'var,'opSym)qAbs * ('index,'bindex,'varSort,'var,'opSym)qAbs
\<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTermItem"
where
"aux_alpha_ignoreSecond K ==
case K of Inl(X,Y) \<Rightarrow> termIn X
|Inr(A,B) \<Rightarrow> absIn A"
lemma aux_alpha_ignoreSecond_qTermLessQSwapped_wf:
"wf(inv_image qTermQSwappedLess aux_alpha_ignoreSecond)"
using qTermQSwappedLess_wf wf_inv_image by auto
(* *)
function
alpha and alphaAbs
where
"alpha (qVar xs x) (qVar xs' x') \<longleftrightarrow> xs = xs' \<and> x = x'"
|
"alpha (qOp delta inp binp) (qOp delta' inp' binp') \<longleftrightarrow>
delta = delta' \<and> sameDom inp inp' \<and> sameDom binp binp' \<and>
liftAll2 alpha inp inp' \<and>
liftAll2 alphaAbs binp binp'"
|
"alpha (qVar xs x) (qOp delta' inp' binp') \<longleftrightarrow> False"
|
"alpha (qOp delta inp binp) (qVar xs' x') \<longleftrightarrow> False"
|
"alphaAbs (qAbs xs x X) (qAbs xs' x' X') \<longleftrightarrow>
xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs' y X' \<and>
alpha (X #[[y \<and> x]]_xs) (X' #[[y \<and> x']]_xs'))"
by(pat_completeness, auto)
termination
apply(relation "inv_image qTermQSwappedLess aux_alpha_ignoreSecond")
apply(simp add: aux_alpha_ignoreSecond_qTermLessQSwapped_wf)
by(auto simp add: qTermQSwappedLess_def qTermLess_modulo_def
aux_alpha_ignoreSecond_def qSwap_qSwapped)
abbreviation alpha_abbrev (infix "#=" 50) where "X #= Y \<equiv> alpha X Y"
abbreviation alphaAbs_abbrev (infix "$=" 50) where "A $= B \<equiv> alphaAbs A B"
(*********************************************)
context FixVars
begin
-subsubsection {* Simplification and elimination rules *}
+subsubsection \<open>Simplification and elimination rules\<close>
lemma alpha_inp_None:
"qOp delta inp binp #= qOp delta' inp' binp' \<Longrightarrow>
(inp i = None) = (inp' i = None)"
by(auto simp add: sameDom_def)
lemma alpha_binp_None:
"qOp delta inp binp #= qOp delta' inp' binp' \<Longrightarrow>
(binp i = None) = (binp' i = None)"
by(auto simp add: sameDom_def)
lemma qVar_alpha_iff:
"qVar xs x #= X' \<longleftrightarrow> X' = qVar xs x"
by(cases X', auto)
lemma alpha_qVar_iff:
"X #= qVar xs' x' \<longleftrightarrow> X = qVar xs' x'"
by(cases X, auto)
lemma qOp_alpha_iff:
"qOp delta inp binp #= X' \<longleftrightarrow>
(\<exists> inp' binp'.
X' = qOp delta inp' binp' \<and> sameDom inp inp' \<and> sameDom binp binp' \<and>
liftAll2 (\<lambda>Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda>A A'. A $= A') binp binp')"
by(cases X') auto
lemma alpha_qOp_iff:
"X #= qOp delta' inp' binp' \<longleftrightarrow>
(\<exists> inp binp. X = qOp delta' inp binp \<and> sameDom inp inp' \<and> sameDom binp binp' \<and>
liftAll2 (\<lambda>Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda>A A'. A $= A') binp binp')"
by(cases X) auto
lemma qAbs_alphaAbs_iff:
"qAbs xs x X $= A' \<longleftrightarrow>
(\<exists> x' y X'. A' = qAbs xs x' X' \<and>
y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs))"
by(cases A') auto
lemma alphaAbs_qAbs_iff:
"A $= qAbs xs' x' X' \<longleftrightarrow>
(\<exists> x y X. A = qAbs xs' x X \<and>
y \<notin> {x,x'} \<and> qAFresh xs' y X \<and> qAFresh xs' y X' \<and>
(X #[[y \<and> x]]_xs') #= (X' #[[y \<and> x']]_xs'))"
by(cases A) auto
-subsubsection {* Basic properties *}
+subsubsection \<open>Basic properties\<close>
-text{* In a nutshell: ``alpha" is included in the kernel of ``qSkel", is
+text\<open>In a nutshell: ``alpha" is included in the kernel of ``qSkel", is
an equivalence on good quasi-terms, preserves goodness,
-and all operators and relations (except ``qAFresh") preserve alpha. *}
+and all operators and relations (except ``qAFresh") preserve alpha.\<close>
lemma alphaAll_qSkelAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(\<forall> X'. X #= X' \<longrightarrow> qSkel X = qSkel X') \<and>
(\<forall> A'. A $= A' \<longrightarrow> qSkelAbs A = qSkelAbs A')"
proof(induction rule: qTerm_induct)
case (Var xs x)
then show ?case unfolding qVar_alpha_iff by simp
next
case (Op delta inp binp)
show ?case proof safe
fix X'
assume "qOp delta inp binp #= X'"
then obtain inp' binp' where X'eq: "X' = qOp delta inp' binp'" and
1: "sameDom inp inp' \<and> sameDom binp binp'" and
2: "liftAll2 (\<lambda> Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda> A A'. A $= A') binp binp'"
unfolding qOp_alpha_iff by auto
from Op.IH 1 2
show "qSkel (qOp delta inp binp) = qSkel X'"
by (simp add: X'eq fun_eq_iff option.case_eq_if
lift_def liftAll_def sameDom_def liftAll2_def)
qed
next
case (Abs xs x X)
show ?case
proof safe
fix A' assume "qAbs xs x X $= A'"
then obtain X' x' y where A'eq: "A' = qAbs xs x' X'" and
*: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)" unfolding qAbs_alphaAbs_iff by auto
moreover have "(X, X #[[y \<and> x]]_xs) \<in> qSwapped" using qSwap_qSwapped by fastforce
ultimately have "qSkel(X #[[y \<and> x]]_xs) = qSkel(X' #[[y \<and> x']]_xs)"
using Abs.IH by blast
hence "qSkel X = qSkel X'" by(auto simp add: qSkel_qSwap)
thus "qSkelAbs (qAbs xs x X) = qSkelAbs A'" unfolding A'eq by simp
qed
qed
corollary alpha_qSkel:
fixes X X' :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
shows "X #= X' \<Longrightarrow> qSkel X = qSkel X'"
by(simp add: alphaAll_qSkelAll)
-text{* Symmetry of alpha is a property that holds for arbitrary
-(not necessarily good) quasi-terms. *}
+text\<open>Symmetry of alpha is a property that holds for arbitrary
+(not necessarily good) quasi-terms.\<close>
lemma alphaAll_sym:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(\<forall> X'. X #= X' \<longrightarrow> X' #= X) \<and> (\<forall> A'. A $= A' \<longrightarrow> A' $= A)"
proof(induction rule: qTerm_induct)
case (Var xs x)
then show ?case unfolding qVar_alpha_iff by simp
next
case (Op delta inp binp)
show ?case proof safe
fix X' assume "qOp delta inp binp #= X'"
then obtain inp' binp' where X': "X' = qOp delta inp' binp'" and
1: "sameDom inp inp' \<and> sameDom binp binp'"
and 2: "liftAll2 (\<lambda>Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda>A A'. A $= A') binp binp'"
unfolding qOp_alpha_iff by auto
thus "X' #= qOp delta inp binp"
unfolding X' using Op.IH 1 2
by (auto simp add: fun_eq_iff option.case_eq_if
lift_def liftAll_def sameDom_def liftAll2_def)
qed
next
case (Abs xs x X)
show ?case proof safe
fix A' assume "qAbs xs x X $= A'"
then obtain x' y X' where
1: "A' = qAbs xs x' X' \<and> y \<notin> {x, x'} \<and> qAFresh xs y X \<and> qAFresh xs y X'" and
"(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
unfolding qAbs_alphaAbs_iff by auto
moreover have "(X, X #[[y \<and> x]]_xs) \<in> qSwapped" by (simp add: qSwap_qSwapped)
ultimately have "(X' #[[y \<and> x']]_xs) #= (X #[[y \<and> x]]_xs)" using Abs.IH by simp
thus "A' $= qAbs xs x X" using 1 by auto
qed
qed
corollary alpha_sym:
fixes X X' :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
shows "X #= X' \<Longrightarrow> X' #= X"
by(simp add: alphaAll_sym)
corollary alphaAbs_sym:
fixes A A' ::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows "A $= A' \<Longrightarrow> A' $= A"
by(simp add: alphaAll_sym)
-text{* Reflexivity does not hold for arbitrary quasi-terms, but onl;y for good
+text\<open>Reflexivity does not hold for arbitrary quasi-terms, but onl;y for good
ones. Indeed, the proof requires picking a fresh variable,
- guaranteed to be possible only if the quasi-term is good. *}
+ guaranteed to be possible only if the quasi-term is good.\<close>
lemma alphaAll_refl:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(qGood X \<longrightarrow> X #= X) \<and> (qGoodAbs A \<longrightarrow> A $= A)"
apply(rule qGood_qTerm_induct, simp_all)
unfolding liftAll_def sameDom_def liftAll2_def apply auto
proof-
fix xs x X
assume "qGood X" and
IH: "\<And>Y. (X,Y) \<in> qSwapped \<Longrightarrow> Y #= Y"
then obtain y where 1: "y \<noteq> x \<and> qAFresh xs y X"
using obtain_qFresh[of "{x}" "{X}"] by auto
hence "(X, X #[[y \<and> x]]_xs) \<in> qSwapped" using qSwap_qSwapped by auto
hence "(X #[[y \<and> x]]_xs) #= (X #[[y \<and> x]]_xs)" using IH by simp
thus "\<exists>y. y \<noteq> x \<and> qAFresh xs y X \<and> (X #[[y \<and> x]]_xs) #= (X #[[y \<and> x]]_xs)"
using 1 by blast
qed
corollary alpha_refl:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
shows "qGood X \<Longrightarrow> X #= X"
by(simp add: alphaAll_refl)
corollary alphaAbs_refl:
fixes A ::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows "qGoodAbs A \<Longrightarrow> A $= A"
by(simp add: alphaAll_refl)
lemma alphaAll_preserves_qGoodAll1:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(qGood X \<longrightarrow> (\<forall> X'. X #= X' \<longrightarrow> qGood X')) \<and>
(qGoodAbs A \<longrightarrow> (\<forall> A'. A $= A' \<longrightarrow> qGoodAbs A'))"
apply(rule qTerm_induct, auto)
unfolding qVar_alpha_iff apply(auto)
proof-
fix delta inp binp X'
assume
IH1: "liftAll (\<lambda>Y. qGood Y \<longrightarrow> (\<forall>Y'. Y #= Y' \<longrightarrow> qGood Y')) inp"
and IH2: "liftAll (\<lambda>A. qGoodAbs A \<longrightarrow> (\<forall>A'. A $= A' \<longrightarrow> qGoodAbs A')) binp"
and *: "liftAll qGood inp" "liftAll qGoodAbs binp"
and **: "|{i. \<exists>Y. inp i = Some Y}| <o |UNIV :: 'var set|"
"|{i. \<exists>A. binp i = Some A}| <o |UNIV :: 'var set|"
and "qOp delta inp binp #= X'"
then obtain inp' binp' where
X'eq: "X' = qOp delta inp' binp'" and
2: "sameDom inp inp' \<and> sameDom binp binp'" and
3: "liftAll2 (\<lambda>Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda>A A'. A $= A') binp binp'"
unfolding qOp_alpha_iff by auto
show "qGood X'"
unfolding X'eq apply simp unfolding liftAll_def apply auto
proof-
fix i Y' assume inp': "inp' i = Some Y'"
then obtain Y where inp: "inp i = Some Y"
using 2 unfolding sameDom_def by fastforce
hence "Y #= Y'" using inp' 3 unfolding liftAll2_def by blast
moreover have "qGood Y" using * inp unfolding liftAll_def by simp
ultimately show "qGood Y'" using IH1 inp unfolding liftAll_def by blast
next
fix i A' assume binp': "binp' i = Some A'"
then obtain A where binp: "binp i = Some A"
using 2 unfolding sameDom_def by fastforce
hence "A $= A'" using binp' 3 unfolding liftAll2_def by blast
moreover have "qGoodAbs A" using * binp unfolding liftAll_def by simp
ultimately show "qGoodAbs A'" using IH2 binp unfolding liftAll_def by blast
next
have "{i. \<exists>Y'. inp' i = Some Y'} = {i. \<exists>Y. inp i = Some Y}"
using 2 unfolding sameDom_def by force
thus "|{i. \<exists>Y'. inp' i = Some Y'}| <o |UNIV :: 'var set|" using ** by simp
next
have "{i. \<exists>A'. binp' i = Some A'} = {i. \<exists>A. binp i = Some A}"
using 2 unfolding sameDom_def by force
thus "|{i. \<exists>A'. binp' i = Some A'}| <o |UNIV :: 'var set|" using ** by simp
qed
next
fix xs x X A'
assume IH: "\<And>Y. (X,Y) \<in> qSwapped \<Longrightarrow> qGood Y \<longrightarrow> (\<forall>X'. Y #= X' \<longrightarrow> qGood X')"
and *: "qGood X" and "qAbs xs x X $= A'"
then obtain x' y X' where "A' = qAbs xs x' X'" and
1: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
unfolding qAbs_alphaAbs_iff by auto
thus "qGoodAbs A'"
proof(auto)
have "(X, X #[[y \<and> x]]_xs) \<in> qSwapped" by(auto simp add: qSwap_qSwapped)
moreover have "qGood(X #[[y \<and> x]]_xs)" using * qSwap_preserves_qGood by auto
ultimately have "qGood(X' #[[y \<and> x']]_xs)" using 1 IH by auto
thus "qGood X'" using * qSwap_preserves_qGood by auto
qed
qed
corollary alpha_preserves_qGood1:
"\<lbrakk>X #= X'; qGood X\<rbrakk> \<Longrightarrow> qGood X'"
using alphaAll_preserves_qGoodAll1 by blast
corollary alphaAbs_preserves_qGoodAbs1:
"\<lbrakk>A $= A'; qGoodAbs A\<rbrakk> \<Longrightarrow> qGoodAbs A'"
using alphaAll_preserves_qGoodAll1 by blast
lemma alpha_preserves_qGood2:
"\<lbrakk>X #= X'; qGood X'\<rbrakk> \<Longrightarrow> qGood X"
using alpha_sym alpha_preserves_qGood1 by blast
lemma alphaAbs_preserves_qGoodAbs2:
"\<lbrakk>A $= A'; qGoodAbs A'\<rbrakk> \<Longrightarrow> qGoodAbs A"
using alphaAbs_sym alphaAbs_preserves_qGoodAbs1 by blast
lemma alpha_preserves_qGood:
"X #= X' \<Longrightarrow> qGood X = qGood X'"
using alpha_preserves_qGood1 alpha_preserves_qGood2 by blast
lemma alphaAbs_preserves_qGoodAbs:
"A $= A' \<Longrightarrow> qGoodAbs A = qGoodAbs A'"
using alphaAbs_preserves_qGoodAbs1 alphaAbs_preserves_qGoodAbs2 by blast
lemma alpha_qSwap_preserves_qGood1:
assumes ALPHA: "(X #[[y \<and> x]]_zs) #= (X' #[[y' \<and> x']]_zs')" and
GOOD: "qGood X"
shows "qGood X'"
proof-
have "qGood(X #[[y \<and> x]]_zs)" using GOOD qSwap_preserves_qGood by auto
hence "qGood (X' #[[y' \<and> x']]_zs')" using ALPHA alpha_preserves_qGood by auto
thus "qGood X'" using qSwap_preserves_qGood by auto
qed
lemma alpha_qSwap_preserves_qGood2:
assumes ALPHA: "(X #[[y \<and> x]]_zs) #= (X' #[[y' \<and> x']]_zs')" and
GOOD': "qGood X'"
shows "qGood X"
proof-
have "qGood(X' #[[y' \<and> x']]_zs')" using GOOD' qSwap_preserves_qGood by auto
hence "qGood (X #[[y \<and> x]]_zs)" using ALPHA alpha_preserves_qGood by auto
thus "qGood X" using qSwap_preserves_qGood by auto
qed
lemma alphaAbs_qSwapAbs_preserves_qGoodAbs2:
assumes ALPHA: "(A $[[y \<and> x]]_zs) $= (A' $[[y' \<and> x']]_zs')" and
GOOD': "qGoodAbs A'"
shows "qGoodAbs A"
proof-
have "qGoodAbs(A' $[[y' \<and> x']]_zs')" using GOOD' qSwapAbs_preserves_qGoodAbs by auto
hence "qGoodAbs (A $[[y \<and> x]]_zs)" using ALPHA alphaAbs_preserves_qGoodAbs by auto
thus "qGoodAbs A" using qSwapAbs_preserves_qGoodAbs by auto
qed
lemma alpha_qSwap_preserves_qGood:
assumes ALPHA: "(X #[[y \<and> x]]_zs) #= (X' #[[y' \<and> x']]_zs')"
shows "qGood X = qGood X'"
using assms alpha_qSwap_preserves_qGood1
alpha_qSwap_preserves_qGood2 by auto
lemma qSwapAll_preserves_alphaAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and z1 z2 zs
shows
"(qGood X \<longrightarrow> (\<forall> X' zs z1 z2. X #= X' \<longrightarrow>
(X #[[z1 \<and> z2]]_zs) #= (X' #[[z1 \<and> z2]]_zs))) \<and>
(qGoodAbs A \<longrightarrow> (\<forall> A' zs z1 z2. A $= A' \<longrightarrow>
(A $[[z1 \<and> z2]]_zs) $= (A' $[[z1 \<and> z2]]_zs)))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
then show ?case unfolding qVar_alpha_iff by simp
next
case (Op delta inp binp)
show ?case proof safe
fix X' zs z1 z2
assume "qOp delta inp binp #= X'" term X' term binp
then obtain inp' binp' where X'eq: "X' = qOp delta inp' binp'" and
1: "sameDom inp inp' \<and> sameDom binp binp'"
and 2: "liftAll2 (\<lambda> Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda> A A'. A $= A') binp binp'"
unfolding qOp_alpha_iff by auto
thus "((qOp delta inp binp) #[[z1 \<and> z2]]_zs) #= (X' #[[z1 \<and> z2]]_zs)"
unfolding X'eq using Op.IH
by (auto simp add: fun_eq_iff option.case_eq_if
lift_def liftAll_def sameDom_def liftAll2_def)
qed
next
case (Abs xs x X)
show ?case proof safe
fix A' zs z1 z2 assume "qAbs xs x X $= A'"
then obtain x' y X' where A': "A' = qAbs xs x' X'" and
y_not: "y \<notin> {x, x'}" and y_fresh: "qAFresh xs y X \<and> qAFresh xs y X'" and
alpha: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
unfolding qAbs_alphaAbs_iff by auto
- hence goodX': "qGood X'" using `qGood X` alpha_qSwap_preserves_qGood by fastforce
+ hence goodX': "qGood X'" using \<open>qGood X\<close> alpha_qSwap_preserves_qGood by fastforce
(* *)
obtain u where u_notin: "u \<notin> {x,x',z1,z2,y}" and
u_freshXX': "qAFresh xs u X \<and> qAFresh xs u X'"
- using `qGood X` goodX' obtain_qFresh[of "{x,x',z1,z2,y}" "{X,X'}"] by auto
+ using \<open>qGood X\<close> goodX' obtain_qFresh[of "{x,x',z1,z2,y}" "{X,X'}"] by auto
hence u_not: "u \<noteq> (x @xs[z1 \<and> z2]_zs) \<and> u \<noteq> (x' @xs[z1 \<and> z2]_zs)"
unfolding sw_def using u_notin by auto
have u_fresh: "qAFresh xs u (X #[[z1 \<and> z2]]_zs) \<and> qAFresh xs u (X' #[[z1 \<and> z2]]_zs)"
using u_freshXX' u_notin by(auto simp add: qSwap_preserves_qAFresh_distinct)
(* *)
have "((X #[[z1 \<and> z2]]_zs) #[[u \<and> (x @xs[z1 \<and> z2]_zs)]]_xs) =
(((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) #[[z1 \<and> z2]]_zs)"
using y_fresh u_freshXX' u_notin by (simp add: qSwap_3commute)
moreover
{have 1: "(X, X #[[y \<and> x]]_xs) \<in> qSwapped" by(simp add: qSwap_qSwapped)
hence "((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) #= ((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs)"
using alpha Abs.IH by auto
moreover have "(X, (X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) \<in> qSwapped"
using 1 by(auto simp add: qSwapped.Swap)
ultimately have "(((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) #[[z1 \<and> z2]]_zs) #=
(((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs) #[[z1 \<and> z2]]_zs)"
using Abs.IH by auto
}
moreover
have "(((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs) #[[z1 \<and> z2]]_zs) =
((X' #[[z1 \<and> z2]]_zs) #[[u \<and> (x' @xs[z1 \<and> z2]_zs)]]_xs)"
using y_fresh u_freshXX' u_notin by (auto simp add: qSwap_3commute)
ultimately have "((X #[[z1 \<and> z2]]_zs) #[[u \<and> (x @xs[z1 \<and> z2]_zs)]]_xs) #=
((X' #[[z1 \<and> z2]]_zs) #[[u \<and> (x' @xs[z1 \<and> z2]_zs)]]_xs)" by simp
thus "((qAbs xs x X) $[[z1 \<and> z2]]_zs) $= (A' $[[z1 \<and> z2]]_zs)"
unfolding A' using u_not u_fresh by auto
qed
qed
corollary qSwap_preserves_alpha:
assumes "qGood X \<or> qGood X'" and "X #= X'"
shows "(X #[[z1 \<and> z2]]_zs) #= (X' #[[z1 \<and> z2]]_zs)"
using assms alpha_preserves_qGood qSwapAll_preserves_alphaAll by blast
corollary qSwapAbs_preserves_alphaAbs:
assumes "qGoodAbs A \<or> qGoodAbs A'" and "A $= A'"
shows "(A $[[z1 \<and> z2]]_zs) $= (A' $[[z1 \<and> z2]]_zs)"
using assms alphaAbs_preserves_qGoodAbs qSwapAll_preserves_alphaAll by blast
lemma qSwap_twice_preserves_alpha:
assumes "qGood X \<or> qGood X'" and "X #= X'"
shows "((X #[[z1 \<and> z2]]_zs) #[[u1 \<and> u2]]_us) #= ((X' #[[z1 \<and> z2]]_zs) #[[u1 \<and> u2]]_us)"
by (simp add: assms qSwap_preserves_alpha qSwap_preserves_qGood)
lemma alphaAll_trans:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows
"(qGood X \<longrightarrow> (\<forall> X' X''. X #= X' \<and> X' #= X'' \<longrightarrow> X #= X'')) \<and>
(qGoodAbs A \<longrightarrow> (\<forall> A' A''. A $= A' \<and> A' $= A'' \<longrightarrow> A $= A''))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
then show ?case by (simp add: qVar_alpha_iff)
next
case (Op delta inp binp)
show ?case proof safe
fix X' X'' assume "qOp delta inp binp #= X'" and *: "X' #= X''"
then obtain inp' binp' where
1: "X' = qOp delta inp' binp'" and
2: "sameDom inp inp' \<and> sameDom binp binp'" and
3: "liftAll2 (\<lambda>Y Y'. Y #= Y') inp inp' \<and>
liftAll2 (\<lambda>A A'. A $= A') binp binp'"
unfolding qOp_alpha_iff by auto
obtain inp'' binp'' where
11: "X'' = qOp delta inp'' binp''" and
22: "sameDom inp' inp'' \<and> sameDom binp' binp''" and
33: "liftAll2 (\<lambda>Y' Y''. Y' #= Y'') inp' inp'' \<and>
liftAll2 (\<lambda>A' A''. A' $= A'') binp' binp''"
using * unfolding 1 unfolding qOp_alpha_iff by auto
have "liftAll2 (#=) inp inp''" unfolding liftAll2_def proof safe
fix i Y Y''
assume inp: "inp i = Some Y" and inp'': "inp'' i = Some Y''"
then obtain Y' where inp': "inp' i = Some Y'"
using 2 unfolding sameDom_def by force
hence "Y #= Y'" using inp 3 unfolding liftAll2_def by blast
moreover have "Y' #= Y''" using inp' inp'' 33 unfolding liftAll2_def by blast
ultimately show "Y #= Y''" using inp Op.IH unfolding liftAll_def by blast
qed
moreover have "liftAll2 ($=) binp binp''" unfolding liftAll2_def proof safe
fix i A A''
assume binp: "binp i = Some A" and binp'': "binp'' i = Some A''"
then obtain A' where binp': "binp' i = Some A'"
using 2 unfolding sameDom_def by force
hence "A $= A'" using binp 3 unfolding liftAll2_def by blast
moreover have "A' $= A''" using binp' binp'' 33 unfolding liftAll2_def by blast
ultimately show "A $= A''" using binp Op.IH unfolding liftAll_def by blast
qed
ultimately show "qOp delta inp binp #= X''"
by (simp add: 11 2 22 sameDom_trans[of inp inp'] sameDom_trans[of binp binp'])
qed
next
case (Abs xs x X)
show ?case proof safe
fix A' A''
assume "qAbs xs x X $= A'" and *: "A' $= A''"
then obtain x' y X' where A': "A' = qAbs xs x' X'" and y_not: "y \<notin> {x, x'}" and
y_fresh: "qAFresh xs y X \<and> qAFresh xs y X'" and
alpha: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
unfolding qAbs_alphaAbs_iff by auto
obtain x'' z X'' where A'': "A'' = qAbs xs x'' X''" and z_not: "z \<notin> {x', x''}" and
z_fresh: "qAFresh xs z X' \<and> qAFresh xs z X''" and
alpha': "(X' #[[z \<and> x']]_xs) #= (X'' #[[z \<and> x'']]_xs)"
using * unfolding A' qAbs_alphaAbs_iff by auto
have goodX': "qGood X'"
- using alpha `qGood X` alpha_qSwap_preserves_qGood by fastforce
+ using alpha \<open>qGood X\<close> alpha_qSwap_preserves_qGood by fastforce
hence goodX'': "qGood X''"
using alpha' alpha_qSwap_preserves_qGood by fastforce
have good: "qGood((X #[[y \<and> x]]_xs)) \<and> qGood((X' #[[z \<and> x']]_xs))"
- using `qGood X` goodX' qSwap_preserves_qGood by auto
+ using \<open>qGood X\<close> goodX' qSwap_preserves_qGood by auto
(* *)
obtain u where u_not: "u \<notin> {x,x',x'',y,z}" and
u_fresh: "qAFresh xs u X \<and> qAFresh xs u X' \<and> qAFresh xs u X''"
- using `qGood X` goodX' goodX''
+ using \<open>qGood X\<close> goodX' goodX''
using obtain_qFresh[of "{x,x',x'',y,z}" "{X, X', X''}"] by auto
(* *)
{have "(X #[[u \<and> x]]_xs) = ((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs)"
using u_fresh y_fresh by (auto simp add: qAFresh_qSwap_compose)
moreover
have "((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) #= ((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs)"
using good alpha qSwap_preserves_alpha by fastforce
moreover have "((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs) = (X' #[[u \<and> x']]_xs)"
using u_fresh y_fresh by (auto simp add: qAFresh_qSwap_compose)
ultimately have "(X #[[u \<and> x]]_xs) #= (X' #[[u \<and> x']]_xs)" by simp
}
moreover
{have "(X' #[[u \<and> x']]_xs) = ((X' #[[z \<and> x']]_xs) #[[u \<and> z]]_xs)"
using u_fresh z_fresh by (auto simp add: qAFresh_qSwap_compose)
moreover
have "((X' #[[z \<and> x']]_xs) #[[u \<and> z]]_xs) #= ((X'' #[[z \<and> x'']]_xs) #[[u \<and> z]]_xs)"
using good alpha' qSwap_preserves_alpha by fastforce
moreover have "((X'' #[[z \<and> x'']]_xs) #[[u \<and> z]]_xs) = (X'' #[[u \<and> x'']]_xs)"
using u_fresh z_fresh by (auto simp add: qAFresh_qSwap_compose)
ultimately have "(X' #[[u \<and> x']]_xs) #= (X'' #[[u \<and> x'']]_xs)" by simp
}
moreover have "(X, X #[[u \<and> x]]_xs) \<in> qSwapped" by (simp add: qSwap_qSwapped)
ultimately have "(X #[[u \<and> x]]_xs) #= (X'' #[[u \<and> x'']]_xs)"
using Abs.IH by blast
thus "qAbs xs x X $= A''"
unfolding A'' using u_not u_fresh by auto
qed
qed
corollary alpha_trans:
assumes "qGood X \<or> qGood X' \<or> qGood X''" "X #= X'" "X' #= X''"
shows "X #= X''"
by (meson alphaAll_trans alpha_preserves_qGood assms)
corollary alphaAbs_trans:
assumes "qGoodAbs A \<or> qGoodAbs A' \<or> qGoodAbs A''"
and "A $= A'" "A' $= A''"
shows "A $= A''"
using assms alphaAbs_preserves_qGoodAbs alphaAll_trans by blast
lemma alpha_trans_twice:
"\<lbrakk>qGood X \<or> qGood X' \<or> qGood X'' \<or> qGood X''';
X #= X'; X' #= X''; X'' #= X'''\<rbrakk> \<Longrightarrow> X #= X'''"
using alpha_trans by blast
lemma alphaAbs_trans_twice:
"\<lbrakk>qGoodAbs A \<or> qGoodAbs A' \<or> qGoodAbs A'' \<or> qGoodAbs A''';
A $= A'; A' $= A''; A'' $= A'''\<rbrakk> \<Longrightarrow> A $= A'''"
using alphaAbs_trans by blast
lemma qAbs_preserves_alpha:
assumes ALPHA: "X #= X'" and GOOD: "qGood X \<or> qGood X'"
shows "qAbs xs x X $= qAbs xs x X'"
proof-
have "qGood X \<and> qGood X'" using GOOD ALPHA by(auto simp add: alpha_preserves_qGood)
then obtain y where y_not: "y \<noteq> x" and
y_fresh: "qAFresh xs y X \<and> qAFresh xs y X'"
using GOOD obtain_qFresh[of "{x}" "{X,X'}"] by auto
hence "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x]]_xs)"
using ALPHA GOOD by(simp add: qSwap_preserves_alpha)
thus ?thesis using y_not y_fresh by auto
qed
corollary qAbs_preserves_alpha2:
assumes ALPHA: "X #= X'" and GOOD: "qGoodAbs(qAbs xs x X) \<or> qGoodAbs (qAbs xs x X')"
shows "qAbs xs x X $= qAbs xs x X'"
using assms by (intro qAbs_preserves_alpha) auto
-subsubsection {* Picking fresh representatives *}
+subsubsection \<open>Picking fresh representatives\<close>
lemma qAbs_alphaAbs_qSwap_qAFresh:
assumes GOOD: "qGood X" and FRESH: "qAFresh ys x' X"
shows "qAbs ys x X $= qAbs ys x' (X #[[x' \<and> x]]_ys)"
proof-
obtain y where 1: "y \<notin> {x,x'}" and 2: "qAFresh ys y X"
using GOOD obtain_qFresh[of "{x,x'}" "{X}"] by auto
hence 3: "qAFresh ys y (X #[[x' \<and> x]]_ys)"
by (auto simp add: qSwap_preserves_qAFresh_distinct)
(* *)
have "(X #[[y \<and> x]]_ys) = ((X #[[x' \<and> x]]_ys) #[[y \<and> x']]_ys)"
using FRESH 2 by (auto simp add: qAFresh_qSwap_compose)
moreover have "qGood (X #[[y \<and> x]]_ys)"
using 1 GOOD qSwap_preserves_qGood by auto
ultimately have "(X #[[y \<and> x]]_ys) #= ((X #[[x' \<and> x]]_ys) #[[y \<and> x']]_ys)"
using alpha_refl by simp
(* *)
thus ?thesis using 1 2 3 assms by auto
qed
lemma qAbs_ex_qAFresh_rep:
assumes GOOD: "qGood X" and FRESH: "qAFresh xs x' X"
shows "\<exists> X'. qGood X' \<and> qAbs xs x X $= qAbs xs x' X'"
proof-
have 1: "qGood (X #[[x' \<and> x]]_xs)" using assms qSwap_preserves_qGood by auto
show ?thesis
apply(rule exI[of _ "X #[[x' \<and> x]]_xs"])
using assms 1 qAbs_alphaAbs_qSwap_qAFresh by fastforce
qed
-subsection {* Properties of swapping and freshness modulo alpha *}
+subsection \<open>Properties of swapping and freshness modulo alpha\<close>
lemma qFreshAll_imp_ex_qAFreshAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs fZs
assumes FIN: "finite V"
shows
"(qGood X \<longrightarrow>
((\<forall> z \<in> V. \<forall> zs \<in> fZs z. qFresh zs z X) \<longrightarrow>
(\<exists> X'. X #= X' \<and> (\<forall> z \<in> V. \<forall> zs \<in> fZs z. qAFresh zs z X')))) \<and>
(qGoodAbs A \<longrightarrow>
((\<forall> z \<in> V. \<forall> zs \<in> fZs z. qFreshAbs zs z A) \<longrightarrow>
(\<exists> A'. A $= A' \<and> (\<forall> z \<in> V. \<forall> zs \<in> fZs z. qAFreshAbs zs z A'))))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
show ?case
by (metis alpha_qVar_iff qAFreshAll_simps(1) qFreshAll_simps(1))
next
case (Op delta inp binp)
show ?case proof safe
assume *: "\<forall>z\<in>V. \<forall>zs\<in>fZs z. qFresh zs z (qOp delta inp binp)"
define phi and phiAbs where
"phi \<equiv> (\<lambda>(Y::('index,'bindex,'varSort,'var,'opSym)qTerm) Y'.
Y #= Y' \<and> (\<forall>z\<in>V. \<forall>zs\<in>fZs z. qAFresh zs z Y'))" and
"phiAbs \<equiv> (\<lambda>(A::('index,'bindex,'varSort,'var,'opSym)qAbs) A'.
A $= A' \<and> (\<forall>z\<in>V. \<forall>zs\<in>fZs z. qAFreshAbs zs z A'))"
have ex_phi: "\<And> i Y. inp i = Some Y \<Longrightarrow> \<exists>Y'. phi Y Y'"
unfolding phi_def using Op.IH * by (auto simp add: liftAll_def)
have ex_phiAbs: "\<And> i A. binp i = Some A \<Longrightarrow> \<exists>A'. phiAbs A A'"
unfolding phiAbs_def using Op.IH * by (auto simp add: liftAll_def)
define inp' and binp' where
"inp' \<equiv> \<lambda> i. case inp i of Some Y \<Rightarrow> Some (SOME Y'. phi Y Y') |None \<Rightarrow> None" and
"binp' \<equiv> \<lambda> i. case binp i of Some A \<Rightarrow> Some (SOME A'. phiAbs A A') |None \<Rightarrow> None"
show "\<exists>X'. qOp delta inp binp #= X' \<and> (\<forall>z\<in>V. \<forall>zs\<in>fZs z. qAFresh zs z X')"
by (intro exI[of _ "qOp delta inp' binp'"])
(auto simp add: inp'_def binp'_def option.case_eq_if sameDom_def liftAll_def liftAll2_def,
(meson ex_phi phi_def ex_phiAbs phiAbs_def some_eq_ex)+)
qed
next
case (Abs xs x X)
show ?case proof safe
assume *: "\<forall>z\<in>V. \<forall>zs\<in>fZs z. qFreshAbs zs z (qAbs xs x X)"
obtain y where y_not_x: "y \<noteq> x" and y_not_V: "y \<notin> V"
and y_afresh: "qAFresh xs y X"
- using FIN `qGood X` obtain_qFresh[of "V \<union> {x}" "{X}"] by auto
+ using FIN \<open>qGood X\<close> obtain_qFresh[of "V \<union> {x}" "{X}"] by auto
hence y_fresh: "qFresh xs y X" using qAFresh_imp_qFresh by fastforce
obtain Y where Y_def: "Y = (X #[[y \<and> x ]]_xs)" by blast
have alphaXY: "qAbs xs x X $= qAbs xs y Y"
- using `qGood X` y_afresh qAbs_alphaAbs_qSwap_qAFresh unfolding Y_def by fastforce
+ using \<open>qGood X\<close> y_afresh qAbs_alphaAbs_qSwap_qAFresh unfolding Y_def by fastforce
have "\<forall>z\<in>V. \<forall>zs \<in> fZs z. qFresh zs z Y"
unfolding Y_def
by (metis * not_equals_and_not_equals_not_in qAFresh_imp_qFresh qAFresh_qSwap_exchange1
qFreshAbs.simps qSwap_preserves_qFresh_distinct y_afresh y_not_V)
moreover have "(X,Y) \<in> qSwapped" unfolding Y_def by(simp add: qSwap_qSwapped)
ultimately obtain Y' where "Y #= Y'" and **: "\<forall>z\<in>V. \<forall>zs \<in> fZs z. qAFresh zs z Y'"
using Abs.IH by blast
- moreover have "qGood Y" unfolding Y_def using `qGood X` qSwap_preserves_qGood by auto
+ moreover have "qGood Y" unfolding Y_def using \<open>qGood X\<close> qSwap_preserves_qGood by auto
ultimately have "qAbs xs y Y $= qAbs xs y Y'" using qAbs_preserves_alpha by blast
- moreover have "qGoodAbs(qAbs xs x X)" using `qGood X` by simp
+ moreover have "qGoodAbs(qAbs xs x X)" using \<open>qGood X\<close> by simp
ultimately have "qAbs xs x X $= qAbs xs y Y'" using alphaXY alphaAbs_trans by blast
moreover have "\<forall>z\<in>V. \<forall>zs \<in> fZs z. qAFreshAbs zs z (qAbs xs y Y')" using ** y_not_V by auto
ultimately show "\<exists>A'. qAbs xs x X $= A' \<and> (\<forall>z\<in>V. \<forall>zs \<in> fZs z. qAFreshAbs zs z A')"
by blast
qed
qed
corollary qFresh_imp_ex_qAFresh:
assumes "finite V" and "qGood X" and "\<forall> z \<in> V. \<forall>zs \<in> fZs z. qFresh zs z X"
shows "\<exists> X'. qGood X' \<and> X #= X' \<and> (\<forall> z \<in> V. \<forall>zs \<in> fZs z. qAFresh zs z X')"
by (metis alphaAll_preserves_qGoodAll1 assms qFreshAll_imp_ex_qAFreshAll)
corollary qFreshAbs_imp_ex_qAFreshAbs:
assumes "finite V" and "qGoodAbs A" and "\<forall> z \<in> V. \<forall>zs \<in> fZs z. qFreshAbs zs z A"
shows "\<exists> A'. qGoodAbs A' \<and> A $= A' \<and> (\<forall> z \<in> V. \<forall>zs \<in> fZs z. qAFreshAbs zs z A')"
by (metis alphaAll_preserves_qGoodAll1 assms qFreshAll_imp_ex_qAFreshAll)
lemma qFresh_imp_ex_qAFresh1:
assumes "qGood X" and "qFresh zs z X"
shows "\<exists> X'. qGood X' \<and> X #= X' \<and> qAFresh zs z X'"
using assms qFresh_imp_ex_qAFresh[of "{z}" _ "undefined(z := {zs})"] by fastforce
lemma qFreshAbs_imp_ex_qAFreshAbs1:
assumes "finite V" and "qGoodAbs A" and "qFreshAbs zs z A"
shows "\<exists> A'. qGoodAbs A' \<and> A $= A' \<and> qAFreshAbs zs z A'"
using assms qFreshAbs_imp_ex_qAFreshAbs[of "{z}" _ "undefined(z := {zs})"] by fastforce
lemma qFresh_imp_ex_qAFresh2:
assumes "qGood X" and "qFresh xs x X" and "qFresh ys y X"
shows "\<exists> X'. qGood X' \<and> X #= X' \<and> qAFresh xs x X' \<and> qAFresh ys y X'"
using assms
qFresh_imp_ex_qAFresh[of "{x}" _ "undefined(x := {xs,ys})"]
qFresh_imp_ex_qAFresh[of "{x,y}" _ "(undefined(x := {xs}))(y := {ys})"]
by (cases "x = y") auto
lemma qFreshAbs_imp_ex_qAFreshAbs2:
assumes "finite V" and "qGoodAbs A" and "qFreshAbs xs x A" and "qFreshAbs ys y A"
shows "\<exists> A'. qGoodAbs A' \<and> A $= A' \<and> qAFreshAbs xs x A' \<and> qAFreshAbs ys y A'"
using assms
qFreshAbs_imp_ex_qAFreshAbs[of "{x}" _ "undefined(x := {xs,ys})"]
qFreshAbs_imp_ex_qAFreshAbs[of "{x,y}" _ "(undefined(x := {xs}))(y := {ys})"]
by (cases "x = y") auto
lemma qAFreshAll_qFreshAll_preserves_alphaAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs z
shows
"(qGood X \<longrightarrow>
(qAFresh zs z X \<longrightarrow> (\<forall> X'. X #= X' \<longrightarrow> qFresh zs z X'))) \<and>
(qGoodAbs A \<longrightarrow>
(qAFreshAbs zs z A \<longrightarrow> (\<forall> A'. A $= A' \<longrightarrow> qFreshAbs zs z A')))"
proof(induction rule: qGood_qTerm_induct)
case (Var xs x)
thus ?case unfolding qVar_alpha_iff by simp
next
case (Op delta inp binp)
show ?case proof safe
fix X'
assume afresh: "qAFresh zs z (qOp delta inp binp)"
and "qOp delta inp binp #= X'"
then obtain inp' and binp' where X'eq: "X' = qOp delta inp' binp'" and
*: "(\<forall>i. (inp i = None) = (inp' i = None)) \<and>
(\<forall>i. (binp i = None) = (binp' i = None))" and
**: "(\<forall>i Y Y'. inp i = Some Y \<and> inp' i = Some Y' \<longrightarrow> Y #= Y') \<and>
(\<forall>i A A'. binp i = Some A \<and> binp' i = Some A' \<longrightarrow> A $= A')"
unfolding qOp_alpha_iff sameDom_def liftAll2_def by auto
{fix i Y' assume inp': "inp' i = Some Y'"
then obtain Y where inp: "inp i = Some Y" using * by fastforce
hence "Y #= Y'" using inp' ** by blast
hence "qFresh zs z Y'" using inp Op.IH afresh by (auto simp: liftAll_def)
}
moreover
{fix i A' assume binp': "binp' i = Some A'"
then obtain A where binp: "binp i = Some A" using * by fastforce
hence "A $= A'" using binp' ** by blast
hence "qFreshAbs zs z A'" using binp Op.IH afresh by (auto simp: liftAll_def)
}
ultimately show "qFresh zs z X'"
unfolding X'eq apply simp unfolding liftAll_def by simp
qed
next
case (Abs xs x X)
show ?case proof safe
fix A'
assume "qAbs xs x X $= A'" and afresh: "qAFreshAbs zs z (qAbs xs x X)"
then obtain x' y X' where A'eq: "A' = qAbs xs x' X'" and
ynot: "y \<notin> {x, x'}" and y_afresh: "qAFresh xs y X \<and> qAFresh xs y X'" and
alpha: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
unfolding qAbs_alphaAbs_iff by auto
(* *)
- have goodXxy: "qGood(X #[[y \<and> x]]_xs)" using `qGood X` qSwap_preserves_qGood by auto
+ have goodXxy: "qGood(X #[[y \<and> x]]_xs)" using \<open>qGood X\<close> qSwap_preserves_qGood by auto
hence goodX'yx': "qGood(X' #[[y \<and> x']]_xs)" using alpha alpha_preserves_qGood by auto
hence "qGood X'" using qSwap_preserves_qGood by auto
then obtain u where u_afresh: "qAFresh xs u X \<and> qAFresh xs u X'"
- and unot: "u \<notin> {x,x',z}" using `qGood X` obtain_qFresh[of "{x,x',z}" "{X,X'}"] by auto
+ and unot: "u \<notin> {x,x',z}" using \<open>qGood X\<close> obtain_qFresh[of "{x,x',z}" "{X,X'}"] by auto
(* *)
have "(X #[[u \<and> x]]_xs) = ((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) \<and>
(X' #[[u \<and> x']]_xs) = ((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs)"
using u_afresh y_afresh qAFresh_qSwap_compose by fastforce
moreover have "((X #[[y \<and> x]]_xs) #[[u \<and> y]]_xs) #= ((X' #[[y \<and> x']]_xs) #[[u \<and> y]]_xs)"
using goodXxy goodX'yx' alpha qSwap_preserves_alpha by fastforce
ultimately have alpha: "(X #[[u \<and> x]]_xs) #= (X' #[[u \<and> x']]_xs)" by simp
(* *)
moreover have "(X, X #[[u \<and> x]]_xs) \<in> qSwapped" by (simp add: qSwap_qSwapped)
moreover have "qAFresh zs z (X #[[u \<and> x]]_xs)"
using unot afresh by(auto simp add: qSwap_preserves_qAFresh_distinct)
ultimately have "qFresh zs z (X' #[[u \<and> x']]_xs)" using afresh Abs.IH by simp
hence "zs = xs \<and> z = x' \<or> qFresh zs z X'"
using unot afresh qSwap_preserves_qFresh_distinct[of zs xs z] by fastforce
thus "qFreshAbs zs z A'" unfolding A'eq by simp
qed
qed
corollary qAFresh_qFresh_preserves_alpha:
"\<lbrakk>qGood X; qAFresh zs z X; X #= X'\<rbrakk> \<Longrightarrow> qFresh zs z X'"
by(simp add: qAFreshAll_qFreshAll_preserves_alphaAll)
corollary qAFreshAbs_imp_qFreshAbs_preserves_alphaAbs:
"\<lbrakk>qGoodAbs A; qAFreshAbs zs z A; A $= A'\<rbrakk> \<Longrightarrow> qFreshAbs zs z A'"
by(simp add: qAFreshAll_qFreshAll_preserves_alphaAll)
lemma qFresh_preserves_alpha1:
assumes "qGood X" and "qFresh zs z X" and "X #= X'"
shows "qFresh zs z X'"
by (meson alpha_sym alpha_trans assms qAFresh_qFresh_preserves_alpha qFresh_imp_ex_qAFresh1)
lemma qFreshAbs_preserves_alphaAbs1:
assumes "qGoodAbs A" and "qFreshAbs zs z A" and "A $= A'"
shows "qFreshAbs zs z A'"
by (meson alphaAbs_sym alphaAbs_trans assms finite.emptyI
qAFreshAbs_imp_qFreshAbs_preserves_alphaAbs qFreshAbs_imp_ex_qAFreshAbs1)
lemma qFresh_preserves_alpha:
assumes "qGood X \<or> qGood X'" and "X #= X'"
shows "qFresh zs z X \<longleftrightarrow> qFresh zs z X'"
using alpha_preserves_qGood alpha_sym assms qFresh_preserves_alpha1 by blast
lemma qFreshAbs_preserves_alphaAbs:
assumes "qGoodAbs A \<or> qGoodAbs A'" and "A $= A'"
shows "qFreshAbs zs z A = qFreshAbs zs z A'"
using assms alphaAbs_preserves_qGoodAbs alphaAbs_sym qFreshAbs_preserves_alphaAbs1 by blast
lemma alpha_qFresh_qSwap_id:
assumes "qGood X" and "qFresh zs z1 X" and "qFresh zs z2 X"
shows "(X #[[z1 \<and> z2]]_zs) #= X"
proof-
obtain X' where 1: "X #= X'" and "qAFresh zs z1 X' \<and> qAFresh zs z2 X'"
using assms qFresh_imp_ex_qAFresh2 by force
hence "(X' #[[z1 \<and> z2]]_zs) = X'" using qAFresh_qSwap_id by auto
moreover have "(X #[[z1 \<and> z2]]_zs) #= (X' #[[z1 \<and> z2]]_zs)"
using assms 1 by (auto simp add: qSwap_preserves_alpha)
moreover have "X' #= X" using 1 alpha_sym by auto
moreover have "qGood(X #[[z1 \<and> z2]]_zs)" using assms qSwap_preserves_qGood by auto
ultimately show ?thesis using alpha_trans by auto
qed
lemma alphaAbs_qFreshAbs_qSwapAbs_id:
assumes "qGoodAbs A" and "qFreshAbs zs z1 A" and "qFreshAbs zs z2 A"
shows "(A $[[z1 \<and> z2]]_zs) $= A"
proof-
obtain A' where 1: "A $= A'" and "qAFreshAbs zs z1 A' \<and> qAFreshAbs zs z2 A'"
using assms qFreshAbs_imp_ex_qAFreshAbs2 by force
hence "(A' $[[z1 \<and> z2]]_zs) = A'" using qAFreshAll_qSwapAll_id by fastforce
moreover have "(A $[[z1 \<and> z2]]_zs) $= (A' $[[z1 \<and> z2]]_zs)"
using assms 1 by (auto simp add: qSwapAbs_preserves_alphaAbs)
moreover have "A' $= A" using 1 alphaAbs_sym by auto
moreover have "qGoodAbs (A $[[z1 \<and> z2]]_zs)" using assms qSwapAbs_preserves_qGoodAbs by auto
ultimately show ?thesis using alphaAbs_trans by auto
qed
lemma alpha_qFresh_qSwap_compose:
assumes GOOD: "qGood X" and "qFresh zs y X" and "qFresh zs z X"
shows "((X #[[y \<and> x]]_zs) #[[z \<and> y]]_zs) #= (X #[[z \<and> x]]_zs)"
proof-
obtain X' where 1: "X #= X'" and "qAFresh zs y X' \<and> qAFresh zs z X'"
using assms qFresh_imp_ex_qAFresh2 by force
hence "((X' #[[y \<and> x]]_zs) #[[z \<and> y]]_zs) = (X' #[[z \<and> x]]_zs)"
using qAFresh_qSwap_compose by auto
moreover have "((X #[[y \<and> x]]_zs) #[[z \<and> y]]_zs) #= ((X' #[[y \<and> x]]_zs) #[[z \<and> y]]_zs)"
using GOOD 1 by (auto simp add: qSwap_twice_preserves_alpha)
moreover have "(X' #[[z \<and> x]]_zs) #= (X #[[z \<and> x]]_zs)"
using GOOD 1 by (auto simp add: qSwap_preserves_alpha alpha_sym)
moreover have "qGood ((X #[[y \<and> x]]_zs) #[[z \<and> y]]_zs)"
using GOOD by (auto simp add: qSwap_twice_preserves_qGood)
ultimately show ?thesis using alpha_trans by auto
qed
lemma qAbs_alphaAbs_qSwap_qFresh:
assumes GOOD: "qGood X" and FRESH: "qFresh xs x' X"
shows "qAbs xs x X $= qAbs xs x' (X #[[x' \<and> x]]_xs)"
proof-
obtain Y where good_Y: "qGood Y" and alpha: "X #= Y" and fresh_Y: "qAFresh xs x' Y"
using assms qFresh_imp_ex_qAFresh1 by fastforce
hence "qAbs xs x Y $= qAbs xs x' (Y #[[x' \<and> x]]_xs)"
using qAbs_alphaAbs_qSwap_qAFresh by blast
moreover have "qAbs xs x X $= qAbs xs x Y"
using GOOD alpha qAbs_preserves_alpha by fastforce
moreover
{have "Y #[[x' \<and> x]]_xs #= X #[[x' \<and> x]]_xs"
using GOOD alpha by (auto simp add: qSwap_preserves_alpha alpha_sym)
moreover have "qGood (Y #[[x' \<and> x]]_xs)" using good_Y qSwap_preserves_qGood by auto
ultimately have "qAbs xs x' (Y #[[x' \<and> x]]_xs) $= qAbs xs x' (X #[[x' \<and> x]]_xs)"
using qAbs_preserves_alpha by blast
}
moreover have "qGoodAbs (qAbs xs x X)" using GOOD by simp
ultimately show ?thesis using alphaAbs_trans by blast
qed
lemma alphaAbs_qAbs_ex_qFresh_rep:
assumes GOOD: "qGood X" and FRESH: "qFresh xs x' X"
shows "\<exists> X'. (X,X') \<in> qSwapped \<and> qGood X' \<and> qAbs xs x X $= qAbs xs x' X'"
proof-
have 1: "qGood (X #[[x' \<and> x]]_xs)" using assms qSwap_preserves_qGood by auto
have 2: "(X,X #[[x' \<and> x]]_xs) \<in> qSwapped" by(simp add: qSwap_qSwapped)
show ?thesis
apply(rule exI[of _ "X #[[x' \<and> x]]_xs"])
using assms 1 2 qAbs_alphaAbs_qSwap_qFresh by fastforce
qed
-subsection {* Alternative statements of the alpha-clause for bound arguments *}
+subsection \<open>Alternative statements of the alpha-clause for bound arguments\<close>
-text{* These alternatives are essentially variations with forall/exists and and qFresh/qAFresh. *}
+text\<open>These alternatives are essentially variations with forall/exists and and qFresh/qAFresh.\<close>
(* FIXME: In this subsection I may have proved quite a few useless things. *)
-subsubsection {* First for ``qAFresh" *}
+subsubsection \<open>First for ``qAFresh"\<close>
definition alphaAbs_ex_equal_or_qAFresh
where
"alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. (y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X') \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_ex_qAFresh
where
"alphaAbs_ex_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. qAFresh xs y X \<and> qAFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_ex_distinct_qAFresh
where
"alphaAbs_ex_distinct_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_equal_or_qAFresh
where
"alphaAbs_all_equal_or_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. (y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_qAFresh
where
"alphaAbs_all_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_distinct_qAFresh
where
"alphaAbs_all_distinct_qAFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
lemma alphaAbs_weakestEx_imp_strongestAll:
assumes GOOD_X: "qGood X" and "alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
shows "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
proof-
obtain y where xs: "xs = xs'" and
yEqFresh: "(y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X')" and
alpha: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
using assms by (auto simp add: alphaAbs_ex_equal_or_qAFresh_def)
show ?thesis
using xs unfolding alphaAbs_all_equal_or_qAFresh_def
proof(intro conjI allI impI, simp)
fix z assume zFresh: "(z = x \<or> qAFresh xs z X) \<and> (z = x' \<or> qAFresh xs z X')"
have "(X #[[z \<and> x]]_xs) = ((X #[[y \<and> x]]_xs) #[[z \<and> y]]_xs)"
proof(cases "z = x")
assume Case1: "z = x"
thus ?thesis by(auto simp add: qSwap_sym)
next
assume Case2: "z \<noteq> x"
hence z_fresh: "qAFresh xs z X" using zFresh by auto
show ?thesis
proof(cases "y = x")
assume Case21: "y = x"
show ?thesis unfolding Case21 by simp
next
assume Case22: "y \<noteq> x"
hence "qAFresh xs y X" using yEqFresh by auto
thus ?thesis using z_fresh qAFresh_qSwap_compose by fastforce
qed
qed
moreover
have "(X' #[[z \<and> x']]_xs) = ((X' #[[y \<and> x']]_xs) #[[z \<and> y]]_xs)"
proof(cases "z = x'")
assume Case1: "z = x'"
thus ?thesis by(auto simp add: qSwap_sym)
next
assume Case2: "z \<noteq> x'"
hence z_fresh: "qAFresh xs z X'" using zFresh by auto
show ?thesis
proof(cases "y = x'")
assume Case21: "y = x'"
show ?thesis unfolding Case21 by simp
next
assume Case22: "y \<noteq> x'"
hence "qAFresh xs y X'" using yEqFresh by auto
thus ?thesis using z_fresh qAFresh_qSwap_compose by fastforce
qed
qed
moreover
{have "qGood (X #[[y \<and> x]]_xs)" using GOOD_X qSwap_preserves_qGood by auto
hence "((X #[[y \<and> x]]_xs) #[[z \<and> y]]_xs) #= ((X' #[[y \<and> x']]_xs) #[[z \<and> y]]_xs)"
using alpha qSwap_preserves_alpha by fastforce
}
ultimately show "(X #[[z \<and> x]]_xs) #= (X' #[[z \<and> x']]_xs)" by simp
qed
qed
lemma alphaAbs_weakestAll_imp_strongestEx:
assumes GOOD: "qGood X" "qGood X'"
and "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
shows "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
proof-
have xs: "xs = xs'"
using assms unfolding alphaAbs_all_distinct_qAFresh_def by auto
obtain y where y_not: "y \<notin> {x,x'}" and
yFresh: "qAFresh xs y X \<and> qAFresh xs y X'"
using GOOD obtain_qFresh[of "{x,x'}" "{X,X'}"] by auto
hence "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
using assms unfolding alphaAbs_all_distinct_qAFresh_def by auto
thus ?thesis unfolding alphaAbs_ex_distinct_qAFresh_def using xs y_not yFresh by auto
qed
(* Note: I do not infer the following from the previous two because
I do not want to use "qGood X'": *)
lemma alphaAbs_weakestEx_imp_strongestEx:
assumes GOOD: "qGood X"
and "alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
shows "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
proof-
obtain y where xs: "xs = xs'" and
yEqFresh: "(y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X')" and
alpha: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
using assms unfolding alphaAbs_ex_equal_or_qAFresh_def by blast
hence goodX': "qGood X'"
using GOOD alpha_qSwap_preserves_qGood by fastforce
then obtain z where zNot: "z \<notin> {x,x',y}" and
zFresh: "qAFresh xs z X \<and> qAFresh xs z X'"
using GOOD obtain_qFresh[of "{x,x',y}" "{X,X'}"] by auto
have "(X #[[z \<and> x]]_xs) = ((X #[[y \<and> x]]_xs) #[[z \<and> y]]_xs)"
proof(cases "y = x", simp)
assume "y \<noteq> x" hence "qAFresh xs y X" using yEqFresh by auto
thus ?thesis using zFresh qAFresh_qSwap_compose by fastforce
qed
moreover have "(X' #[[z \<and> x']]_xs) = ((X' #[[y \<and> x']]_xs) #[[z \<and> y]]_xs)"
proof(cases "y = x'", simp add: qSwap_ident)
assume "y \<noteq> x'" hence "qAFresh xs y X'" using yEqFresh by auto
thus ?thesis using zFresh qAFresh_qSwap_compose by fastforce
qed
moreover
{have "qGood (X #[[y \<and> x]]_xs)" using GOOD qSwap_preserves_qGood by auto
hence "((X #[[y \<and> x]]_xs) #[[z \<and> y]]_xs) #= ((X' #[[y \<and> x']]_xs) #[[z \<and> y]]_xs)"
using alpha by (auto simp add: qSwap_preserves_alpha)
}
ultimately have "(X #[[z \<and> x]]_xs) #= (X' #[[z \<and> x']]_xs)" by simp
thus ?thesis unfolding alphaAbs_ex_distinct_qAFresh_def using xs zNot zFresh by auto
qed
lemma alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh:
"(qAbs xs x X $= qAbs xs' x' X') = alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
unfolding alphaAbs_ex_distinct_qAFresh_def by auto
corollary alphaAbs_qAbs_iff_ex_distinct_qAFresh:
"(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
unfolding alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh
alphaAbs_ex_distinct_qAFresh_def by fastforce
lemma alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh:
assumes "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?Right = "alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right" unfolding alphaAbs_ex_equal_or_qAFresh_def by auto
moreover have "?Right \<Longrightarrow> ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh[of _ _ X]
alphaAbs_weakestEx_imp_strongestEx by auto
ultimately show ?thesis by auto
qed
corollary alphaAbs_qAbs_iff_ex_equal_or_qAFresh:
assumes "qGood X"
shows
"(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. (y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X') \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh by fastforce
thus ?thesis unfolding alphaAbs_ex_equal_or_qAFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_ex_qAFresh:
assumes "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') = alphaAbs_ex_qAFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?Middle = "alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_ex_qAFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right" unfolding alphaAbs_ex_qAFresh_def by auto
moreover have "?Right \<Longrightarrow> ?Middle"
unfolding alphaAbs_ex_qAFresh_def alphaAbs_ex_equal_or_qAFresh_def by auto
moreover have "?Middle = ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh[of X] by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_ex_qAFresh:
assumes "qGood X"
shows
"(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. qAFresh xs y X \<and> qAFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') = alphaAbs_ex_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_qAFresh by fastforce
thus ?thesis unfolding alphaAbs_ex_qAFresh_def .
qed
lemma alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qAFresh:
assumes "qGood X" and "qAbs xs x X $= qAbs xs' x' X'"
shows "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh
alphaAbs_weakestEx_imp_strongestAll by fastforce
corollary alphaAbs_qAbs_imp_all_equal_or_qAFresh:
assumes "qGood X" and "(qAbs xs x X $= qAbs xs' x' X')"
shows
"(xs = xs' \<and>
(\<forall> y. (y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_equal_or_qAFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?MiddleEx = "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
let ?MiddleAll = "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qAFresh by blast
moreover have "?Right \<Longrightarrow> ?MiddleAll"
unfolding alphaAbs_all_equal_or_qAFresh_def alphaAbs_all_distinct_qAFresh_def by auto
moreover have "?MiddleAll \<Longrightarrow> ?MiddleEx"
using assms alphaAbs_weakestAll_imp_strongestEx by fastforce
moreover have "?MiddleEx \<Longrightarrow> ?Left"
using alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh[of _ _ X] by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_all_equal_or_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. (y = x \<or> qAFresh xs y X) \<and> (y = x' \<or> qAFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_equal_or_qAFresh_def .
qed
lemma alphaAbs_qAbs_imp_alphaAbs_all_qAFresh:
assumes "qGood X" and "qAbs xs x X $= qAbs xs' x' X'"
shows "alphaAbs_all_qAFresh xs x X xs' x' X'"
proof-
have "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_qAFresh_def alphaAbs_all_equal_or_qAFresh_def by auto
qed
corollary alphaAbs_qAbs_imp_all_qAFresh:
assumes "qGood X" and "(qAbs xs x X $= qAbs xs' x' X')"
shows
"(xs = xs' \<and>
(\<forall> y. qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "alphaAbs_all_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_qAFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_all_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') = alphaAbs_all_qAFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?MiddleEx = "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
let ?MiddleAll = "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_qAFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right" using assms alphaAbs_qAbs_imp_alphaAbs_all_qAFresh by blast
moreover have "?Right \<Longrightarrow> ?MiddleAll"
unfolding alphaAbs_all_qAFresh_def alphaAbs_all_distinct_qAFresh_def by auto
moreover have "?MiddleAll \<Longrightarrow> ?MiddleEx"
using assms alphaAbs_weakestAll_imp_strongestEx by fastforce
moreover have "?MiddleEx \<Longrightarrow> ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh[of _ _ X] by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_all_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_all_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_qAFresh_def .
qed
lemma alphaAbs_qAbs_imp_alphaAbs_all_distinct_qAFresh:
assumes "qGood X" and "qAbs xs x X $= qAbs xs' x' X'"
shows "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
proof-
have "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qAFresh by blast
thus ?thesis
unfolding alphaAbs_all_distinct_qAFresh_def alphaAbs_all_equal_or_qAFresh_def by auto
qed
corollary alphaAbs_qAbs_imp_all_distinct_qAFresh:
assumes "qGood X" and "(qAbs xs x X $= qAbs xs' x' X')"
shows
"(xs = xs' \<and>
(\<forall> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_distinct_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_distinct_qAFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_all_distinct_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?MiddleEx = "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
let ?MiddleAll = "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right"
using assms alphaAbs_qAbs_imp_alphaAbs_all_distinct_qAFresh by blast
moreover have "?Right \<Longrightarrow> ?MiddleAll"
unfolding alphaAbs_all_distinct_qAFresh_def alphaAbs_all_distinct_qAFresh_def by auto
moreover have "?MiddleAll \<Longrightarrow> ?MiddleEx"
using assms alphaAbs_weakestAll_imp_strongestEx by fastforce
moreover have "?MiddleEx \<Longrightarrow> ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qAFresh[of _ _ X] by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_all_distinct_qAFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. y \<notin> {x,x'} \<and> qAFresh xs y X \<and> qAFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_all_distinct_qAFresh by blast
thus ?thesis unfolding alphaAbs_all_distinct_qAFresh_def .
qed
-subsubsection{* Then for ``qFresh" *}
+subsubsection\<open>Then for ``qFresh"\<close>
definition alphaAbs_ex_equal_or_qFresh
where
"alphaAbs_ex_equal_or_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. (y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X') \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_ex_qFresh
where
"alphaAbs_ex_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. qFresh xs y X \<and> qFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_ex_distinct_qFresh
where
"alphaAbs_ex_distinct_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> qFresh xs y X \<and> qFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_equal_or_qFresh
where
"alphaAbs_all_equal_or_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. (y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_qFresh
where
"alphaAbs_all_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. qFresh xs y X \<and> qFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
definition alphaAbs_all_distinct_qFresh
where
"alphaAbs_all_distinct_qFresh xs x X xs' x' X' ==
(xs = xs' \<and>
(\<forall> y. y \<notin> {x,x'} \<and> qFresh xs y X \<and> qFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
lemma alphaAbs_ex_equal_or_qAFresh_imp_qFresh:
"alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
unfolding alphaAbs_ex_equal_or_qAFresh_def alphaAbs_ex_equal_or_qFresh_def
using qAFresh_imp_qFresh[of _ _ X] qAFresh_imp_qFresh[of _ _ X'] by blast
lemma alphaAbs_ex_distinct_qAFresh_imp_qFresh:
"alphaAbs_ex_distinct_qAFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_ex_distinct_qFresh xs x X xs' x' X'"
unfolding alphaAbs_ex_distinct_qAFresh_def alphaAbs_ex_distinct_qFresh_def
using qAFresh_imp_qFresh[of _ _ X] qAFresh_imp_qFresh[of _ _ X'] by blast
lemma alphaAbs_ex_qAFresh_imp_qFresh:
"alphaAbs_ex_qAFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_ex_qFresh xs x X xs' x' X'"
unfolding alphaAbs_ex_qAFresh_def alphaAbs_ex_qFresh_def
using qAFresh_imp_qFresh[of _ _ X] qAFresh_imp_qFresh[of _ _ X'] by blast
lemma alphaAbs_all_equal_or_qFresh_imp_qAFresh:
"alphaAbs_all_equal_or_qFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
unfolding alphaAbs_all_equal_or_qAFresh_def alphaAbs_all_equal_or_qFresh_def
using qAFresh_imp_qFresh[of _ _ X] qAFresh_imp_qFresh[of _ _ X'] by blast
lemma alphaAbs_all_distinct_qFresh_imp_qAFresh:
"alphaAbs_all_distinct_qFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_all_distinct_qAFresh xs x X xs' x' X'"
using qAFresh_imp_qFresh
unfolding alphaAbs_all_distinct_qAFresh_def alphaAbs_all_distinct_qFresh_def by fastforce
lemma alphaAbs_all_qFresh_imp_qAFresh:
"alphaAbs_all_qFresh xs x X xs' x' X' \<Longrightarrow>
alphaAbs_all_qAFresh xs x X xs' x' X'"
using qAFresh_imp_qFresh
unfolding alphaAbs_all_qAFresh_def alphaAbs_all_qFresh_def by fastforce
lemma alphaAbs_ex_equal_or_qFresh_imp_alphaAbs_qAbs:
assumes GOOD: "qGood X" and "alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
shows "qAbs xs x X $= qAbs xs' x' X'"
proof-
obtain y where xs: "xs = xs'" and
yEqFresh: "(y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X')" and
alphaXX'yx: "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
using assms unfolding alphaAbs_ex_equal_or_qFresh_def by blast
have "\<exists> Y. X #= Y \<and> (y = x \<or> qAFresh xs y Y)"
proof(cases "y = x")
assume Case1: "y = x" hence "X #= X" using GOOD alpha_refl by auto
thus ?thesis using Case1 by fastforce
next
assume Case2: "y \<noteq> x" hence "qFresh xs y X" using yEqFresh by blast
then obtain Y where "X #= Y" and "qAFresh xs y Y"
using GOOD qFresh_imp_ex_qAFresh1 by fastforce
thus ?thesis by auto
qed
then obtain Y where alphaXY: "X #= Y" and yEqAFresh: "y = x \<or> qAFresh xs y Y" by blast
hence "(X #[[y \<and> x]]_xs) #= (Y #[[y \<and> x]]_xs)"
using GOOD qSwap_preserves_alpha by fastforce
hence alphaYXyx: "(Y #[[y \<and> x]]_xs) #= (X #[[y \<and> x]]_xs)" using alpha_sym by auto
have goodY: "qGood Y" using alphaXY GOOD alpha_preserves_qGood by auto
hence goodYyx: "qGood(Y #[[y \<and> x]]_xs)" using qSwap_preserves_qGood by auto
(* *)
have good': "qGood X'"
using GOOD alphaXX'yx alpha_qSwap_preserves_qGood by fastforce
have "\<exists> Y'. X' #= Y' \<and> (y = x' \<or> qAFresh xs y Y')"
proof(cases "y = x'")
assume Case1: "y = x'" hence "X' #= X'" using good' alpha_refl by auto
thus ?thesis using Case1 by fastforce
next
assume Case2: "y \<noteq> x'" hence "qFresh xs y X'" using yEqFresh by blast
then obtain Y' where "X' #= Y'" and "qAFresh xs y Y'"
using good' qFresh_imp_ex_qAFresh1 by fastforce
thus ?thesis by auto
qed
then obtain Y' where alphaX'Y': "X' #= Y'" and
yEqAFresh': "y = x' \<or> qAFresh xs y Y'" by blast
hence "(X' #[[y \<and> x']]_xs) #= (Y' #[[y \<and> x']]_xs)"
using good' by (auto simp add: qSwap_preserves_alpha)
hence "(Y #[[y \<and> x]]_xs) #= (Y' #[[y \<and> x']]_xs)"
using goodYyx alphaYXyx alphaXX'yx alpha_trans by blast
hence "alphaAbs_ex_equal_or_qAFresh xs x Y xs x' Y'"
unfolding alphaAbs_ex_equal_or_qAFresh_def using yEqAFresh yEqAFresh' by fastforce
hence "qAbs xs x Y $= qAbs xs x' Y'"
using goodY alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh[of Y xs x xs] by fastforce
moreover have "qAbs xs x X $= qAbs xs x Y"
using alphaXY GOOD qAbs_preserves_alpha by fastforce
moreover
{have 1: "Y' #= X'" using alphaX'Y' alpha_sym by auto
hence "qGood Y'" using good' alpha_preserves_qGood by auto
hence "qAbs xs x' Y' $= qAbs xs x' X'"
using 1 GOOD qAbs_preserves_alpha by fastforce
}
moreover have "qGoodAbs(qAbs xs x X)" using GOOD by simp
ultimately have "qAbs xs x X $= qAbs xs x' X'"
using alphaAbs_trans_twice by blast
thus ?thesis using xs by simp
qed
lemma alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?Middle = "alphaAbs_ex_equal_or_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
have "?Right \<Longrightarrow> ?Left"
using assms alphaAbs_ex_equal_or_qFresh_imp_alphaAbs_qAbs by blast
moreover have "?Left \<Longrightarrow> ?Middle"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qAFresh by blast
moreover have "?Middle \<Longrightarrow> ?Right" using
alphaAbs_ex_equal_or_qAFresh_imp_qFresh by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_ex_equal_or_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. (y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X') \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_equal_or_qFresh by blast
thus ?thesis unfolding alphaAbs_ex_equal_or_qFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_ex_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_qFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?Middle1 = "alphaAbs_ex_qAFresh xs x X xs' x' X'"
let ?Middle2 = "alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_ex_qFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Middle1" unfolding alphaAbs_ex_qAFresh_def by auto
moreover have "?Middle1 \<Longrightarrow> ?Right" using alphaAbs_ex_qAFresh_imp_qFresh by fastforce
moreover have "?Right \<Longrightarrow> ?Middle2"
unfolding alphaAbs_ex_qFresh_def alphaAbs_ex_equal_or_qFresh_def by auto
moreover have "?Middle2 \<Longrightarrow> ?Left"
using assms alphaAbs_ex_equal_or_qFresh_imp_alphaAbs_qAbs by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_ex_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. qFresh xs y X \<and> qFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_qFresh by blast
thus ?thesis unfolding alphaAbs_ex_qFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_distinct_qFresh xs x X xs' x' X'"
proof-
let ?Left = "qAbs xs x X $= qAbs xs' x' X'"
let ?Middle1 = "alphaAbs_ex_distinct_qAFresh xs x X xs' x' X'"
let ?Middle2 = "alphaAbs_ex_equal_or_qFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_ex_distinct_qFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Middle1" unfolding alphaAbs_ex_distinct_qAFresh_def by auto
moreover have "?Middle1 \<Longrightarrow> ?Right" using alphaAbs_ex_distinct_qAFresh_imp_qFresh by fastforce
moreover have "?Right \<Longrightarrow> ?Middle2"
unfolding alphaAbs_ex_distinct_qFresh_def alphaAbs_ex_equal_or_qFresh_def by auto
moreover have "?Middle2 \<Longrightarrow> ?Left"
using assms alphaAbs_ex_equal_or_qFresh_imp_alphaAbs_qAbs by fastforce
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_ex_distinct_qFresh:
assumes GOOD: "qGood X"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. y \<notin> {x, x'} \<and> qFresh xs y X \<and> qFresh xs y X' \<and>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_ex_distinct_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_ex_distinct_qFresh by blast
thus ?thesis unfolding alphaAbs_ex_distinct_qFresh_def .
qed
lemma alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qFresh:
assumes "qGood X" and "qAbs xs x X $= qAbs xs' x' X'"
shows "alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
proof-
have "qGoodAbs(qAbs xs x X)" using assms by auto
hence "qGoodAbs(qAbs xs' x' X')" using assms alphaAbs_preserves_qGoodAbs by blast
hence GOOD: "qGood X \<and> qGood X'" using assms by auto
have xs: "xs = xs'" using assms by auto
show ?thesis
unfolding alphaAbs_all_equal_or_qFresh_def using xs
proof(intro conjI impI allI, simp)
fix y
assume yEqFresh: "(y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X')"
have "\<exists> Y. X #= Y \<and> (y = x \<or> qAFresh xs y Y)"
proof(cases "y = x")
assume Case1: "y = x" hence "X #= X" using GOOD alpha_refl by auto
thus ?thesis using Case1 by fastforce
next
assume Case2: "y \<noteq> x" hence "qFresh xs y X" using yEqFresh by blast
then obtain Y where "X #= Y" and "qAFresh xs y Y"
using GOOD qFresh_imp_ex_qAFresh1 by blast
thus ?thesis by auto
qed
then obtain Y where alphaXY: "X #= Y" and yEqAFresh: "y = x \<or> qAFresh xs y Y" by blast
hence alphaXYyx: "(X #[[y \<and> x]]_xs) #= (Y #[[y \<and> x]]_xs)"
using GOOD by (auto simp add: qSwap_preserves_alpha)
have goodY: "qGood Y" using GOOD alphaXY alpha_preserves_qGood by auto
(* *)
have "\<exists> Y'. X' #= Y' \<and> (y = x' \<or> qAFresh xs y Y')"
proof(cases "y = x'")
assume Case1: "y = x'" hence "X' #= X'" using GOOD alpha_refl by auto
thus ?thesis using Case1 by fastforce
next
assume Case2: "y \<noteq> x'" hence "qFresh xs y X'" using yEqFresh by blast
then obtain Y' where "X' #= Y'" and "qAFresh xs y Y'"
using GOOD qFresh_imp_ex_qAFresh1 by blast
thus ?thesis by auto
qed
then obtain Y' where alphaX'Y': "X' #= Y'" and
yEqAFresh': "y = x' \<or> qAFresh xs y Y'" by blast
hence "(X' #[[y \<and> x']]_xs) #= (Y' #[[y \<and> x']]_xs)"
using GOOD by (auto simp add: qSwap_preserves_alpha)
hence alphaY'X'yx': "(Y' #[[y \<and> x']]_xs) #= (X' #[[y \<and> x']]_xs)" using alpha_sym by auto
have goodY': "qGood Y'" using GOOD alphaX'Y' alpha_preserves_qGood by auto
(* *)
have 1: "Y #= X" using alphaXY alpha_sym by auto
hence "qGood Y" using GOOD alpha_preserves_qGood by auto
hence 2: "qAbs xs x Y $= qAbs xs x X"
using 1 GOOD qAbs_preserves_alpha by blast
moreover have "qAbs xs x' X' $= qAbs xs x' Y'"
using alphaX'Y' GOOD qAbs_preserves_alpha by blast
moreover
{have "qGoodAbs(qAbs xs x X)" using GOOD by simp
hence "qGoodAbs(qAbs xs x Y)" using 2 alphaAbs_preserves_qGoodAbs by fastforce
}
ultimately have "qAbs xs x Y $= qAbs xs x' Y'"
using assms xs alphaAbs_trans_twice by blast
hence "alphaAbs_all_equal_or_qAFresh xs x Y xs x' Y'"
using goodY goodY' alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qAFresh by blast
hence "(Y #[[y \<and> x]]_xs) #= (Y' #[[y \<and> x']]_xs)"
unfolding alphaAbs_all_equal_or_qAFresh_def
using yEqAFresh yEqAFresh' by auto
moreover have "qGood (X #[[y \<and> x]]_xs)" using GOOD qSwap_preserves_qGood by auto
ultimately show "(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)"
using alphaXYyx alphaY'X'yx' alpha_trans_twice by blast
qed
qed
corollary alphaAbs_qAbs_imp_all_equal_or_qFresh:
assumes "qGood X" and "(qAbs xs x X $= qAbs xs' x' X')"
shows
"(xs = xs' \<and>
(\<forall> y. (y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qFresh by blast
thus ?thesis unfolding alphaAbs_all_equal_or_qFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
proof-
let ?Left = "(qAbs xs x X $= qAbs xs' x' X')"
let ?Middle = "alphaAbs_all_equal_or_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qFresh by blast
moreover have "?Right \<Longrightarrow> ?Middle"
using alphaAbs_all_equal_or_qFresh_imp_qAFresh by fastforce
moreover have "?Middle ==> ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qAFresh by blast
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_all_equal_or_qFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. (y = x \<or> qFresh xs y X) \<and> (y = x' \<or> qFresh xs y X') \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_all_equal_or_qFresh by blast
thus ?thesis unfolding alphaAbs_all_equal_or_qFresh_def .
qed
lemma alphaAbs_qAbs_imp_alphaAbs_all_qFresh:
assumes "qGood X" and "qAbs xs x X $= qAbs xs' x' X'"
shows "alphaAbs_all_qFresh xs x X xs' x' X'"
proof-
let ?Left = "(qAbs xs x X $= qAbs xs' x' X')"
let ?Middle = "alphaAbs_all_equal_or_qFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_qFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Middle"
using assms alphaAbs_qAbs_imp_alphaAbs_all_equal_or_qFresh by blast
moreover have "?Middle \<Longrightarrow> ?Right"
unfolding alphaAbs_all_equal_or_qFresh_def alphaAbs_all_qFresh_def by auto
ultimately show ?thesis using assms by blast
qed
corollary alphaAbs_qAbs_imp_all_qFresh:
assumes "qGood X" and "(qAbs xs x X $= qAbs xs' x' X')"
shows
"(xs = xs' \<and>
(\<forall> y. qFresh xs y X \<and> qFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "alphaAbs_all_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_imp_alphaAbs_all_qFresh by blast
thus ?thesis unfolding alphaAbs_all_qFresh_def .
qed
lemma alphaAbs_qAbs_iff_alphaAbs_all_qFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_qFresh xs x X xs' x' X'"
proof-
let ?Left = "(qAbs xs x X $= qAbs xs' x' X')"
let ?Middle = "alphaAbs_all_qAFresh xs x X xs' x' X'"
let ?Right = "alphaAbs_all_qFresh xs x X xs' x' X'"
have "?Left \<Longrightarrow> ?Right"
using assms alphaAbs_qAbs_imp_alphaAbs_all_qFresh by blast
moreover have "?Right \<Longrightarrow> ?Middle"
using alphaAbs_all_qFresh_imp_qAFresh by fastforce
moreover have "?Middle \<Longrightarrow> ?Left"
using assms alphaAbs_qAbs_iff_alphaAbs_all_qAFresh by blast
ultimately show ?thesis by blast
qed
corollary alphaAbs_qAbs_iff_all_qFresh:
assumes "qGood X" and "qGood X'"
shows "(qAbs xs x X $= qAbs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. qFresh xs y X \<and> qFresh xs y X' \<longrightarrow>
(X #[[y \<and> x]]_xs) #= (X' #[[y \<and> x']]_xs)))"
proof-
have "(qAbs xs x X $= qAbs xs' x' X') =
alphaAbs_all_qFresh xs x X xs' x' X'"
using assms alphaAbs_qAbs_iff_alphaAbs_all_qFresh by blast
thus ?thesis unfolding alphaAbs_all_qFresh_def .
qed
end (* context FixVars *)
end
diff --git a/thys/Binding_Syntax_Theory/QuasiTerms_Swap_Fresh.thy b/thys/Binding_Syntax_Theory/QuasiTerms_Swap_Fresh.thy
--- a/thys/Binding_Syntax_Theory/QuasiTerms_Swap_Fresh.thy
+++ b/thys/Binding_Syntax_Theory/QuasiTerms_Swap_Fresh.thy
@@ -1,685 +1,685 @@
-section {* Quasi-Terms with Swapping and Freshness *}
+section \<open>Quasi-Terms with Swapping and Freshness\<close>
theory QuasiTerms_Swap_Fresh imports Preliminaries
begin
-text{*
+text\<open>
This section defines and studies the (totally free) datatype of quasi-terms
and the notions of freshness and
swapping variables for them.
``Quasi" refers to the fact that these items are not (yet) factored to alpha-equivalence.
- We shall later call ``terms" those alpha-equivalence classes. *}
+ We shall later call ``terms" those alpha-equivalence classes.\<close>
-subsection {* The datatype of quasi-terms with bindings *}
+subsection \<open>The datatype of quasi-terms with bindings\<close>
datatype
('index,'bindex,'varSort,'var,'opSym)qTerm =
qVar 'varSort 'var
|qOp 'opSym "('index, (('index,'bindex,'varSort,'var,'opSym)qTerm))input"
"('bindex, (('index,'bindex,'varSort,'var,'opSym)qAbs)) input"
and
('index,'bindex,'varSort,'var,'opSym)qAbs =
qAbs 'varSort 'var "('index,'bindex,'varSort,'var,'opSym)qTerm"
-text{* Above:
+text\<open>Above:
\begin{itemize}
\item ``Var" stands for ``variable injection"
\item ``Op" stands for ``operation"
\item ``opSym" stands for ``operation symbol"
\item ``q" stands for ``quasi"
\item ``Abs" stands for ``abstraction"
\end{itemize}
Thus, a quasi-term is either (an injection of) a variable, or an operation symbol applied
to a term-input and an abstraction-input
(where, for any type $T$, $T$-inputs are partial
maps from indexes to $T$. A quasi-abstraction is
essentially a pair (variable,quasi-term).
- *}
+\<close>
type_synonym ('index,'bindex,'varSort,'var,'opSym)qTermItem =
"('index,'bindex,'varSort,'var,'opSym)qTerm +
('index,'bindex,'varSort,'var,'opSym)qAbs"
abbreviation termIn ::
"('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTermItem"
where "termIn X == Inl X"
abbreviation absIn ::
"('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTermItem"
where "absIn A == Inr A"
-subsection {* Induction principles *}
+subsection \<open>Induction principles\<close>
definition qTermLess :: "('index,'bindex,'varSort,'var,'opSym)qTermItem rel"
where
"qTermLess == {(termIn X, termIn(qOp delta inp binp))| X delta inp binp i. inp i = Some X} \<union>
{(absIn A, termIn(qOp delta inp binp))| A delta inp binp i. binp i = Some A} \<union>
{(termIn X, absIn (qAbs xs x X))| X xs x. True}"
-text{* This induction will be used only temporarily, until we
- get a better one, involving swapping: *}
+text\<open>This induction will be used only temporarily, until we
+ get a better one, involving swapping:\<close>
lemma qTerm_rawInduct[case_names Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm" and
A :: "('index,'bindex,'varSort,'var,'opSym)qAbs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>liftAll phi inp; liftAll phiAbs binp\<rbrakk> \<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. phi X \<Longrightarrow> phiAbs (qAbs xs x X)"
shows "phi X \<and> phiAbs A"
by (induct rule: qTerm_qAbs.induct)
(fastforce intro!: Var Op Abs rangeI simp: liftAll_def)+
lemma qTermLess_wf: "wf qTermLess"
unfolding wf_def proof safe
fix chi item
assume *: "\<forall>item. (\<forall>item'. (item', item) \<in> qTermLess \<longrightarrow> chi item') \<longrightarrow> chi item"
show "chi item"
proof-
{fix X A
have "chi (termIn X) \<and> chi (absIn A)"
apply(induct rule: qTerm_rawInduct)
using * unfolding qTermLess_def liftAll_def by blast+
}
thus ?thesis by(cases item) auto
qed
qed
lemma qTermLessPlus_wf: "wf (qTermLess ^+)"
using qTermLess_wf wf_trancl by auto
-text{* The skeleton of a quasi-term item -- this is the generalization
+text\<open>The skeleton of a quasi-term item -- this is the generalization
of the size function from the case of finitary syntax.
- We use the skeleton later for proving correct various recursive function definitions, notably that of ``alpha". *}
+ We use the skeleton later for proving correct various recursive function definitions, notably that of ``alpha".\<close>
function
qSkel :: "('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> ('index,'bindex)tree"
and
qSkelAbs :: "('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> ('index,'bindex)tree"
where
"qSkel (qVar xs x) = Branch (\<lambda>i. None) (\<lambda>i. None)"
|
"qSkel (qOp delta inp binp) = Branch (lift qSkel inp) (lift qSkelAbs binp)"
|
"qSkelAbs (qAbs xs x X) = Branch (\<lambda>i. Some(qSkel X)) (\<lambda>i. None)"
by(pat_completeness, auto)
termination by(relation qTermLess, simp add: qTermLess_wf, auto simp add: qTermLess_def)
-text{* Next is a template for generating induction principles whenever we come up
- with relation on terms included in the kernel of the skeleton operator. *}
+text\<open>Next is a template for generating induction principles whenever we come up
+ with relation on terms included in the kernel of the skeleton operator.\<close>
lemma qTerm_templateInduct[case_names Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A :: "('index,'bindex,'varSort,'var,'opSym)qAbs"
and phi phiAbs and rel
assumes
REL: "\<And> X Y. (X,Y) \<in> rel \<Longrightarrow> qSkel Y = qSkel X" and
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. (\<And> Y. (X,Y) \<in> rel \<Longrightarrow> phi Y) \<Longrightarrow> phiAbs (qAbs xs x X)"
shows "phi X \<and> phiAbs A"
proof-
{fix T
have "\<forall> X A. (T = qSkel X \<longrightarrow> phi X) \<and> (T = qSkelAbs A \<longrightarrow> phiAbs A)"
proof(induct rule: treeLess_induct)
case (1 T')
show ?case apply safe
subgoal for X _
using assms 1 unfolding treeLess_def liftAll_def
by (cases X) (auto simp add: lift_def, metis option.simps(5))
subgoal for _ A apply (cases A)
using assms 1 unfolding treeLess_def by simp .
qed
}
thus ?thesis by blast
qed
-text{* A modification of the canonical immediate-subterm
-relation on quasi-terms, that takes into account a relation assumed included in the skeleton kernel. *}
+text\<open>A modification of the canonical immediate-subterm
+relation on quasi-terms, that takes into account a relation assumed included in the skeleton kernel.\<close>
definition qTermLess_modulo ::
"('index,'bindex,'varSort,'var,'opSym)qTerm rel \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTermItem rel"
where
"qTermLess_modulo rel ==
{(termIn X, termIn(qOp delta inp binp))| X delta inp binp i. inp i = Some X} \<union>
{(absIn A, termIn(qOp delta inp binp))| A delta inp binp j. binp j = Some A} \<union>
{(termIn Y, absIn (qAbs xs x X))| X Y xs x. (X,Y) \<in> rel}"
lemma qTermLess_modulo_wf:
fixes rel::"('index,'bindex,'varSort,'var,'opSym)qTerm rel"
assumes "\<And> X Y. (X,Y) \<in> rel \<Longrightarrow> qSkel Y = qSkel X"
shows "wf (qTermLess_modulo rel)"
proof(unfold wf_def, auto)
fix chi item
assume *:
"\<forall>item. (\<forall>item'. (item', item) \<in> qTermLess_modulo rel \<longrightarrow> chi item')
\<longrightarrow> chi item"
show "chi item"
proof-
obtain phi where phi_def: "phi = (\<lambda> X. chi (termIn X))" by blast
obtain phiAbs where phiAbs_def: "phiAbs = (\<lambda> A. chi (absIn A))" by blast
{fix X A
have "chi (termIn X) \<and> chi (absIn A)"
apply(induct rule: qTerm_templateInduct[of rel])
using * assms unfolding qTermLess_modulo_def liftAll_def by blast+
}
thus ?thesis unfolding phi_def phiAbs_def
by(cases item, auto)
qed
qed
-subsection {* Swap and substitution on variables *}
+subsection \<open>Swap and substitution on variables\<close>
definition sw :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow> 'var \<Rightarrow> 'var"
where
"sw ys y1 y2 xs x ==
if ys = xs then if x = y1 then y2
else if x = y2 then y1
else x
else x"
abbreviation sw_abbrev :: "'var \<Rightarrow> 'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow> 'var"
("_ @_[_ \<and> _]'__" 200)
where "(x @xs[y1 \<and> y2]_ys) == sw ys y1 y2 xs x"
definition sb :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow> 'var \<Rightarrow> 'var"
where
"sb ys y1 y2 xs x ==
if ys = xs then if x = y2 then y1
else x
else x"
abbreviation sb_abbrev :: "'var \<Rightarrow> 'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow> 'var"
("_ @_[_ '/ _]'__" 200)
where "(x @xs[y1 / y2]_ys) == sb ys y1 y2 xs x"
theorem sw_simps1[simp]: "(x @xs[x \<and> y]_xs) = y"
unfolding sw_def by simp
theorem sw_simps2[simp]: "(x @xs[y \<and> x]_xs) = y"
unfolding sw_def by simp
theorem sw_simps3[simp]:
"(zs \<noteq> xs \<or> x \<notin> {z1,z2}) \<Longrightarrow> (x @xs[z1 \<and> z2]_zs) = x"
unfolding sw_def by simp
lemmas sw_simps = sw_simps1 sw_simps2 sw_simps3
theorem sw_ident[simp]: "(x @xs[y \<and> y]_ys) = x"
unfolding sw_def by auto
theorem sw_compose:
"((z @zs[x \<and> y]_xs) @zs[x' \<and> y']_xs') =
((z @zs[x' \<and> y']_xs') @zs[(x @xs[x' \<and> y']_xs') \<and> (y @xs[x' \<and> y']_xs')]_xs)"
by(unfold sw_def, auto)
theorem sw_commute:
assumes "zs \<noteq> zs' \<or> {x,y} Int {x',y'} = {}"
shows "((u @us[x \<and> y]_zs) @us[x' \<and> y']_zs') = ((u @us[x' \<and> y']_zs') @us[x \<and> y]_zs)"
using assms by(unfold sw_def, auto)
theorem sw_involutive[simp]:
"((z @zs[x \<and> y]_xs) @zs[x \<and> y]_xs) = z"
by(unfold sw_def, auto)
theorem sw_inj[simp]:
"((z @zs[x \<and> y]_xs) = (z' @zs[x \<and> y]_xs)) = (z = z')"
by (simp add: sw_def)
lemma sw_preserves_mship[simp]:
assumes "{y1,y2} \<subseteq> Var ys"
shows "((x @xs[y1 \<and> y2]_ys) \<in> Var xs) = (x \<in> Var xs)"
using assms unfolding sw_def by auto
theorem sw_sym:
"(z @zs[x \<and> y]_xs) = (z @zs[y \<and> x]_xs)"
by (unfold sw_def) auto
theorem sw_involutive2[simp]:
"((z @zs[x \<and> y]_xs) @zs[y \<and> x]_xs) = z"
by (unfold sw_def) auto
theorem sw_trans:
"us \<noteq> zs \<or> u \<notin> {y,z} \<Longrightarrow>
((u @us[y \<and> x]_zs) @us[z \<and> y]_zs) = (u @us[z \<and> x]_zs)"
by (unfold sw_def) auto
lemmas sw_otherSimps =
sw_ident sw_involutive sw_inj sw_preserves_mship sw_involutive2
theorem sb_simps1[simp]: "(x @xs[y / x]_xs) = y"
unfolding sb_def by simp
theorem sb_simps2[simp]:
"(zs \<noteq> xs \<or> z2 \<noteq> x) \<Longrightarrow> (x @xs[z1 / z2]_zs) = x"
unfolding sb_def by auto
lemmas sb_simps = sb_simps1 sb_simps2
theorem sb_ident[simp]: "(x @xs[y / y]_ys) = x"
unfolding sb_def by auto
theorem sb_compose1:
"((z @zs[y1 / x]_xs) @zs[y2 / x]_xs) = (z @zs[(y1 @xs[y2 / x]_xs) / x]_xs)"
by(unfold sb_def, auto)
theorem sb_compose2:
"ys \<noteq> xs \<or> (x2 \<notin> {y1,y2}) \<Longrightarrow>
((z @zs[x1 / x2]_xs) @zs[y1 / y2]_ys) =
((z @zs[y1 / y2]_ys) @zs[(x1 @xs[y1 / y2]_ys) / x2]_xs)"
by (unfold sb_def) auto
theorem sb_commute:
assumes "zs \<noteq> zs' \<or> {x,y} Int {x',y'} = {}"
shows "((u @us[x / y]_zs) @us[x' / y']_zs') = ((u @us[x' / y']_zs') @us[x / y]_zs)"
using assms by (unfold sb_def) auto
theorem sb_idem[simp]:
"((z @zs[x / y]_xs) @zs[x / y]_xs) = (z @zs[x / y]_xs)"
by (unfold sb_def) auto
lemma sb_preserves_mship[simp]:
assumes "{y1,y2} \<subseteq> Var ys"
shows "((x @xs[y1 / y2]_ys) \<in> Var xs) = (x \<in> Var xs)"
using assms by (unfold sb_def) auto
theorem sb_trans:
"us \<noteq> zs \<or> u \<noteq> y \<Longrightarrow>
((u @us[y / x]_zs) @us[z / y]_zs) = (u @us[z / x]_zs)"
by (unfold sb_def) auto
lemmas sb_otherSimps =
sb_ident sb_idem sb_preserves_mship
-subsection {* The swapping and freshness operators *}
+subsection \<open>The swapping and freshness operators\<close>
-text {* For establishing the preliminary results quickly, we use both the notion of
+text \<open>For establishing the preliminary results quickly, we use both the notion of
binding-sensitive freshness (operator ``qFresh")
and that of ``absolute" freshness, ignoring bindings (operator ``qAFresh"). Later,
- for alpha-equivalence classes, ``qAFresh" will not make sense. *}
+ for alpha-equivalence classes, ``qAFresh" will not make sense.\<close>
definition
aux_qSwap_ignoreFirst3 ::
"'varSort * 'var * 'var * ('index,'bindex,'varSort,'var,'opSym)qTerm +
'varSort * 'var * 'var * ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTermItem"
where
"aux_qSwap_ignoreFirst3 K =
(case K of Inl(zs,x,y,X) \<Rightarrow> termIn X
|Inr(zs,x,y,A) \<Rightarrow> absIn A)"
lemma qTermLess_ingoreFirst3_wf:
"wf(inv_image qTermLess aux_qSwap_ignoreFirst3)"
using qTermLess_wf wf_inv_image by auto
function
qSwap :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTerm"
and
qSwapAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qAbs"
where
"qSwap zs x y (qVar zs' z) = qVar zs' (z @zs'[x \<and> y]_zs)"
|
"qSwap zs x y (qOp delta inp binp) =
qOp delta (lift (qSwap zs x y) inp) (lift (qSwapAbs zs x y) binp)"
|
"qSwapAbs zs x y (qAbs zs' z X) = qAbs zs' (z @zs'[x \<and> y]_zs) (qSwap zs x y X)"
by(pat_completeness, auto)
termination
by(relation "inv_image qTermLess aux_qSwap_ignoreFirst3",
simp add: qTermLess_ingoreFirst3_wf,
auto simp add: qTermLess_def aux_qSwap_ignoreFirst3_def)
lemmas qSwapAll_simps = qSwap.simps qSwapAbs.simps
abbreviation qSwap_abbrev ::
"('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTerm" ("_ #[[_ \<and> _]]'__" 200)
where "(X #[[z1 \<and> z2]]_zs) == qSwap zs z1 z2 X"
abbreviation qSwapAbs_abbrev ::
"('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> 'varSort \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qAbs" ("_ $[[_ \<and> _]]'__" 200)
where "(A $[[z1 \<and> z2]]_zs) == qSwapAbs zs z1 z2 A"
definition
aux_qFresh_ignoreFirst2 ::
"'varSort * 'var * ('index,'bindex,'varSort,'var,'opSym)qTerm +
'varSort * 'var * ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)qTermItem"
where
"aux_qFresh_ignoreFirst2 K =
(case K of Inl(zs,x,X) \<Rightarrow> termIn X
|Inr (zs,x,A) \<Rightarrow> absIn A)"
lemma qTermLess_ingoreFirst2_wf: "wf(inv_image qTermLess aux_qFresh_ignoreFirst2)"
using qTermLess_wf wf_inv_image by auto
-text{* The quasi absolutely-fresh predicate:
+text\<open>The quasi absolutely-fresh predicate:
(note that this is not an oxymoron: ``quasi" refers
to being an operator on quasi-terms, and not on
terms, i.e., on alpha-equivalence classes;
``absolutely'' refers to not ignoring bindings in the notion of freshness,
-and thus counting absolutely all the variables. *}
+and thus counting absolutely all the variables.\<close>
function
qAFresh :: "'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> bool"
and
qAFreshAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> bool"
where
"qAFresh xs x (qVar ys y) = (xs \<noteq> ys \<or> x \<noteq> y)"
|
"qAFresh xs x (qOp delta inp binp) =
(liftAll (qAFresh xs x) inp \<and> liftAll (qAFreshAbs xs x) binp)"
|
"qAFreshAbs xs x (qAbs ys y X) = ((xs \<noteq> ys \<or> x \<noteq> y) \<and> qAFresh xs x X)"
by(pat_completeness, auto)
termination
by(relation "inv_image qTermLess aux_qFresh_ignoreFirst2",
simp add: qTermLess_ingoreFirst2_wf,
auto simp add: qTermLess_def aux_qFresh_ignoreFirst2_def)
lemmas qAFreshAll_simps = qAFresh.simps qAFreshAbs.simps
-text{* The next is standard freshness -- note that its definition differs from that
-of absolute freshness only at the clause for abstractions. *}
+text\<open>The next is standard freshness -- note that its definition differs from that
+of absolute freshness only at the clause for abstractions.\<close>
function
qFresh :: "'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qTerm \<Rightarrow> bool"
and
qFreshAbs :: "'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)qAbs \<Rightarrow> bool"
where
"qFresh xs x (qVar ys y) = (xs \<noteq> ys \<or> x \<noteq> y)"
|
"qFresh xs x (qOp delta inp binp) =
(liftAll (qFresh xs x) inp \<and> liftAll (qFreshAbs xs x) binp)"
|
"qFreshAbs xs x (qAbs ys y X) = ((xs = ys \<and> x = y) \<or> qFresh xs x X)"
by(pat_completeness, auto)
termination
by(relation "inv_image qTermLess aux_qFresh_ignoreFirst2",
simp add: qTermLess_ingoreFirst2_wf,
auto simp add: qTermLess_def aux_qFresh_ignoreFirst2_def)
lemmas qFreshAll_simps = qFresh.simps qFreshAbs.simps
-subsection {* Compositional properties of swapping *}
+subsection \<open>Compositional properties of swapping\<close>
lemma qSwapAll_ident:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"
shows "(X #[[x \<and> x]]_zs) = X \<and> (A $[[x \<and> x]]_zs) = A"
by (induct rule: qTerm_rawInduct)
(auto simp add: liftAll_def lift_cong lift_ident)
corollary qSwap_ident[simp]: "(X #[[x \<and> x]]_zs) = X"
by(simp add: qSwapAll_ident)
lemma qSwapAll_compose:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs x y x' y'
shows
"((X #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs') =
((X #[[x' \<and> y']]_zs') #[[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]]_zs)
\<and>
((A $[[x \<and> y]]_zs) $[[x' \<and> y']]_zs') =
((A $[[x' \<and> y']]_zs') $[[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]]_zs)"
proof(induct rule: qTerm_rawInduct[of _ _ X A])
case (Op delta inp binp)
then show ?case by (auto intro!: lift_cong simp: liftAll_def lift_comp)
qed (auto simp add: sw_def sw_compose)
corollary qSwap_compose:
"((X #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs') =
((X #[[x' \<and> y']]_zs') #[[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]]_zs)"
by (meson qSwapAll_compose)
lemma qSwap_commute:
assumes "zs \<noteq> zs' \<or> {x,y} Int {x',y'} = {}"
shows "((X #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs') = ((X #[[x' \<and> y']]_zs') #[[x \<and> y]]_zs)"
by (metis assms disjoint_insert(1) qSwapAll_compose sw_simps3)
lemma qSwapAll_involutive:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs x y
shows "((X #[[x \<and> y]]_zs) #[[x \<and> y]]_zs) = X \<and>
((A $[[x \<and> y]]_zs) $[[x \<and> y]]_zs) = A"
proof(induct rule: qTerm_rawInduct[of _ _ X A])
case (Op delta inp binp)
then show ?case
unfolding qSwapAll_simps(2) liftAll_lift_ext
lift_comp o_def
by (simp add: lift_ident)
qed(auto)
corollary qSwap_involutive[simp]:
"((X #[[x \<and> y]]_zs) #[[x \<and> y]]_zs) = X"
by(simp add: qSwapAll_involutive)
lemma qSwapAll_sym:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs x y
shows "(X #[[x \<and> y]]_zs) = (X #[[y \<and> x]]_zs) \<and>
(A $[[x \<and> y]]_zs) = (A $[[y \<and> x]]_zs)"
by (induct rule: qTerm_rawInduct[of _ _ X A])
(auto simp: sw_sym lift_comp liftAll_lift_ext)
corollary qSwap_sym:
"(X #[[x \<and> y]]_zs) = (X #[[y \<and> x]]_zs)"
by(simp add: qSwapAll_sym)
lemma qAFreshAll_qSwapAll_id:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and zs z1 z2
shows "(qAFresh zs z1 X \<and> qAFresh zs z2 X \<longrightarrow> (X #[[z1 \<and> z2]]_zs) = X) \<and>
(qAFreshAbs zs z1 A \<and> qAFreshAbs zs z2 A \<longrightarrow> (A $[[z1 \<and> z2]]_zs) = A)"
by (induct rule: qTerm_rawInduct[of _ _ X A])
(auto intro!: ext simp: liftAll_def lift_def option.case_eq_if)
corollary qAFresh_qSwap_id[simp]:
"\<lbrakk>qAFresh zs z1 X; qAFresh zs z2 X\<rbrakk> \<Longrightarrow> (X #[[z1 \<and> z2]]_zs) = X"
by(simp add: qAFreshAll_qSwapAll_id)
lemma qAFreshAll_qSwapAll_compose:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs"and zs x y z
shows "(qAFresh zs y X \<and> qAFresh zs z X \<longrightarrow>
((X #[[y \<and> x]]_zs) #[[z \<and> y]]_zs) = (X #[[z \<and> x]]_zs)) \<and>
(qAFreshAbs zs y A \<and> qAFreshAbs zs z A \<longrightarrow>
((A $[[y \<and> x]]_zs) $[[z \<and> y]]_zs) = (A $[[z \<and> x]]_zs))"
by (induct rule: qTerm_rawInduct[of _ _ X A])
(auto intro!: ext simp: sw_trans lift_comp lift_def liftAll_def option.case_eq_if)
corollary qAFresh_qSwap_compose:
"\<lbrakk>qAFresh zs y X; qAFresh zs z X\<rbrakk> \<Longrightarrow>
((X #[[y \<and> x]]_zs) #[[z \<and> y]]_zs) = (X #[[z \<and> x]]_zs)"
by(simp add: qAFreshAll_qSwapAll_compose)
-subsection {* Induction and well-foundedness modulo swapping *}
+subsection \<open>Induction and well-foundedness modulo swapping\<close>
lemma qSkel_qSwapAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and x y zs
shows "qSkel(X #[[x \<and> y]]_zs) = qSkel X \<and>
qSkelAbs(A $[[x \<and> y]]_zs) = qSkelAbs A"
proof(induct rule: qTerm_rawInduct[of _ _ X A])
case (Op delta inp binp)
then show ?case
unfolding qSwapAll_simps(2) liftAll_lift_ext qSkel.simps(2)
lift_comp comp_apply by simp
qed auto
corollary qSkel_qSwap: "qSkel(X #[[x \<and> y]]_zs) = qSkel X"
by(simp add: qSkel_qSwapAll)
-text{*
+text\<open>
For induction modulo swapping, one may wish to swap not just once,
but several times at the
induction hypothesis (an example of this will be the proof of compatibility
of ``qSwap" with alpha) -- for this, we introduce the following relation
(the suffix ``Raw" signifies the fact that the involved variables are
- not required to be well-sorted): *}
+ not required to be well-sorted):\<close>
inductive_set qSwapped :: "('index,'bindex,'varSort,'var,'opSym)qTerm rel"
where
Refl: "(X,X) \<in> qSwapped"
|
Trans: "\<lbrakk>(X,Y) \<in> qSwapped; (Y,Z) \<in> qSwapped\<rbrakk> \<Longrightarrow> (X,Z) \<in> qSwapped"
|
Swap: "(X,Y) \<in> qSwapped \<Longrightarrow> (X, Y #[[x \<and> y]]_zs) \<in> qSwapped"
lemmas qSwapped_Clauses = qSwapped.Refl qSwapped.Trans qSwapped.Swap
lemma qSwap_qSwapped: "(X, X #[[x \<and> y]]_zs): qSwapped"
by (auto simp add: qSwapped_Clauses)
lemma qSwapped_qSkel:
"(X,Y) \<in> qSwapped \<Longrightarrow> qSkel Y = qSkel X"
by(erule qSwapped.induct, auto simp add: qSkel_qSwap)
-text{* The following is henceforth our main induction principle for quasi-terms. At the
+text\<open>The following is henceforth our main induction principle for quasi-terms. At the
clause for abstractions, the user may choose among the two
induction hypotheses (IHs):
\\-(1) IH for all swapped terms
\\-(2) IH for all terms with the same skeleton.
The user may choose only one of the above, and ignore the others, but may of course also
assume both. (2) is stronger than (1),
but we offer both of them for convenience in proofs.
-Most of the times, (1) will be the most convenient. *}
+Most of the times, (1) will be the most convenient.\<close>
lemma qTerm_induct[case_names Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)qTerm"
and A :: "('index,'bindex,'varSort,'var,'opSym)qAbs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (qVar xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (qOp delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>\<And> Y. (X,Y) \<in> qSwapped \<Longrightarrow> phi Y;
\<And> Y. qSkel Y = qSkel X \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (qAbs xs x X)"
shows "phi X \<and> phiAbs A"
by (induct rule: qTerm_templateInduct[of "qSwapped \<union> {(X,Y). qSkel Y = qSkel X}"],
auto simp add: qSwapped_qSkel assms)
-text{* The following relation will be needed for proving alpha-equivalence well-defined: *}
+text\<open>The following relation will be needed for proving alpha-equivalence well-defined:\<close>
definition qTermQSwappedLess :: "('index,'bindex,'varSort,'var,'opSym)qTermItem rel"
where "qTermQSwappedLess = qTermLess_modulo qSwapped"
lemma qTermQSwappedLess_wf: "wf qTermQSwappedLess"
unfolding qTermQSwappedLess_def
using qSwapped_qSkel qTermLess_modulo_wf[of qSwapped] by blast
-subsection{* More properties connecting swapping and freshness *}
+subsection\<open>More properties connecting swapping and freshness\<close>
lemma qSwap_3commute:
assumes *: "qAFresh ys y X" and **: "qAFresh ys y0 X"
and ***: "ys \<noteq> zs \<or> y0 \<notin> {z1,z2}"
shows "((X #[[z1 \<and> z2]]_zs) #[[y0 \<and> x @ys[z1 \<and> z2]_zs]]_ys) =
(((X #[[y \<and> x]]_ys) #[[y0 \<and> y]]_ys) #[[z1 \<and> z2]]_zs)"
proof-
have "y0 = (y0 @ys[z1 \<and> z2]_zs)" using *** unfolding sw_def by auto
hence "((X #[[z1 \<and> z2]]_zs) #[[y0 \<and> x @ys[z1 \<and> z2]_zs]]_ys) =
((X #[[y0 \<and> x]]_ys) #[[z1 \<and> z2]]_zs)"
by(simp add: qSwap_compose[of _ z1])
also have "((X #[[y0 \<and> x]]_ys) #[[z1 \<and> z2]]_zs) =
(((X #[[y \<and> x]]_ys) #[[y0 \<and> y]]_ys) #[[z1 \<and> z2]]_zs)"
using * ** by (simp add: qAFresh_qSwap_compose)
finally show ?thesis .
qed
lemma qAFreshAll_imp_qFreshAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and xs x
shows "(qAFresh xs x X \<longrightarrow> qFresh xs x X) \<and>
(qAFreshAbs xs x A \<longrightarrow> qFreshAbs xs x A)"
by(induct rule: qTerm_rawInduct, auto simp add: liftAll_def)
corollary qAFresh_imp_qFresh:
"qAFresh xs x X \<Longrightarrow> qFresh xs x X"
by(simp add: qAFreshAll_imp_qFreshAll)
lemma qSwapAll_preserves_qAFreshAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and ys y zs z1 z2
shows
"(qAFresh ys (y @ys[z1 \<and> z2]_zs) (X #[[z1 \<and> z2]]_zs) = qAFresh ys y X) \<and>
(qAFreshAbs ys (y @ys[z1 \<and> z2]_zs) (A $[[z1 \<and> z2]]_zs) = qAFreshAbs ys y A)"
proof(induct rule: qTerm_rawInduct[of _ _ X A])
case (Op delta inp binp)
then show ?case
unfolding qAFreshAll_simps(2) qSwapAll_simps(2) liftAll_lift_comp o_def
unfolding liftAll_def by presburger
qed(auto simp add: sw_def)
corollary qSwap_preserves_qAFresh[simp]:
"(qAFresh ys (y @ys[z1 \<and> z2]_zs) (X #[[z1 \<and> z2]]_zs) = qAFresh ys y X)"
by(simp add: qSwapAll_preserves_qAFreshAll)
lemma qSwap_preserves_qAFresh_distinct:
assumes "ys \<noteq> zs \<or> y \<notin> {z1,z2}"
shows "qAFresh ys y (X #[[z1 \<and> z2]]_zs) = qAFresh ys y X"
proof-
have "y = (y @ys[z1 \<and> z2]_zs)" using assms unfolding sw_def by auto
thus ?thesis using qSwap_preserves_qAFresh[of ys zs z1 z2 y] by auto
qed
lemma qAFresh_qSwap_exchange1:
"qAFresh zs z2 (X #[[z1 \<and> z2]]_zs) = qAFresh zs z1 X"
proof-
have "z2 = (z1 @zs[z1 \<and> z2]_zs)" unfolding sw_def by auto
thus ?thesis using qSwap_preserves_qAFresh[of zs zs z1 z2 z1 X] by auto
qed
lemma qAFresh_qSwap_exchange2:
"qAFresh zs z2 (X #[[z2 \<and> z1]]_zs) = qAFresh zs z1 X"
by(auto simp add: qAFresh_qSwap_exchange1 qSwap_sym)
lemma qSwapAll_preserves_qFreshAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)qTerm" and
A::"('index,'bindex,'varSort,'var,'opSym)qAbs" and ys y zs z1 z2
shows
"(qFresh ys (y @ys[z1 \<and> z2]_zs) (X #[[z1 \<and> z2]]_zs) = qFresh ys y X) \<and>
(qFreshAbs ys (y @ys[z1 \<and> z2]_zs) (A $[[z1 \<and> z2]]_zs) = qFreshAbs ys y A)"
proof(induct rule: qTerm_rawInduct[of _ _ X A])
case (Op delta inp binp)
then show ?case
unfolding qFreshAll_simps(2) qSwapAll_simps(2) liftAll_lift_comp o_def
unfolding liftAll_def by presburger
qed (auto simp add: sw_def)
corollary qSwap_preserves_qFresh:
"(qFresh ys (y @ys[z1 \<and> z2]_zs) (X #[[z1 \<and> z2]]_zs) = qFresh ys y X)"
by(simp add: qSwapAll_preserves_qFreshAll)
lemma qSwap_preserves_qFresh_distinct:
assumes "ys \<noteq> zs \<or> y \<notin> {z1,z2}"
shows "qFresh ys y (X #[[z1 \<and> z2]]_zs) = qFresh ys y X"
proof-
have "y = (y @ys[z1 \<and> z2]_zs)" using assms unfolding sw_def by auto
thus ?thesis using qSwap_preserves_qFresh[of ys zs z1 z2 y] by auto
qed
lemma qFresh_qSwap_exchange1:
"qFresh zs z2 (X #[[z1 \<and> z2]]_zs) = qFresh zs z1 X"
proof-
have "z2 = (z1 @zs[z1 \<and> z2]_zs)" unfolding sw_def by auto
thus ?thesis using qSwap_preserves_qFresh[of zs zs z1 z2 z1 X] by auto
qed
lemma qFresh_qSwap_exchange2:
"qFresh zs z1 X = qFresh zs z2 (X #[[z2 \<and> z1]]_zs)"
by (auto simp add: qFresh_qSwap_exchange1 qSwap_sym)
lemmas qSwap_qAFresh_otherSimps =
qSwap_ident qSwap_involutive qAFresh_qSwap_id qSwap_preserves_qAFresh
end
diff --git a/thys/Binding_Syntax_Theory/Recursion.thy b/thys/Binding_Syntax_Theory/Recursion.thy
--- a/thys/Binding_Syntax_Theory/Recursion.thy
+++ b/thys/Binding_Syntax_Theory/Recursion.thy
@@ -1,2110 +1,2110 @@
-section {* General Recursion *}
+section \<open>General Recursion\<close>
theory Recursion imports Iteration
begin
-text{* The initiality theorems from the previous section support iteration principles.
+text\<open>The initiality theorems from the previous section support iteration principles.
Next we extend the results to general recursion. The difference between
general recursion and iteration is that the former also considers
the (source) ``items" (terms and abstractions), and not only the
(target) generalized items, appear in the recursive clauses.
(Here is an example illustrating the above difference for the standard case
of natural numbers:
\\- Given a number n, the operator ``add-n" can be defined by iteration:
\\--- ``add-n 0 = n",
\\--- ``add-n (Suc m) = Suc (add-n m)".
Notice that, in right-hand side of the recursive clause, ``m" is not used ``directly", but only
via ``add-n" -- this makes the definition iterative. By contrast, the following
definition of predecessor is trivial form of recursion (namely, case analysis),
but is {\em not} iteration:
\\--- ``pred 0 = 0",
\\--- ``pred (Suc n) = n".
)
We achieve our desired extension by augmenting the notion of model
and then essentially inferring recursion (as customary)
from
[iteration having as target the product between the term model and the original model].
As a matter of notation: remember we are using for generalized items
the same meta-variables as for ``items" (terms and abstractions).
But now the model operators will take both items and generalized items.
We shall prime the meta-variables for items (as in X', A', etc).
- *}
+\<close>
-subsection {* Raw models *}
+subsection \<open>Raw models\<close>
record ('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs)model =
gWls :: "'sort \<Rightarrow> 'gTerm \<Rightarrow> bool"
gWlsAbs :: "'varSort \<times> 'sort \<Rightarrow> 'gAbs \<Rightarrow> bool"
(* *)
gVar :: "'varSort \<Rightarrow> 'var \<Rightarrow> 'gTerm"
gAbs ::
"'varSort \<Rightarrow> 'var \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow>
'gAbs"
gOp ::
"'opSym \<Rightarrow>
('index,('index,'bindex,'varSort,'var,'opSym)term)input \<Rightarrow> ('index,'gTerm)input \<Rightarrow>
('bindex,('index,'bindex,'varSort,'var,'opSym)abs)input \<Rightarrow> ('bindex,'gAbs)input \<Rightarrow>
'gTerm"
(* *)
gFresh ::
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow> bool"
gFreshAbs ::
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow> 'gAbs \<Rightarrow> bool"
(* *)
gSwap ::
"'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow>
'gTerm"
gSwapAbs ::
"'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow> 'gAbs \<Rightarrow>
'gAbs"
(* *)
gSubst ::
"'varSort \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow>
'var \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow>
'gTerm"
gSubstAbs ::
"'varSort \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm \<Rightarrow>
'var \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow> 'gAbs \<Rightarrow>
'gAbs"
-subsection {* Well-sorted models of various kinds *}
+subsection \<open>Well-sorted models of various kinds\<close>
-text{* Lifting the model operations to inputs *}
+text\<open>Lifting the model operations to inputs\<close>
definition gFreshInp where
"gFreshInp MOD ys y inp' inp \<equiv> liftAll2 (gFresh MOD ys y) inp' inp"
definition gFreshBinp where
"gFreshBinp MOD ys y binp' binp \<equiv> liftAll2 (gFreshAbs MOD ys y) binp' binp"
definition gSwapInp where
"gSwapInp MOD zs z1 z2 inp' inp \<equiv> lift2 (gSwap MOD zs z1 z2) inp' inp"
definition gSwapBinp where
"gSwapBinp MOD zs z1 z2 binp' binp \<equiv> lift2 (gSwapAbs MOD zs z1 z2) binp' binp"
definition gSubstInp where
"gSubstInp MOD ys Y' Y y inp' inp \<equiv> lift2 (gSubst MOD ys Y' Y y) inp' inp"
definition gSubstBinp where
"gSubstBinp MOD ys Y' Y y binp' binp \<equiv> lift2 (gSubstAbs MOD ys Y' Y y) binp' binp"
(* *************************************************** *)
context FixSyn
begin
definition gWlsInp where
"gWlsInp MOD delta inp \<equiv>
wlsOpS delta \<and> sameDom (arOf delta) inp \<and> liftAll2 (gWls MOD) (arOf delta) inp"
lemmas gWlsInp_defs = gWlsInp_def sameDom_def liftAll2_def
definition gWlsBinp where
"gWlsBinp MOD delta binp \<equiv>
wlsOpS delta \<and> sameDom (barOf delta) binp \<and> liftAll2 (gWlsAbs MOD) (barOf delta) binp"
lemmas gWlsBinp_defs = gWlsBinp_def sameDom_def liftAll2_def
-text{* Basic properties of the lifted model operations *}
+text\<open>Basic properties of the lifted model operations\<close>
-text{* . for free inputs: *}
+text\<open>. for free inputs:\<close>
lemma sameDom_swapInp_gSwapInp[simp]:
assumes "wlsInp delta inp'" and "gWlsInp MOD delta inp"
shows "sameDom (swapInp zs z1 z2 inp') (gSwapInp MOD zs z1 z2 inp' inp)"
using assms by(simp add: wlsInp_iff gWlsInp_def swapInp_def gSwapInp_def
liftAll2_def lift_def lift2_def sameDom_def split: option.splits)
lemma sameDom_substInp_gSubstInp[simp]:
assumes "wlsInp delta inp'" and "gWlsInp MOD delta inp"
shows "sameDom (substInp ys Y' y inp') (gSubstInp MOD ys Y' Y y inp' inp)"
using assms by(simp add: wlsInp_iff gWlsInp_def substInp_def2 gSubstInp_def
liftAll2_def lift_def lift2_def sameDom_def split: option.splits)
-text{* . for bound inputs: *}
+text\<open>. for bound inputs:\<close>
lemma sameDom_swapBinp_gSwapBinp[simp]:
assumes "wlsBinp delta binp'" and "gWlsBinp MOD delta binp"
shows "sameDom (swapBinp zs z1 z2 binp') (gSwapBinp MOD zs z1 z2 binp' binp)"
using assms by(simp add: wlsBinp_iff gWlsBinp_def swapBinp_def gSwapBinp_def
liftAll2_def lift_def lift2_def sameDom_def split: option.splits)
lemma sameDom_substBinp_gSubstBinp[simp]:
assumes "wlsBinp delta binp'" and "gWlsBinp MOD delta binp"
shows "sameDom (substBinp ys Y' y binp') (gSubstBinp MOD ys Y' Y y binp' binp)"
using assms by(simp add: wlsBinp_iff gWlsBinp_def substBinp_def2 gSubstBinp_def
liftAll2_def lift_def lift2_def sameDom_def split: option.splits)
lemmas sameDom_gInput_simps =
sameDom_swapInp_gSwapInp sameDom_substInp_gSubstInp
sameDom_swapBinp_gSwapBinp sameDom_substBinp_gSubstBinp
-text{* Domain disjointness: *}
+text\<open>Domain disjointness:\<close>
definition gWlsDisj where
"gWlsDisj MOD \<equiv> \<forall> s s' X. gWls MOD s X \<and> gWls MOD s' X \<longrightarrow> s = s'"
definition gWlsAbsDisj where
"gWlsAbsDisj MOD \<equiv> \<forall> xs s xs' s' A.
isInBar (xs,s) \<and> isInBar (xs',s') \<and>
gWlsAbs MOD (xs,s) A \<and> gWlsAbs MOD (xs',s') A
\<longrightarrow> xs = xs' \<and> s = s'"
definition gWlsAllDisj where
"gWlsAllDisj MOD \<equiv> gWlsDisj MOD \<and> gWlsAbsDisj MOD"
lemmas gWlsAllDisj_defs =
gWlsAllDisj_def gWlsDisj_def gWlsAbsDisj_def
-text {* Abstraction domains inhabited only within bound arities: *}
+text \<open>Abstraction domains inhabited only within bound arities:\<close>
definition gWlsAbsIsInBar where
"gWlsAbsIsInBar MOD \<equiv> \<forall> us s A. gWlsAbs MOD (us,s) A \<longrightarrow> isInBar (us,s)"
-text{* Domain preservation by the operators *}
+text\<open>Domain preservation by the operators\<close>
-text{* The constructs preserve the domains: *}
+text\<open>The constructs preserve the domains:\<close>
definition gVarPresGWls where
"gVarPresGWls MOD \<equiv> \<forall> xs x. gWls MOD (asSort xs) (gVar MOD xs x)"
definition gAbsPresGWls where
"gAbsPresGWls MOD \<equiv> \<forall> xs s x X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
gWlsAbs MOD (xs,s) (gAbs MOD xs x X' X)"
definition gOpPresGWls where
"gOpPresGWls MOD \<equiv> \<forall> delta inp' inp binp' binp.
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<and> wlsBinp delta binp' \<and> gWlsBinp MOD delta binp
\<longrightarrow> gWls MOD (stOf delta) (gOp MOD delta inp' inp binp' binp)"
definition gConsPresGWls where
"gConsPresGWls MOD \<equiv> gVarPresGWls MOD \<and> gAbsPresGWls MOD \<and> gOpPresGWls MOD"
lemmas gConsPresGWls_defs = gConsPresGWls_def
gVarPresGWls_def gAbsPresGWls_def gOpPresGWls_def
-text{* ``swap" preserves the domains: *}
+text\<open>``swap" preserves the domains:\<close>
definition gSwapPresGWls where
"gSwapPresGWls MOD \<equiv> \<forall> zs z1 z2 s X' X.
wls s X' \<and> gWls MOD s X \<longrightarrow>
gWls MOD s (gSwap MOD zs z1 z2 X' X)"
definition gSwapAbsPresGWlsAbs where
"gSwapAbsPresGWlsAbs MOD \<equiv> \<forall> zs z1 z2 us s A' A.
isInBar (us,s) \<and> wlsAbs (us,s) A' \<and> gWlsAbs MOD (us,s) A \<longrightarrow>
gWlsAbs MOD (us,s) (gSwapAbs MOD zs z1 z2 A' A)"
definition gSwapAllPresGWlsAll where
"gSwapAllPresGWlsAll MOD \<equiv> gSwapPresGWls MOD \<and> gSwapAbsPresGWlsAbs MOD"
lemmas gSwapAllPresGWlsAll_defs =
gSwapAllPresGWlsAll_def gSwapPresGWls_def gSwapAbsPresGWlsAbs_def
-text{* ``subst" preserves the domains: *}
+text\<open>``subst" preserves the domains:\<close>
definition gSubstPresGWls where
"gSubstPresGWls MOD \<equiv> \<forall> ys Y' Y y s X' X.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
gWls MOD s (gSubst MOD ys Y' Y y X' X)"
definition gSubstAbsPresGWlsAbs where
"gSubstAbsPresGWlsAbs MOD \<equiv> \<forall> ys Y' Y y us s A' A.
isInBar (us,s) \<and>
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and> wlsAbs (us,s) A' \<and> gWlsAbs MOD (us,s) A \<longrightarrow>
gWlsAbs MOD (us,s) (gSubstAbs MOD ys Y' Y y A' A)"
definition gSubstAllPresGWlsAll where
"gSubstAllPresGWlsAll MOD \<equiv> gSubstPresGWls MOD \<and> gSubstAbsPresGWlsAbs MOD"
lemmas gSubstAllPresGWlsAll_defs =
gSubstAllPresGWlsAll_def gSubstPresGWls_def gSubstAbsPresGWlsAbs_def
-text{* Clauses for fresh: *}
+text\<open>Clauses for fresh:\<close>
definition gFreshGVar where
"gFreshGVar MOD \<equiv> \<forall> ys y xs x.
(ys \<noteq> xs \<or> y \<noteq> x) \<longrightarrow>
gFresh MOD ys y (Var xs x) (gVar MOD xs x)"
definition gFreshGAbs1 where
"gFreshGAbs1 MOD \<equiv> \<forall> ys y s X' X.
isInBar (ys,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
gFreshAbs MOD ys y (Abs ys y X') (gAbs MOD ys y X' X)"
definition gFreshGAbs2 where
"gFreshGAbs2 MOD \<equiv> \<forall> ys y xs x s X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
fresh ys y X' \<and> gFresh MOD ys y X' X \<longrightarrow>
gFreshAbs MOD ys y (Abs xs x X') (gAbs MOD xs x X' X)"
definition gFreshGOp where
"gFreshGOp MOD \<equiv> \<forall> ys y delta inp' inp binp' binp.
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<and> wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
freshInp ys y inp' \<and> gFreshInp MOD ys y inp' inp \<and>
freshBinp ys y binp' \<and> gFreshBinp MOD ys y binp' binp \<longrightarrow>
gFresh MOD ys y (Op delta inp' binp') (gOp MOD delta inp' inp binp' binp)"
definition gFreshCls where
"gFreshCls MOD \<equiv> gFreshGVar MOD \<and> gFreshGAbs1 MOD \<and> gFreshGAbs2 MOD \<and> gFreshGOp MOD"
lemmas gFreshCls_defs = gFreshCls_def
gFreshGVar_def gFreshGAbs1_def gFreshGAbs2_def gFreshGOp_def
(* Clauses for swap: fully-conditional versions and less-conditional,
stronger versions (suffix ``STR") *)
definition gSwapGVar where
"gSwapGVar MOD \<equiv> \<forall> zs z1 z2 xs x.
gSwap MOD zs z1 z2 (Var xs x) (gVar MOD xs x) =
gVar MOD xs (x @xs[z1 \<and> z2]_zs)"
definition gSwapGAbs where
"gSwapGAbs MOD \<equiv> \<forall> zs z1 z2 xs x s X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
gSwapAbs MOD zs z1 z2 (Abs xs x X') (gAbs MOD xs x X' X) =
gAbs MOD xs (x @xs[z1 \<and> z2]_zs) (X' #[z1 \<and> z2]_zs) (gSwap MOD zs z1 z2 X' X)"
definition gSwapGOp where
"gSwapGOp MOD \<equiv> \<forall> zs z1 z2 delta inp' inp binp' binp.
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<and> wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
gSwap MOD zs z1 z2 (Op delta inp' binp') (gOp MOD delta inp' inp binp' binp) =
gOp MOD delta
(inp' %[z1 \<and> z2]_zs) (gSwapInp MOD zs z1 z2 inp' inp)
(binp' %%[z1 \<and> z2]_zs) (gSwapBinp MOD zs z1 z2 binp' binp)"
definition gSwapCls where
"gSwapCls MOD \<equiv> gSwapGVar MOD \<and> gSwapGAbs MOD \<and> gSwapGOp MOD"
lemmas gSwapCls_defs = gSwapCls_def
gSwapGVar_def gSwapGAbs_def gSwapGOp_def
(* Clauses for subst: *)
definition gSubstGVar1 where
"gSubstGVar1 MOD \<equiv> \<forall> ys y Y' Y xs x.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<longrightarrow>
(ys \<noteq> xs \<or> y \<noteq> x) \<longrightarrow>
gSubst MOD ys Y' Y y (Var xs x) (gVar MOD xs x) =
gVar MOD xs x"
definition gSubstGVar2 where
"gSubstGVar2 MOD \<equiv> \<forall> ys y Y' Y.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<longrightarrow>
gSubst MOD ys Y' Y y (Var ys y) (gVar MOD ys y) = Y"
definition gSubstGAbs where
"gSubstGAbs MOD \<equiv> \<forall> ys y Y' Y xs x s X' X.
isInBar (xs,s) \<and>
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and>
wls s X' \<and> gWls MOD s X \<longrightarrow>
(xs \<noteq> ys \<or> x \<noteq> y) \<and> fresh xs x Y' \<and> gFresh MOD xs x Y' Y \<longrightarrow>
gSubstAbs MOD ys Y' Y y (Abs xs x X') (gAbs MOD xs x X' X) =
gAbs MOD xs x (X' #[Y' / y]_ys) (gSubst MOD ys Y' Y y X' X)"
definition gSubstGOp where
"gSubstGOp MOD \<equiv> \<forall> ys y Y' Y delta inp' inp binp' binp.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and>
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<and>
wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
gSubst MOD ys Y' Y y (Op delta inp' binp') (gOp MOD delta inp' inp binp' binp) =
gOp MOD delta
(inp' %[Y' / y]_ys) (gSubstInp MOD ys Y' Y y inp' inp)
(binp' %%[Y' / y]_ys) (gSubstBinp MOD ys Y' Y y binp' binp)"
definition gSubstCls where
"gSubstCls MOD \<equiv> gSubstGVar1 MOD \<and> gSubstGVar2 MOD \<and> gSubstGAbs MOD \<and> gSubstGOp MOD"
lemmas gSubstCls_defs = gSubstCls_def
gSubstGVar1_def gSubstGVar2_def gSubstGAbs_def gSubstGOp_def
(* Freshness-based congruence for abstractions: *)
(* ... employing swap: *)
definition gAbsCongS where
"gAbsCongS MOD \<equiv> \<forall> xs x x2 y s X' X X2' X2.
isInBar (xs,s) \<and>
wls s X' \<and> gWls MOD s X \<and>
wls s X2' \<and> gWls MOD s X2 \<longrightarrow>
fresh xs y X' \<and> gFresh MOD xs y X' X \<and>
fresh xs y X2' \<and> gFresh MOD xs y X2' X2 \<and>
(X' #[y \<and> x]_xs) = (X2' #[y \<and> x2]_xs) \<longrightarrow>
gSwap MOD xs y x X' X = gSwap MOD xs y x2 X2' X2 \<longrightarrow>
gAbs MOD xs x X' X = gAbs MOD xs x2 X2' X2"
(* ... Note: no need for congruence employing subst (as it is not used in the
definition of rmorphisms *)
(* Abstraction renaming: *)
definition gAbsRen where
"gAbsRen MOD \<equiv> \<forall> xs y x s X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
fresh xs y X' \<and> gFresh MOD xs y X' X \<longrightarrow>
gAbs MOD xs y (X' #[y // x]_xs) (gSubst MOD xs (Var xs y) (gVar MOD xs y) x X' X) =
gAbs MOD xs x X' X"
-text {* Well-sorted fresh-swap models: *}
+text \<open>Well-sorted fresh-swap models:\<close>
definition wlsFSw where
"wlsFSw MOD \<equiv> gWlsAllDisj MOD \<and> gWlsAbsIsInBar MOD \<and>
gConsPresGWls MOD \<and> gSwapAllPresGWlsAll MOD \<and>
gFreshCls MOD \<and> gSwapCls MOD \<and> gAbsCongS MOD"
lemmas wlsFSw_defs1 = wlsFSw_def
gWlsAllDisj_def gWlsAbsIsInBar_def
gConsPresGWls_def gSwapAllPresGWlsAll_def
gFreshCls_def gSwapCls_def gAbsCongS_def
lemmas wlsFSw_defs = wlsFSw_def
gWlsAllDisj_defs gWlsAbsIsInBar_def
gConsPresGWls_defs gSwapAllPresGWlsAll_defs
gFreshCls_defs gSwapCls_defs gAbsCongS_def
-text {* Well-sorted fresh-subst models: *}
+text \<open>Well-sorted fresh-subst models:\<close>
definition wlsFSb where
"wlsFSb MOD \<equiv> gWlsAllDisj MOD \<and> gWlsAbsIsInBar MOD \<and>
gConsPresGWls MOD \<and> gSubstAllPresGWlsAll MOD \<and>
gFreshCls MOD \<and> gSubstCls MOD \<and> gAbsRen MOD"
lemmas wlsFSb_defs1 = wlsFSb_def
gWlsAllDisj_def gWlsAbsIsInBar_def
gConsPresGWls_def gSubstAllPresGWlsAll_def
gFreshCls_def gSubstCls_def gAbsRen_def
lemmas wlsFSb_defs = wlsFSb_def
gWlsAllDisj_defs gWlsAbsIsInBar_def
gConsPresGWls_defs gSubstAllPresGWlsAll_defs
gFreshCls_defs gSubstCls_defs gAbsRen_def
-text {* Well-sorted fresh-swap-subst-models *}
+text \<open>Well-sorted fresh-swap-subst-models\<close>
definition wlsFSwSb where
"wlsFSwSb MOD \<equiv> wlsFSw MOD \<and> gSubstAllPresGWlsAll MOD \<and> gSubstCls MOD"
lemmas wlsFSwSb_defs1 = wlsFSwSb_def
wlsFSw_def gSubstAllPresGWlsAll_def gSubstCls_def
lemmas wlsFSwSb_defs = wlsFSwSb_def
wlsFSw_def gSubstAllPresGWlsAll_defs gSubstCls_defs
-text {* Well-sorted fresh-subst-swap-models *}
+text \<open>Well-sorted fresh-subst-swap-models\<close>
definition wlsFSbSw where
"wlsFSbSw MOD \<equiv> wlsFSb MOD \<and> gSwapAllPresGWlsAll MOD \<and> gSwapCls MOD"
lemmas wlsFSbSw_defs1 = wlsFSbSw_def
wlsFSw_def gSwapAllPresGWlsAll_def gSwapCls_def
lemmas wlsFSbSw_defs = wlsFSbSw_def
wlsFSw_def gSwapAllPresGWlsAll_defs gSwapCls_defs
-text{* Extension of domain preservation (by swap and subst) to inputs: *}
+text\<open>Extension of domain preservation (by swap and subst) to inputs:\<close>
-text {* First for free inputs: *}
+text \<open>First for free inputs:\<close>
definition gSwapInpPresGWlsInp where
"gSwapInpPresGWlsInp MOD \<equiv> \<forall> zs z1 z2 delta inp' inp.
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<longrightarrow>
gWlsInp MOD delta (gSwapInp MOD zs z1 z2 inp' inp)"
definition gSubstInpPresGWlsInp where
"gSubstInpPresGWlsInp MOD \<equiv> \<forall> ys y Y' Y delta inp' inp.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and>
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<longrightarrow>
gWlsInp MOD delta (gSubstInp MOD ys Y' Y y inp' inp)"
lemma imp_gSwapInpPresGWlsInp:
"gSwapPresGWls MOD \<Longrightarrow> gSwapInpPresGWlsInp MOD"
by (auto simp: lift2_def liftAll2_def sameDom_def wlsInp_iff gWlsInp_def
gSwapPresGWls_def gSwapInpPresGWlsInp_def gSwapInp_def
split: option.splits)
lemma imp_gSubstInpPresGWlsInp:
"gSubstPresGWls MOD \<Longrightarrow> gSubstInpPresGWlsInp MOD"
by (auto simp: lift2_def liftAll2_def sameDom_def wlsInp_iff gWlsInp_def
gSubstPresGWls_def gSubstInpPresGWlsInp_def gSubstInp_def
split: option.splits)
-text {* Then for bound inputs: *}
+text \<open>Then for bound inputs:\<close>
definition gSwapBinpPresGWlsBinp where
"gSwapBinpPresGWlsBinp MOD \<equiv> \<forall> zs z1 z2 delta binp' binp.
wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
gWlsBinp MOD delta (gSwapBinp MOD zs z1 z2 binp' binp)"
definition gSubstBinpPresGWlsBinp where
"gSubstBinpPresGWlsBinp MOD \<equiv> \<forall> ys y Y' Y delta binp' binp.
wls (asSort ys) Y' \<and> gWls MOD (asSort ys) Y \<and>
wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
gWlsBinp MOD delta (gSubstBinp MOD ys Y' Y y binp' binp)"
lemma imp_gSwapBinpPresGWlsBinp:
"gSwapAbsPresGWlsAbs MOD \<Longrightarrow> gSwapBinpPresGWlsBinp MOD"
by (auto simp: lift2_def liftAll2_def sameDom_def wlsBinp_iff gWlsBinp_def
gSwapAbsPresGWlsAbs_def gSwapBinpPresGWlsBinp_def gSwapBinp_def
split: option.splits)
lemma imp_gSubstBinpPresGWlsBinp:
"gSubstAbsPresGWlsAbs MOD \<Longrightarrow> gSubstBinpPresGWlsBinp MOD"
by (auto simp: lift2_def liftAll2_def sameDom_def wlsBinp_iff gWlsBinp_def
gSubstAbsPresGWlsAbs_def gSubstBinpPresGWlsBinp_def gSubstBinp_def
split: option.splits)
-subsection{* Model morphisms from the term model *}
+subsection\<open>Model morphisms from the term model\<close>
definition presWls where
"presWls h MOD \<equiv> \<forall> s X. wls s X \<longrightarrow> gWls MOD s (h X)"
definition presWlsAbs where
"presWlsAbs hA MOD \<equiv> \<forall> us s A. wlsAbs (us,s) A \<longrightarrow> gWlsAbs MOD (us,s) (hA A)"
definition presWlsAll where
"presWlsAll h hA MOD \<equiv> presWls h MOD \<and> presWlsAbs hA MOD"
lemmas presWlsAll_defs = presWlsAll_def presWls_def presWlsAbs_def
definition presVar where
"presVar h MOD \<equiv> \<forall> xs x. h (Var xs x) = gVar MOD xs x"
definition presAbs where
"presAbs h hA MOD \<equiv> \<forall> xs x s X.
isInBar (xs,s) \<and> wls s X \<longrightarrow>
hA (Abs xs x X) = gAbs MOD xs x X (h X)"
definition presOp where
"presOp h hA MOD \<equiv> \<forall> delta inp binp.
wlsInp delta inp \<and> wlsBinp delta binp \<longrightarrow>
h (Op delta inp binp) =
gOp MOD delta inp (lift h inp) binp (lift hA binp)"
definition presCons where
"presCons h hA MOD \<equiv> presVar h MOD \<and> presAbs h hA MOD \<and> presOp h hA MOD"
lemmas presCons_defs = presCons_def
presVar_def presAbs_def presOp_def
definition presFresh where
"presFresh h MOD \<equiv> \<forall> ys y s X.
wls s X \<longrightarrow>
fresh ys y X \<longrightarrow> gFresh MOD ys y X (h X)"
definition presFreshAbs where
"presFreshAbs hA MOD \<equiv> \<forall> ys y us s A.
wlsAbs (us,s) A \<longrightarrow>
freshAbs ys y A \<longrightarrow> gFreshAbs MOD ys y A (hA A)"
definition presFreshAll where
"presFreshAll h hA MOD \<equiv> presFresh h MOD \<and> presFreshAbs hA MOD"
lemmas presFreshAll_defs = presFreshAll_def
presFresh_def presFreshAbs_def
definition presSwap where
"presSwap h MOD \<equiv> \<forall> zs z1 z2 s X.
wls s X \<longrightarrow>
h (X #[z1 \<and> z2]_zs) = gSwap MOD zs z1 z2 X (h X)"
definition presSwapAbs where
"presSwapAbs hA MOD \<equiv> \<forall> zs z1 z2 us s A.
wlsAbs (us,s) A \<longrightarrow>
hA (A $[z1 \<and> z2]_zs) = gSwapAbs MOD zs z1 z2 A (hA A)"
definition presSwapAll where
"presSwapAll h hA MOD \<equiv> presSwap h MOD \<and> presSwapAbs hA MOD"
lemmas presSwapAll_defs = presSwapAll_def
presSwap_def presSwapAbs_def
definition presSubst where
"presSubst h MOD \<equiv> \<forall> ys Y y s X.
wls (asSort ys) Y \<and> wls s X \<longrightarrow>
h (subst ys Y y X) = gSubst MOD ys Y (h Y) y X (h X)"
definition presSubstAbs where
"presSubstAbs h hA MOD \<equiv> \<forall> ys Y y us s A.
wls (asSort ys) Y \<and> wlsAbs (us,s) A \<longrightarrow>
hA (A $[Y / y]_ys) = gSubstAbs MOD ys Y (h Y) y A (hA A)"
definition presSubstAll where
"presSubstAll h hA MOD \<equiv> presSubst h MOD \<and> presSubstAbs h hA MOD"
lemmas presSubstAll_defs = presSubstAll_def
presSubst_def presSubstAbs_def
definition termFSwMorph where
"termFSwMorph h hA MOD \<equiv> presWlsAll h hA MOD \<and> presCons h hA MOD \<and>
presFreshAll h hA MOD \<and> presSwapAll h hA MOD"
lemmas termFSwMorph_defs1 = termFSwMorph_def
presWlsAll_def presCons_def presFreshAll_def presSwapAll_def
lemmas termFSwMorph_defs = termFSwMorph_def
presWlsAll_defs presCons_defs presFreshAll_defs presSwapAll_defs
definition termFSbMorph where
"termFSbMorph h hA MOD \<equiv> presWlsAll h hA MOD \<and> presCons h hA MOD \<and>
presFreshAll h hA MOD \<and> presSubstAll h hA MOD"
lemmas termFSbMorph_defs1 = termFSbMorph_def
presWlsAll_def presCons_def presFreshAll_def presSubstAll_def
lemmas termFSbMorph_defs = termFSbMorph_def
presWlsAll_defs presCons_defs presFreshAll_defs presSubstAll_defs
definition termFSwSbMorph where
"termFSwSbMorph h hA MOD \<equiv> termFSwMorph h hA MOD \<and> presSubstAll h hA MOD"
lemmas termFSwSbMorph_defs1 = termFSwSbMorph_def
termFSwMorph_def presSubstAll_def
lemmas termFSwSbMorph_defs = termFSwSbMorph_def
termFSwMorph_defs presSubstAll_defs
-text{* Extension of domain preservation (by the morphisms) to inputs *}
+text\<open>Extension of domain preservation (by the morphisms) to inputs\<close>
-text{* . for free inputs: *}
+text\<open>. for free inputs:\<close>
lemma presWls_wlsInp:
"wlsInp delta inp \<Longrightarrow> presWls h MOD \<Longrightarrow> gWlsInp MOD delta (lift h inp)"
by(auto simp: wlsInp_iff gWlsInp_def lift_def liftAll2_def sameDom_def
presWls_def split: option.splits)
-text{* . for bound inputs: *}
+text\<open>. for bound inputs:\<close>
lemma presWls_wlsBinp:
"wlsBinp delta binp \<Longrightarrow> presWlsAbs hA MOD \<Longrightarrow> gWlsBinp MOD delta (lift hA binp)"
by(auto simp: wlsBinp_iff gWlsBinp_def lift_def liftAll2_def sameDom_def
presWlsAbs_def split: option.splits)
-subsection {* From models to iterative models *}
+subsection \<open>From models to iterative models\<close>
-text {* The transition map: *}
+text \<open>The transition map:\<close>
definition fromMOD ::
"('index,'bindex,'varSort,'sort,'opSym,'var,'gTerm,'gAbs) model
\<Rightarrow>
('index,'bindex,'varSort,'sort,'opSym,'var,
('index,'bindex,'varSort,'var,'opSym)term \<times> 'gTerm,
('index,'bindex,'varSort,'var,'opSym)abs \<times> 'gAbs) Iteration.model"
where
"fromMOD MOD \<equiv>
\<lparr>
igWls = \<lambda>s X'X. wls s (fst X'X) \<and> gWls MOD s (snd X'X),
igWlsAbs = \<lambda>us_s A'A. wlsAbs us_s (fst A'A) \<and> gWlsAbs MOD us_s (snd A'A),
igVar = \<lambda>xs x. (Var xs x, gVar MOD xs x),
igAbs = \<lambda>xs x X'X. (Abs xs x (fst X'X), gAbs MOD xs x (fst X'X) (snd X'X)),
igOp =
\<lambda>delta iinp biinp.
(Op delta (lift fst iinp) (lift fst biinp),
gOp MOD delta
(lift fst iinp) (lift snd iinp)
(lift fst biinp) (lift snd biinp)),
igFresh =
\<lambda>ys y X'X. fresh ys y (fst X'X) \<and> gFresh MOD ys y (fst X'X) (snd X'X),
igFreshAbs =
\<lambda>ys y A'A. freshAbs ys y (fst A'A) \<and> gFreshAbs MOD ys y (fst A'A) (snd A'A),
igSwap =
\<lambda>zs z1 z2 X'X. ((fst X'X) #[z1 \<and> z2]_zs, gSwap MOD zs z1 z2 (fst X'X) (snd X'X)),
igSwapAbs =
\<lambda>zs z1 z2 A'A. ((fst A'A) $[z1 \<and> z2]_zs, gSwapAbs MOD zs z1 z2 (fst A'A) (snd A'A)),
igSubst =
\<lambda>ys Y'Y y X'X.
((fst X'X) #[(fst Y'Y) / y]_ys,
gSubst MOD ys (fst Y'Y) (snd Y'Y) y (fst X'X) (snd X'X)),
igSubstAbs =
\<lambda>ys Y'Y y A'A.
((fst A'A) $[(fst Y'Y) / y]_ys,
gSubstAbs MOD ys (fst Y'Y) (snd Y'Y) y (fst A'A) (snd A'A))
\<rparr>"
-text{* Basic simplification rules: *}
+text\<open>Basic simplification rules:\<close>
lemma fromMOD_basic_simps[simp]:
"igWls (fromMOD MOD) s X'X =
(wls s (fst X'X) \<and> gWls MOD s (snd X'X))"
(* *)
"igWlsAbs (fromMOD MOD) us_s A'A =
(wlsAbs us_s (fst A'A) \<and> gWlsAbs MOD us_s (snd A'A))"
(* *)
"igVar (fromMOD MOD) xs x = (Var xs x, gVar MOD xs x)"
(* *)
"igAbs (fromMOD MOD) xs x X'X = (Abs xs x (fst X'X), gAbs MOD xs x (fst X'X) (snd X'X))"
(* *)
"igOp (fromMOD MOD) delta iinp biinp =
(Op delta (lift fst iinp) (lift fst biinp),
gOp MOD delta
(lift fst iinp) (lift snd iinp)
(lift fst biinp) (lift snd biinp))"
(* *)
"igFresh (fromMOD MOD) ys y X'X =
(fresh ys y (fst X'X) \<and> gFresh MOD ys y (fst X'X) (snd X'X))"
(* *)
"igFreshAbs (fromMOD MOD) ys y A'A =
(freshAbs ys y (fst A'A) \<and> gFreshAbs MOD ys y (fst A'A) (snd A'A))"
(* *)
"igSwap (fromMOD MOD) zs z1 z2 X'X =
((fst X'X) #[z1 \<and> z2]_zs, gSwap MOD zs z1 z2 (fst X'X) (snd X'X))"
(* *)
"igSwapAbs (fromMOD MOD) zs z1 z2 A'A =
((fst A'A) $[z1 \<and> z2]_zs, gSwapAbs MOD zs z1 z2 (fst A'A) (snd A'A))"
(* *)
"igSubst (fromMOD MOD) ys Y'Y y X'X =
((fst X'X) #[(fst Y'Y) / y]_ys,
gSubst MOD ys (fst Y'Y) (snd Y'Y) y (fst X'X) (snd X'X))"
(* *)
"igSubstAbs (fromMOD MOD) ys Y'Y y A'A =
((fst A'A) $[(fst Y'Y) / y]_ys,
gSubstAbs MOD ys (fst Y'Y) (snd Y'Y) y (fst A'A) (snd A'A))"
unfolding fromMOD_def by auto
-text{* Simps for inputs *}
+text\<open>Simps for inputs\<close>
-text{* . for free inputs: *}
+text\<open>. for free inputs:\<close>
lemma igWlsInp_fromMOD[simp]:
"igWlsInp (fromMOD MOD) delta iinp \<longleftrightarrow>
wlsInp delta (lift fst iinp) \<and> gWlsInp MOD delta (lift snd iinp)"
apply (intro iffI)
subgoal apply(simp add: liftAll2_def lift_def sameDom_def
igWlsInp_def wlsInp_iff gWlsInp_def split: option.splits) .
subgoal
unfolding liftAll2_def lift_def sameDom_def
igWlsInp_def wlsInp_iff gWlsInp_def
by simp (metis (no_types, lifting) eq_snd_iff fstI option.case_eq_if
option.distinct(1) option.simps(5)) .
lemma igFreshInp_fromMOD[simp]:
"igFreshInp (fromMOD MOD) ys y iinp \<longleftrightarrow>
freshInp ys y (lift fst iinp) \<and> gFreshInp MOD ys y (lift fst iinp) (lift snd iinp)"
by (auto simp: igFreshInp_def gFreshInp_def freshInp_def
liftAll2_def liftAll_def lift_def split: option.splits)
lemma igSwapInp_fromMOD[simp]:
"igSwapInp (fromMOD MOD) zs z1 z2 iinp =
lift2 Pair
(swapInp zs z1 z2 (lift fst iinp))
(gSwapInp MOD zs z1 z2 (lift fst iinp) (lift snd iinp))"
by(auto simp: igSwapInp_def swapInp_def gSwapInp_def lift_def lift2_def
split: option.splits)
lemma igSubstInp_fromMOD[simp]:
"igSubstInp (fromMOD MOD) ys Y'Y y iinp =
lift2 Pair
(substInp ys (fst Y'Y) y (lift fst iinp))
(gSubstInp MOD ys (fst Y'Y) (snd Y'Y) y (lift fst iinp) (lift snd iinp))"
by(auto simp: igSubstInp_def substInp_def2 gSubstInp_def lift_def lift2_def
split: option.splits)
lemmas input_fromMOD_simps =
igWlsInp_fromMOD igFreshInp_fromMOD igSwapInp_fromMOD igSubstInp_fromMOD
-text{* . for bound inputs: *}
+text\<open>. for bound inputs:\<close>
lemma igWlsBinp_fromMOD[simp]:
"igWlsBinp (fromMOD MOD) delta biinp \<longleftrightarrow>
(wlsBinp delta (lift fst biinp) \<and> gWlsBinp MOD delta (lift snd biinp))"
apply (intro iffI)
subgoal apply(simp add: liftAll2_def lift_def sameDom_def
igWlsBinp_def wlsBinp_iff gWlsBinp_def split: option.splits) .
subgoal
unfolding liftAll2_def lift_def sameDom_def
igWlsBinp_def wlsBinp_iff gWlsBinp_def
by simp (metis (no_types, lifting) eq_snd_iff fstI option.case_eq_if
option.distinct(1) option.simps(5)) .
lemma igFreshBinp_fromMOD[simp]:
"igFreshBinp (fromMOD MOD) ys y biinp \<longleftrightarrow>
(freshBinp ys y (lift fst biinp) \<and>
gFreshBinp MOD ys y (lift fst biinp) (lift snd biinp))"
by (auto simp: igFreshBinp_def gFreshBinp_def freshBinp_def
liftAll2_def liftAll_def lift_def split: option.splits)
lemma igSwapBinp_fromMOD[simp]:
"igSwapBinp (fromMOD MOD) zs z1 z2 biinp =
lift2 Pair
(swapBinp zs z1 z2 (lift fst biinp))
(gSwapBinp MOD zs z1 z2 (lift fst biinp) (lift snd biinp))"
by(auto simp: igSwapBinp_def swapBinp_def gSwapBinp_def lift_def lift2_def
split: option.splits)
lemma igSubstBinp_fromMOD[simp]:
"igSubstBinp (fromMOD MOD) ys Y'Y y biinp =
lift2 Pair
(substBinp ys (fst Y'Y) y (lift fst biinp))
(gSubstBinp MOD ys (fst Y'Y) (snd Y'Y) y (lift fst biinp) (lift snd biinp))"
by(auto simp: igSubstBinp_def substBinp_def2 gSubstBinp_def lift_def lift2_def
split: option.splits)
lemmas binput_fromMOD_simps =
igWlsBinp_fromMOD igFreshBinp_fromMOD igSwapBinp_fromMOD igSubstBinp_fromMOD
-text{* Domain disjointness: *}
+text\<open>Domain disjointness:\<close>
lemma igWlsDisj_fromMOD[simp]:
"gWlsDisj MOD \<Longrightarrow> igWlsDisj (fromMOD MOD)"
unfolding igWlsDisj_def gWlsDisj_def by auto
lemma igWlsAbsDisj_fromMOD[simp]:
"gWlsAbsDisj MOD \<Longrightarrow> igWlsAbsDisj (fromMOD MOD)"
unfolding igWlsAbsDisj_def gWlsAbsDisj_def by fastforce
lemma igWlsAllDisj_fromMOD[simp]:
"gWlsAllDisj MOD \<Longrightarrow> igWlsAllDisj (fromMOD MOD)"
unfolding igWlsAllDisj_def gWlsAllDisj_def by fastforce
lemmas igWlsAllDisj_fromMOD_simps =
igWlsDisj_fromMOD igWlsAbsDisj_fromMOD igWlsAllDisj_fromMOD
-text{* Abstractions only within IsInBar: *}
+text\<open>Abstractions only within IsInBar:\<close>
lemma igWlsAbsIsInBar_fromMOD[simp]:
"gWlsAbsIsInBar MOD \<Longrightarrow> igWlsAbsIsInBar (fromMOD MOD)"
unfolding gWlsAbsIsInBar_def igWlsAbsIsInBar_def by simp
-text{* The constructs preserve the domains: *}
+text\<open>The constructs preserve the domains:\<close>
lemma igVarIPresIGWls_fromMOD[simp]:
"gVarPresGWls MOD \<Longrightarrow> igVarIPresIGWls (fromMOD MOD)"
unfolding igVarIPresIGWls_def gVarPresGWls_def by simp
lemma igAbsIPresIGWls_fromMOD[simp]:
"gAbsPresGWls MOD \<Longrightarrow> igAbsIPresIGWls (fromMOD MOD)"
unfolding igAbsIPresIGWls_def gAbsPresGWls_def by simp
lemma igOpIPresIGWls_fromMOD[simp]:
"gOpPresGWls MOD \<Longrightarrow> igOpIPresIGWls (fromMOD MOD)"
unfolding igOpIPresIGWls_def gOpPresGWls_def by simp
lemma igConsIPresIGWls_fromMOD[simp]:
"gConsPresGWls MOD \<Longrightarrow> igConsIPresIGWls (fromMOD MOD)"
unfolding igConsIPresIGWls_def gConsPresGWls_def by simp
lemmas igConsIPresIGWls_fromMOD_simps =
igVarIPresIGWls_fromMOD igAbsIPresIGWls_fromMOD
igOpIPresIGWls_fromMOD igConsIPresIGWls_fromMOD
-text{* Swap preserves the domains: *}
+text\<open>Swap preserves the domains:\<close>
lemma igSwapIPresIGWls_fromMOD[simp]:
"gSwapPresGWls MOD \<Longrightarrow> igSwapIPresIGWls (fromMOD MOD)"
unfolding igSwapIPresIGWls_def gSwapPresGWls_def by simp
lemma igSwapAbsIPresIGWlsAbs_fromMOD[simp]:
"gSwapAbsPresGWlsAbs MOD \<Longrightarrow> igSwapAbsIPresIGWlsAbs (fromMOD MOD)"
unfolding igSwapAbsIPresIGWlsAbs_def gSwapAbsPresGWlsAbs_def by simp
lemma igSwapAllIPresIGWlsAll_fromMOD[simp]:
"gSwapAllPresGWlsAll MOD \<Longrightarrow> igSwapAllIPresIGWlsAll (fromMOD MOD)"
unfolding igSwapAllIPresIGWlsAll_def gSwapAllPresGWlsAll_def by simp
lemmas igSwapAllIPresIGWlsAll_fromMOD_simps =
igSwapIPresIGWls_fromMOD igSwapAbsIPresIGWlsAbs_fromMOD igSwapAllIPresIGWlsAll_fromMOD
-text{* Subst preserves the domains: *}
+text\<open>Subst preserves the domains:\<close>
lemma igSubstIPresIGWls_fromMOD[simp]:
"gSubstPresGWls MOD \<Longrightarrow> igSubstIPresIGWls (fromMOD MOD)"
unfolding igSubstIPresIGWls_def gSubstPresGWls_def by simp
lemma igSubstAbsIPresIGWlsAbs_fromMOD[simp]:
"gSubstAbsPresGWlsAbs MOD \<Longrightarrow> igSubstAbsIPresIGWlsAbs (fromMOD MOD)"
unfolding igSubstAbsIPresIGWlsAbs_def gSubstAbsPresGWlsAbs_def by simp
lemma igSubstAllIPresIGWlsAll_fromMOD[simp]:
"gSubstAllPresGWlsAll MOD \<Longrightarrow> igSubstAllIPresIGWlsAll (fromMOD MOD)"
unfolding igSubstAllIPresIGWlsAll_def gSubstAllPresGWlsAll_def by simp
lemmas igSubstAllIPresIGWlsAll_fromMOD_simps =
igSubstIPresIGWls_fromMOD igSubstAbsIPresIGWlsAbs_fromMOD igSubstAllIPresIGWlsAll_fromMOD
-text{* The fresh clauses: *}
+text\<open>The fresh clauses:\<close>
lemma igFreshIGVar_fromMOD[simp]:
"gFreshGVar MOD \<Longrightarrow> igFreshIGVar (fromMOD MOD)"
unfolding igFreshIGVar_def gFreshGVar_def by simp
lemma igFreshIGAbs1_fromMOD[simp]:
"gFreshGAbs1 MOD \<Longrightarrow> igFreshIGAbs1 (fromMOD MOD)"
unfolding igFreshIGAbs1_def gFreshGAbs1_def by auto
lemma igFreshIGAbs2_fromMOD[simp]:
"gFreshGAbs2 MOD \<Longrightarrow> igFreshIGAbs2 (fromMOD MOD)"
unfolding igFreshIGAbs2_def gFreshGAbs2_def by auto
lemma igFreshIGOp_fromMOD[simp]:
"gFreshGOp MOD \<Longrightarrow> igFreshIGOp (fromMOD MOD)"
unfolding igFreshIGOp_def gFreshGOp_def by simp
lemma igFreshCls_fromMOD[simp]:
"gFreshCls MOD \<Longrightarrow> igFreshCls (fromMOD MOD)"
unfolding igFreshCls_def gFreshCls_def by simp
lemmas igFreshCls_fromMOD_simps =
igFreshIGVar_fromMOD igFreshIGAbs1_fromMOD igFreshIGAbs2_fromMOD
igFreshIGOp_fromMOD igFreshCls_fromMOD
-text{* The swap clauses *}
+text\<open>The swap clauses\<close>
lemma igSwapIGVar_fromMOD[simp]:
"gSwapGVar MOD \<Longrightarrow> igSwapIGVar (fromMOD MOD)"
unfolding igSwapIGVar_def gSwapGVar_def by simp
lemma igSwapIGAbs_fromMOD[simp]:
"gSwapGAbs MOD \<Longrightarrow> igSwapIGAbs (fromMOD MOD)"
unfolding igSwapIGAbs_def gSwapGAbs_def by auto
lemma igSwapIGOp_fromMOD[simp]:
"gSwapGOp MOD \<Longrightarrow> igSwapIGOp (fromMOD MOD)"
by (auto simp: igSwapIGOp_def gSwapGOp_def lift_lift2)
lemma igSwapCls_fromMOD[simp]:
"gSwapCls MOD \<Longrightarrow> igSwapCls (fromMOD MOD)"
unfolding igSwapCls_def gSwapCls_def by simp
lemmas igSwapCls_fromMOD_simps =
igSwapIGVar_fromMOD igSwapIGAbs_fromMOD
igSwapIGOp_fromMOD igSwapCls_fromMOD
-text{* The subst clauses *}
+text\<open>The subst clauses\<close>
lemma igSubstIGVar1_fromMOD[simp]:
"gSubstGVar1 MOD \<Longrightarrow> igSubstIGVar1 (fromMOD MOD)"
unfolding igSubstIGVar1_def gSubstGVar1_def by simp
lemma igSubstIGVar2_fromMOD[simp]:
"gSubstGVar2 MOD \<Longrightarrow> igSubstIGVar2 (fromMOD MOD)"
unfolding igSubstIGVar2_def gSubstGVar2_def by simp
lemma igSubstIGAbs_fromMOD[simp]:
"gSubstGAbs MOD \<Longrightarrow> igSubstIGAbs (fromMOD MOD)"
unfolding igSubstIGAbs_def gSubstGAbs_def by fastforce+
lemma igSubstIGOp_fromMOD[simp]:
"gSubstGOp MOD \<Longrightarrow> igSubstIGOp (fromMOD MOD)"
by(auto simp: igSubstIGOp_def gSubstGOp_def lift_lift2)
lemma igSubstCls_fromMOD[simp]:
"gSubstCls MOD \<Longrightarrow> igSubstCls (fromMOD MOD)"
unfolding igSubstCls_def gSubstCls_def by simp
lemmas igSubstCls_fromMOD_simps =
igSubstIGVar1_fromMOD igSubstIGVar2_fromMOD igSubstIGAbs_fromMOD
igSubstIGOp_fromMOD igSubstCls_fromMOD
-text{* Abstraction swapping congruence: *}
+text\<open>Abstraction swapping congruence:\<close>
lemma igAbsCongS_fromMOD[simp]:
assumes "gAbsCongS MOD"
shows "igAbsCongS (fromMOD MOD)"
using assms
unfolding igAbsCongS_def gAbsCongS_def
apply simp
apply clarify
by (intro conjI, erule wls_Abs_swap_cong) blast+
-text{* Abstraction renaming: *}
+text\<open>Abstraction renaming:\<close>
lemma igAbsRen_fromMOD[simp]:
"gAbsRen MOD \<Longrightarrow> igAbsRen (fromMOD MOD)"
unfolding igAbsRen_def gAbsRen_def vsubst_def by auto
-text{* Models: *}
+text\<open>Models:\<close>
lemma iwlsFSw_fromMOD[simp]:
"wlsFSw MOD \<Longrightarrow> iwlsFSw (fromMOD MOD)"
unfolding iwlsFSw_def wlsFSw_def by simp
lemma iwlsFSb_fromMOD[simp]:
"wlsFSb MOD \<Longrightarrow> iwlsFSb (fromMOD MOD)"
unfolding iwlsFSb_def wlsFSb_def by simp
lemma iwlsFSwSb_fromMOD[simp]:
"wlsFSwSb MOD \<Longrightarrow> iwlsFSwSb (fromMOD MOD)"
unfolding iwlsFSwSb_def wlsFSwSb_def by simp
lemma iwlsFSbSw_fromMOD[simp]:
"wlsFSbSw MOD \<Longrightarrow> iwlsFSbSw (fromMOD MOD)"
unfolding iwlsFSbSw_def wlsFSbSw_def by simp
lemmas iwlsModel_fromMOD_simps =
iwlsFSw_fromMOD iwlsFSb_fromMOD
iwlsFSwSb_fromMOD iwlsFSbSw_fromMOD
(******************************)
lemmas fromMOD_predicate_simps =
igWlsAllDisj_fromMOD_simps
igConsIPresIGWls_fromMOD_simps
igSwapAllIPresIGWlsAll_fromMOD_simps
igSubstAllIPresIGWlsAll_fromMOD_simps
igFreshCls_fromMOD_simps
igSwapCls_fromMOD_simps
igSubstCls_fromMOD_simps
igAbsCongS_fromMOD
igAbsRen_fromMOD
iwlsModel_fromMOD_simps
lemmas fromMOD_simps =
fromMOD_basic_simps
input_fromMOD_simps
binput_fromMOD_simps
fromMOD_predicate_simps
-subsection {* The recursion-iteration ``identity trick" *}
+subsection \<open>The recursion-iteration ``identity trick"\<close>
-text {* Here we show that any construct-preserving map from terms to ``fromMOD MOD"
+text \<open>Here we show that any construct-preserving map from terms to ``fromMOD MOD"
is the identity on its first projection -- this is the main trick when
-reducing recursion to iteration. *}
+reducing recursion to iteration.\<close>
lemma ipresCons_fromMOD_fst:
assumes "ipresCons h hA (fromMOD MOD)"
shows "(wls s X \<longrightarrow> fst (h X) = X) \<and> (wlsAbs (us,s') A \<longrightarrow> fst (hA A) = A)"
proof(induction rule: wls_rawInduct)
next
case (Op delta inp binp)
hence "lift (fst \<circ> h) inp = inp \<and> lift (fst \<circ> hA) binp = binp"
by (simp add: lift_def fun_eq_iff liftAll2_def
wlsInp_iff wlsBinp_iff sameDom_def split: option.splits)
(metis not_Some_eq old.prod.exhaust)
then show ?case
using assms Op by (auto simp: ipresCons_def ipresOp_def lift_comp)
qed(insert assms, auto simp: ipresVar_def ipresCons_def ipresAbs_def)
lemma ipresCons_fromMOD_fst_simps[simp]:
"\<lbrakk>ipresCons h hA (fromMOD MOD); wls s X\<rbrakk>
\<Longrightarrow> fst (h X) = X"
(* *)
"\<lbrakk>ipresCons h hA (fromMOD MOD); wlsAbs (us,s') A\<rbrakk>
\<Longrightarrow> fst (hA A) = A"
using ipresCons_fromMOD_fst by blast+
lemma ipresCons_fromMOD_fst_inp[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> wlsInp delta inp \<Longrightarrow> lift (fst o h) inp = inp"
by (force simp add: lift_def fun_eq_iff liftAll2_def
wlsInp_iff sameDom_def split: option.splits)
lemma ipresCons_fromMOD_fst_binp[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> wlsBinp delta binp \<Longrightarrow> lift (fst o hA) binp = binp"
by (force simp add: lift_def fun_eq_iff liftAll2_def
wlsBinp_iff sameDom_def split: option.splits)
lemmas ipresCons_fromMOD_fst_all_simps =
ipresCons_fromMOD_fst_simps ipresCons_fromMOD_fst_inp ipresCons_fromMOD_fst_binp
-subsection {* From iteration morphisms to morphisms *}
+subsection \<open>From iteration morphisms to morphisms\<close>
-text{* The transition map: *}
+text\<open>The transition map:\<close>
definition fromIMor ::
"(('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)term \<times> 'gTerm)
\<Rightarrow>
(('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'gTerm)"
where "fromIMor h \<equiv> snd o h"
definition fromIMorAbs ::
"(('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)abs \<times> 'gAbs)
\<Rightarrow>
(('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow> 'gAbs)"
where "fromIMorAbs hA \<equiv> snd o hA"
-text{* Basic simplification rules: *}
+text\<open>Basic simplification rules:\<close>
lemma fromIMor[simp]: "fromIMor h X' = snd (h X')"
unfolding fromIMor_def by simp
lemma fromIMorAbs[simp]: "fromIMorAbs hA A' = snd (hA A')"
unfolding fromIMorAbs_def by simp
lemma fromIMor_snd_inp[simp]:
"wlsInp delta inp \<Longrightarrow> lift (fromIMor h) inp = lift (snd o h) inp"
by (auto simp: lift_def split: option.splits)
lemma fromIMorAbs_snd_binp[simp]:
"wlsBinp delta binp \<Longrightarrow> lift (fromIMorAbs hA) binp = lift (snd o hA) binp"
by (auto simp: lift_def split: option.splits)
lemmas fromIMor_basic_simps =
fromIMor fromIMorAbs fromIMor_snd_inp fromIMorAbs_snd_binp
-text{* Predicate simplification rules *}
+text\<open>Predicate simplification rules\<close>
-text{* Domain preservation *}
+text\<open>Domain preservation\<close>
lemma presWls_fromIMor[simp]:
"ipresWls h (fromMOD MOD) \<Longrightarrow> presWls (fromIMor h) MOD"
unfolding ipresWls_def presWls_def by simp
lemma presWlsAbs_fromIMorAbs[simp]:
"ipresWlsAbs hA (fromMOD MOD) \<Longrightarrow> presWlsAbs (fromIMorAbs hA) MOD"
unfolding ipresWlsAbs_def presWlsAbs_def by simp
lemma presWlsAll_fromIMorAll[simp]:
"ipresWlsAll h hA (fromMOD MOD) \<Longrightarrow> presWlsAll (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresWlsAll_def presWlsAll_def by simp
lemmas presWlsAll_fromIMorAll_simps =
presWls_fromIMor presWlsAbs_fromIMorAbs presWlsAll_fromIMorAll
-text{* Preservation of the constructs *}
+text\<open>Preservation of the constructs\<close>
lemma presVar_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> presVar (fromIMor h) MOD"
unfolding ipresCons_def ipresVar_def presVar_def by simp
lemma presAbs_fromIMor[simp]:
assumes "ipresCons h hA (fromMOD MOD)"
shows "presAbs (fromIMor h) (fromIMorAbs hA) MOD"
using assms unfolding ipresCons_def ipresAbs_def presAbs_def
using assms by fastforce
lemma presOp_fromIMor[simp]:
assumes "ipresCons h hA (fromMOD MOD)"
shows "presOp (fromIMor h) (fromIMorAbs hA) MOD"
using assms unfolding ipresCons_def ipresOp_def presOp_def
using assms by (auto simp: lift_comp)
lemma presCons_fromIMor[simp]:
assumes "ipresCons h hA (fromMOD MOD)"
shows "presCons (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresCons_def presCons_def using assms by simp
lemmas presCons_fromIMor_simps =
presVar_fromIMor presAbs_fromIMor presOp_fromIMor presCons_fromIMor
-text{* Preservation of freshness *}
+text\<open>Preservation of freshness\<close>
lemma presFresh_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresFresh h (fromMOD MOD)
\<Longrightarrow> presFresh (fromIMor h) MOD"
unfolding ipresFresh_def presFresh_def by simp
lemma presFreshAbs_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresFreshAbs hA (fromMOD MOD)
\<Longrightarrow> presFreshAbs (fromIMorAbs hA) MOD"
unfolding ipresFreshAbs_def presFreshAbs_def by simp
lemma presFreshAll_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresFreshAll h hA (fromMOD MOD)
\<Longrightarrow> presFreshAll (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresFreshAll_def presFreshAll_def by simp
lemmas presFreshAll_fromIMor_simps =
presFresh_fromIMor presFreshAbs_fromIMor presFreshAll_fromIMor
-text{* Preservation of swap *}
+text\<open>Preservation of swap\<close>
lemma presSwap_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSwap h (fromMOD MOD)
\<Longrightarrow> presSwap (fromIMor h) MOD"
unfolding ipresSwap_def presSwap_def by simp
lemma presSwapAbs_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSwapAbs hA (fromMOD MOD)
\<Longrightarrow> presSwapAbs (fromIMorAbs hA) MOD"
unfolding ipresSwapAbs_def presSwapAbs_def by simp
lemma presSwapAll_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSwapAll h hA (fromMOD MOD)
\<Longrightarrow> presSwapAll (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresSwapAll_def presSwapAll_def by simp
lemmas presSwapAll_fromIMor_simps =
presSwap_fromIMor presSwapAbs_fromIMor presSwapAll_fromIMor
-text{* Preservation of subst *}
+text\<open>Preservation of subst\<close>
lemma presSubst_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSubst h (fromMOD MOD)
\<Longrightarrow> presSubst (fromIMor h) MOD"
unfolding ipresSubst_def presSubst_def by auto
lemma presSubstAbs_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSubstAbs h hA (fromMOD MOD)
\<Longrightarrow> presSubstAbs (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresSubstAbs_def presSubstAbs_def by auto
lemma presSubstAll_fromIMor[simp]:
"ipresCons h hA (fromMOD MOD) \<Longrightarrow> ipresSubstAll h hA (fromMOD MOD)
\<Longrightarrow> presSubstAll (fromIMor h) (fromIMorAbs hA) MOD"
unfolding ipresSubstAll_def presSubstAll_def by simp
lemmas presSubstAll_fromIMor_simps =
presSubst_fromIMor presSubstAbs_fromIMor presSubstAll_fromIMor
-text{* Morphisms *}
+text\<open>Morphisms\<close>
lemma fromIMor_termFSwMorph[simp]:
"termFSwImorph h hA (fromMOD MOD) \<Longrightarrow> termFSwMorph (fromIMor h) (fromIMorAbs hA) MOD"
unfolding termFSwImorph_def termFSwMorph_def by simp
lemma fromIMor_termFSbMorph[simp]:
"termFSbImorph h hA (fromMOD MOD) \<Longrightarrow> termFSbMorph (fromIMor h) (fromIMorAbs hA) MOD"
unfolding termFSbImorph_def termFSbMorph_def by simp
lemma fromIMor_termFSwSbMorph[simp]:
assumes "termFSwSbImorph h hA (fromMOD MOD)"
shows "termFSwSbMorph (fromIMor h) (fromIMorAbs hA) MOD"
using assms unfolding termFSwSbImorph_defs1
using assms unfolding termFSwSbImorph_def termFSwSbMorph_def by simp
lemmas mor_fromIMor_simps =
fromIMor_termFSwMorph fromIMor_termFSbMorph fromIMor_termFSwSbMorph
(********************************)
lemmas fromIMor_predicate_simps =
presCons_fromIMor_simps
presFreshAll_fromIMor_simps
presSwapAll_fromIMor_simps
presSubstAll_fromIMor_simps
mor_fromIMor_simps
lemmas fromIMor_simps =
fromIMor_basic_simps fromIMor_predicate_simps
-subsection {* The recursion theorem *}
+subsection \<open>The recursion theorem\<close>
-text{* The recursion maps: *}
+text\<open>The recursion maps:\<close>
definition rec where "rec MOD \<equiv> fromIMor (iter (fromMOD MOD))"
definition recAbs where "recAbs MOD \<equiv> fromIMorAbs (iterAbs (fromMOD MOD))"
-text{* Existence: *}
+text\<open>Existence:\<close>
theorem wlsFSw_recAll_termFSwMorph:
"wlsFSw MOD \<Longrightarrow> termFSwMorph (rec MOD) (recAbs MOD) MOD"
by (simp add: rec_def recAbs_def iwlsFSw_iterAll_termFSwImorph)
theorem wlsFSb_recAll_termFSbMorph:
"wlsFSb MOD \<Longrightarrow> termFSbMorph (rec MOD) (recAbs MOD) MOD"
by (simp add: rec_def recAbs_def iwlsFSb_iterAll_termFSbImorph)
theorem wlsFSwSb_recAll_termFSwSbMorph:
"wlsFSwSb MOD \<Longrightarrow> termFSwSbMorph (rec MOD) (recAbs MOD) MOD"
by (simp add: rec_def recAbs_def iwlsFSwSb_iterAll_termFSwSbImorph)
theorem wlsFSbSw_recAll_termFSwSbMorph:
"wlsFSbSw MOD \<Longrightarrow> termFSwSbMorph (rec MOD) (recAbs MOD) MOD"
by (simp add: rec_def recAbs_def iwlsFSbSw_iterAll_termFSwSbImorph)
-text{* Uniqueness: *}
+text\<open>Uniqueness:\<close>
lemma presCons_unique:
assumes "presCons f fA MOD" and "presCons g gA MOD"
shows "(wls s X \<longrightarrow> f X = g X) \<and> (wlsAbs (us,s') A \<longrightarrow> fA A = gA A)"
proof(induction rule: wls_rawInduct)
case (Op delta inp binp)
hence "lift f inp = lift g inp \<and> lift fA binp = lift gA binp"
apply(simp add: lift_def wlsInp_iff wlsBinp_iff sameDom_def liftAll2_def fun_eq_iff split: option.splits)
by (metis not_Some_eq old.prod.exhaust)
then show ?case using assms Op unfolding presCons_def presOp_def by simp
qed(insert assms, auto simp: presVar_def presCons_def presAbs_def )
theorem wlsFSw_recAll_unique_presCons:
assumes "wlsFSw MOD" and "presCons h hA MOD"
shows "(wls s X \<longrightarrow> h X = rec MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = recAbs MOD A)"
using assms wlsFSw_recAll_termFSwMorph
by (intro presCons_unique) (auto simp: termFSwMorph_def)
theorem wlsFSb_recAll_unique_presCons:
assumes "wlsFSb MOD" and "presCons h hA MOD"
shows "(wls s X \<longrightarrow> h X = rec MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = recAbs MOD A)"
using assms wlsFSb_recAll_termFSbMorph
by (intro presCons_unique) (auto simp: termFSbMorph_def)
theorem wlsFSwSb_recAll_unique_presCons:
assumes "wlsFSwSb MOD" and "presCons h hA MOD"
shows "(wls s X \<longrightarrow> h X = rec MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = recAbs MOD A)"
using assms wlsFSw_recAll_unique_presCons unfolding wlsFSwSb_def by blast
theorem wlsFSbSw_recAll_unique_presCons:
assumes "wlsFSbSw MOD" and "presCons h hA MOD"
shows "(wls s X \<longrightarrow> h X = rec MOD X) \<and>
(wlsAbs (us,s') A \<longrightarrow> hA A = recAbs MOD A)"
using assms wlsFSb_recAll_unique_presCons unfolding wlsFSbSw_def by blast
-subsection{* Models that are even ``closer" to the term model *}
+subsection\<open>Models that are even ``closer" to the term model\<close>
-text{* We describe various conditions (later referred to as ``extra clauses"
+text\<open>We describe various conditions (later referred to as ``extra clauses"
or ``extra conditions")
that, when satisfied by models,
yield the recursive maps
(1) freshness-preserving and/or (2) injective and/or (3) surjective, thus bringing the
considered models ``closer" to (being isomorphic to) the term model.
The extreme case, when all of (1)-(3) above are ensured, means indeed isomorphism to
-the term model -- this is in fact an abstract characterization of the term model. *}
+the term model -- this is in fact an abstract characterization of the term model.\<close>
-subsubsection {* Relevant predicates on models *}
+subsubsection \<open>Relevant predicates on models\<close>
-text{* The fresh clauses reversed *}
+text\<open>The fresh clauses reversed\<close>
definition gFreshGVarRev where
"gFreshGVarRev MOD \<equiv> \<forall> xs y x.
gFresh MOD xs y (Var xs x) (gVar MOD xs x) \<longrightarrow> y \<noteq> x"
definition gFreshGAbsRev where
"gFreshGAbsRev MOD \<equiv> \<forall> ys y xs x s X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
gFreshAbs MOD ys y (Abs xs x X') (gAbs MOD xs x X' X) \<longrightarrow>
(ys = xs \<and> y = x) \<or> gFresh MOD ys y X' X"
definition gFreshGOpRev where
"gFreshGOpRev MOD \<equiv> \<forall> ys y delta inp' inp binp' binp.
wlsInp delta inp' \<and> gWlsInp MOD delta inp \<and> wlsBinp delta binp' \<and> gWlsBinp MOD delta binp \<longrightarrow>
gFresh MOD ys y (Op delta inp' binp') (gOp MOD delta inp' inp binp' binp) \<longrightarrow>
gFreshInp MOD ys y inp' inp \<and> gFreshBinp MOD ys y binp' binp"
definition gFreshClsRev where
"gFreshClsRev MOD \<equiv> gFreshGVarRev MOD \<and> gFreshGAbsRev MOD \<and> gFreshGOpRev MOD"
lemmas gFreshClsRev_defs = gFreshClsRev_def
gFreshGVarRev_def gFreshGAbsRev_def gFreshGOpRev_def
-text{* Injectiveness of the construct operators *}
+text\<open>Injectiveness of the construct operators\<close>
definition gVarInj where
"gVarInj MOD \<equiv> \<forall> xs x y. gVar MOD xs x = gVar MOD xs y \<longrightarrow> x = y"
definition gAbsInj where
"gAbsInj MOD \<equiv> \<forall> xs s x X' X X1' X1.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<and> wls s X1' \<and> gWls MOD s X1 \<and>
gAbs MOD xs x X' X = gAbs MOD xs x X1' X1
\<longrightarrow>
X = X1"
definition gOpInj where
"gOpInj MOD \<equiv> \<forall> delta delta1 inp' binp' inp binp inp1' binp1' inp1 binp1.
wlsInp delta inp' \<and> wlsBinp delta binp' \<and> gWlsInp MOD delta inp \<and> gWlsBinp MOD delta binp \<and>
wlsInp delta1 inp1' \<and> wlsBinp delta1 binp1' \<and> gWlsInp MOD delta1 inp1 \<and> gWlsBinp MOD delta1 binp1 \<and>
stOf delta = stOf delta1 \<and>
gOp MOD delta inp' inp binp' binp = gOp MOD delta1 inp1' inp1 binp1' binp1
\<longrightarrow>
delta = delta1 \<and> inp = inp1 \<and> binp = binp1"
definition gVarGOpInj where
"gVarGOpInj MOD \<equiv> \<forall> xs x delta inp' binp' inp binp.
wlsInp delta inp' \<and> wlsBinp delta binp' \<and> gWlsInp MOD delta inp \<and> gWlsBinp MOD delta binp \<and>
asSort xs = stOf delta
\<longrightarrow>
gVar MOD xs x \<noteq> gOp MOD delta inp' inp binp' binp"
definition gConsInj where
"gConsInj MOD \<equiv> gVarInj MOD \<and> gAbsInj MOD \<and> gOpInj MOD \<and> gVarGOpInj MOD"
lemmas gConsInj_defs = gConsInj_def
gVarInj_def gAbsInj_def gOpInj_def gVarGOpInj_def
-text{* Abstraction renaming for swapping *}
+text\<open>Abstraction renaming for swapping\<close>
definition gAbsRenS where
"gAbsRenS MOD \<equiv> \<forall> xs y x s X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<longrightarrow>
fresh xs y X' \<and> gFresh MOD xs y X' X \<longrightarrow>
gAbs MOD xs y (X' #[y \<and> x]_xs) (gSwap MOD xs y x X' X) =
gAbs MOD xs x X' X"
-text{* Indifference to the general-recursive argument *}
+text\<open>Indifference to the general-recursive argument\<close>
-text{* . This ``indifference" property says that the construct operators
+text\<open>. This ``indifference" property says that the construct operators
from the model only depend on
the generalized item (i.e., generalized term or abstraction) argument,
and {\em not} on the ``item" (i.e., concrete term or abstraction) argument.
In other words, the model constructs correspond to {\em iterative clauses},
-and not to the more general notion of ``general-recursive" clause. *}
+and not to the more general notion of ``general-recursive" clause.\<close>
definition gAbsIndif where
"gAbsIndif MOD \<equiv> \<forall> xs s x X1' X2' X.
isInBar (xs,s) \<and> wls s X1' \<and> wls s X2' \<and> gWls MOD s X \<longrightarrow>
gAbs MOD xs x X1' X = gAbs MOD xs x X2' X"
definition gOpIndif where
"gOpIndif MOD \<equiv> \<forall> delta inp1' inp2' inp binp1' binp2' binp.
wlsInp delta inp1' \<and> wlsBinp delta binp1' \<and> wlsInp delta inp2' \<and> wlsBinp delta binp2' \<and>
gWlsInp MOD delta inp \<and> gWlsBinp MOD delta binp
\<longrightarrow>
gOp MOD delta inp1' inp binp1' binp = gOp MOD delta inp2' inp binp2' binp"
definition gConsIndif where
"gConsIndif MOD \<equiv> gOpIndif MOD \<and> gAbsIndif MOD"
lemmas gConsIndif_defs = gConsIndif_def gAbsIndif_def gOpIndif_def
-text{* Inductiveness *}
+text\<open>Inductiveness\<close>
-text{* . Inductiveness of a model means the satisfaction of a minimal inductive
+text\<open>. Inductiveness of a model means the satisfaction of a minimal inductive
principle (``minimal" in the sense that no fancy swapping or freshness
-induction-friendly conditions are involved). *}
+induction-friendly conditions are involved).\<close>
definition gInduct where
"gInduct MOD \<equiv> \<forall> phi phiAbs s X us s' A.
(
(\<forall> xs x. phi (asSort xs) (gVar MOD xs x))
\<and>
(\<forall> delta inp' inp binp' binp.
wlsInp delta inp' \<and> wlsBinp delta binp' \<and> gWlsInp MOD delta inp \<and> gWlsBinp MOD delta binp \<and>
liftAll2 phi (arOf delta) inp \<and> liftAll2 phiAbs (barOf delta) binp
\<longrightarrow> phi (stOf delta) (gOp MOD delta inp' inp binp' binp))
\<and>
(\<forall> xs s x X' X.
isInBar (xs,s) \<and> wls s X' \<and> gWls MOD s X \<and>
phi s X
\<longrightarrow> phiAbs (xs,s) (gAbs MOD xs x X' X))
)
\<longrightarrow>
(gWls MOD s X \<longrightarrow> phi s X) \<and>
(gWlsAbs MOD (us,s') A \<longrightarrow> phiAbs (us,s') A)"
lemma gInduct_elim:
assumes "gInduct MOD" and
Var: "\<And> xs x. phi (asSort xs) (gVar MOD xs x)" and
Op:
"\<And> delta inp' inp binp' binp.
\<lbrakk>wlsInp delta inp'; wlsBinp delta binp'; gWlsInp MOD delta inp; gWlsBinp MOD delta binp;
liftAll2 phi (arOf delta) inp; liftAll2 phiAbs (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (gOp MOD delta inp' inp binp' binp)" and
Abs:
"\<And> xs s x X' X.
\<lbrakk>isInBar (xs,s); wls s X'; gWls MOD s X; phi s X\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (gAbs MOD xs x X' X)"
shows
"(gWls MOD s X \<longrightarrow> phi s X) \<and>
(gWlsAbs MOD (us,s') A \<longrightarrow> phiAbs (us,s') A)"
using assms unfolding gInduct_def
apply(elim allE[of _ phi] allE[of _ phiAbs] allE[of _ s] allE[of _ X])
apply(elim allE[of _ us] allE[of _ s'] allE[of _ A])
by blast
-subsubsection {* Relevant predicates on maps from the term model *}
+subsubsection \<open>Relevant predicates on maps from the term model\<close>
-text{* Reflection of freshness *}
+text\<open>Reflection of freshness\<close>
definition reflFresh where
"reflFresh h MOD \<equiv> \<forall> ys y s X.
wls s X \<longrightarrow>
gFresh MOD ys y X (h X) \<longrightarrow> fresh ys y X"
definition reflFreshAbs where
"reflFreshAbs hA MOD \<equiv> \<forall> ys y us s A.
wlsAbs (us,s) A \<longrightarrow>
gFreshAbs MOD ys y A (hA A) \<longrightarrow> freshAbs ys y A"
definition reflFreshAll where
"reflFreshAll h hA MOD \<equiv> reflFresh h MOD \<and> reflFreshAbs hA MOD"
lemmas reflFreshAll_defs = reflFreshAll_def
reflFresh_def reflFreshAbs_def
-text{* Injectiveness *}
+text\<open>Injectiveness\<close>
definition isInj where
"isInj h \<equiv> \<forall> s X Y.
wls s X \<and> wls s Y \<longrightarrow>
h X = h Y \<longrightarrow> X = Y"
definition isInjAbs where
"isInjAbs hA \<equiv> \<forall> us s A B.
wlsAbs (us,s) A \<and> wlsAbs (us,s) B \<longrightarrow>
hA A = hA B \<longrightarrow> A = B"
definition isInjAll where
"isInjAll h hA \<equiv> isInj h \<and> isInjAbs hA"
lemmas isInjAll_defs = isInjAll_def
isInj_def isInjAbs_def
-text{* Surjectiveness *}
+text\<open>Surjectiveness\<close>
definition isSurj where
"isSurj h MOD \<equiv> \<forall> s X.
gWls MOD s X \<longrightarrow>
(\<exists> X'. wls s X' \<and> h X' = X)"
definition isSurjAbs where
"isSurjAbs hA MOD \<equiv> \<forall> us s A.
gWlsAbs MOD (us,s) A \<longrightarrow>
(\<exists> A'. wlsAbs (us,s) A' \<and> hA A' = A)"
definition isSurjAll where
"isSurjAll h hA MOD \<equiv> isSurj h MOD \<and> isSurjAbs hA MOD"
lemmas isSurjAll_defs = isSurjAll_def
isSurj_def isSurjAbs_def
-subsubsection{* Criterion for the reflection of freshness *}
+subsubsection\<open>Criterion for the reflection of freshness\<close>
-text{* First an auxiliary fact, independent of the type of model: *}
+text\<open>First an auxiliary fact, independent of the type of model:\<close>
lemma gFreshClsRev_recAll_reflFreshAll:
assumes pWls: "presWlsAll (rec MOD) (recAbs MOD) MOD"
and pCons: "presCons (rec MOD) (recAbs MOD) MOD"
and pFresh: "presFreshAll (rec MOD) (recAbs MOD) MOD"
and **: "gFreshClsRev MOD"
shows "reflFreshAll (rec MOD) (recAbs MOD) MOD"
proof-
let ?h = "rec MOD" let ?hA = "recAbs MOD"
have pWlsInps[simp]:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> gWlsInp MOD delta (lift ?h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> gWlsBinp MOD delta (lift ?hA binp)"
using pWls presWls_wlsInp presWls_wlsBinp unfolding presWlsAll_def by auto
{fix ys y s X us s' A
have
"(wls s X \<longrightarrow> gFresh MOD ys y X (rec MOD X) \<longrightarrow> fresh ys y X) \<and>
(wlsAbs (us,s') A \<longrightarrow> gFreshAbs MOD ys y A (recAbs MOD A) \<longrightarrow> freshAbs ys y A)"
proof(induction rule: wls_induct)
case (Var xs x)
then show ?case using assms
by (fastforce simp: presWlsAll_defs presCons_defs gFreshClsRev_def gFreshGVarRev_def)
next
case (Op delta inp binp)
show ?case proof safe
let ?ar = "arOf delta" let ?bar = "barOf delta" let ?st = "stOf delta"
let ?linp = "lift ?h inp" let ?lbinp = "lift ?hA binp"
assume "gFresh MOD ys y (Op delta inp binp) (rec MOD (Op delta inp binp))"
hence "gFresh MOD ys y (Op delta inp binp) (gOp MOD delta inp ?linp binp ?lbinp)"
using assms Op by (simp add: presCons_def presOp_def)
hence "gFreshInp MOD ys y inp ?linp \<and> gFreshBinp MOD ys y binp ?lbinp"
using Op ** by (force simp: gFreshClsRev_def gFreshGOpRev_def)
with Op have freshInp: "freshInp ys y inp \<and> freshBinp ys y binp"
by (simp add: freshInp_def freshBinp_def liftAll_def gFreshInp_def gFreshBinp_def liftAll2_def lift_def
sameDom_def wlsInp_iff wlsBinp_iff split: option.splits) (metis eq_snd_iff not_Some_eq)
thus "fresh ys y (Op delta inp binp)" using Op by auto
qed
next
case (Abs s xs x X)
show ?case proof safe
have hX: "gWls MOD s (?h X)" using Abs pWls unfolding presWlsAll_defs by simp
assume "gFreshAbs MOD ys y (Abs xs x X) (recAbs MOD (Abs xs x X))"
hence "gFreshAbs MOD ys y (Abs xs x X) (gAbs MOD xs x X (rec MOD X))"
using Abs by (metis pCons presAbs_def presCons_def)
moreover have "?hA (Abs xs x X) = gAbs MOD xs x X (?h X)"
using Abs pCons unfolding presCons_defs by blast
ultimately have 1: "gFreshAbs MOD ys y (Abs xs x X) (gAbs MOD xs x X (?h X))" by simp
show "freshAbs ys y (Abs xs x X)"
using assms hX Abs ** unfolding gFreshClsRev_def gFreshGAbsRev_def using 1 by fastforce
qed
qed
}
thus ?thesis unfolding reflFreshAll_defs by blast
qed
-text{* For fresh-swap models *}
+text\<open>For fresh-swap models\<close>
theorem wlsFSw_recAll_reflFreshAll:
"wlsFSw MOD \<Longrightarrow> gFreshClsRev MOD \<Longrightarrow> reflFreshAll (rec MOD) (recAbs MOD) MOD"
using wlsFSw_recAll_termFSwMorph
by (auto simp: termFSwMorph_def intro: gFreshClsRev_recAll_reflFreshAll)
-text{* For fresh-subst models *}
+text\<open>For fresh-subst models\<close>
theorem wlsFSb_recAll_reflFreshAll:
"wlsFSb MOD \<Longrightarrow> gFreshClsRev MOD \<Longrightarrow> reflFreshAll (rec MOD) (recAbs MOD) MOD"
using wlsFSb_recAll_termFSbMorph
by (auto simp: termFSbMorph_def intro: gFreshClsRev_recAll_reflFreshAll)
(* Note: Here and below: No need for corresponding results for FSwSb and FSbSw models, as they
would follow at once from the above. *)
-subsubsection{* Criterion for the injectiveness of the recursive map *}
+subsubsection\<open>Criterion for the injectiveness of the recursive map\<close>
-text{* For fresh-swap models *}
+text\<open>For fresh-swap models\<close>
theorem wlsFSw_recAll_isInjAll:
assumes *: "wlsFSw MOD" "gAbsRenS MOD" and **: "gConsInj MOD"
shows "isInjAll (rec MOD) (recAbs MOD)"
proof-
let ?h = "rec MOD" let ?hA = "recAbs MOD"
have 1: "termFSwMorph ?h ?hA MOD" using * wlsFSw_recAll_termFSwMorph by auto
hence pWls: "presWlsAll ?h ?hA MOD"
and pCons: "presCons ?h ?hA MOD"
and pFresh: "presFreshAll ?h ?hA MOD"
and pSwap: "presSwapAll ?h ?hA MOD" unfolding termFSwMorph_def by auto
hence pWlsInps[simp]:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> gWlsInp MOD delta (lift ?h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> gWlsBinp MOD delta (lift ?hA binp)"
using presWls_wlsInp presWls_wlsBinp unfolding presWlsAll_def by auto
{fix s X us s' A
have
"(wls s X \<longrightarrow> (\<forall> Y. wls s Y \<and> rec MOD X = rec MOD Y \<longrightarrow> X = Y)) \<and>
(wlsAbs (us,s') A \<longrightarrow> (\<forall> B. wlsAbs (us,s') B \<and> recAbs MOD A = recAbs MOD B \<longrightarrow> A = B))"
proof (induction rule: wls_induct)
case (Var xs x)
show ?case proof clarify
fix Y
assume eq: "rec MOD (Var xs x) = rec MOD Y" and Y: "wls (asSort xs) Y"
thus "Var xs x = Y"
proof-
{fix ys y assume Y_def: "Y = Var ys y" and "asSort ys = asSort xs"
hence ys_def: "ys = xs" by simp
have rec_y_def: "rec MOD (Var ys y) = gVar MOD ys y"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq ** 1 unfolding Y_def rec_y_def gConsInj_def gVarInj_def
unfolding ys_def by (simp add: termFSwMorph_defs)
}
moreover
{fix delta1 inp1 binp1 assume inp1s: "wlsInp delta1 inp1" "wlsBinp delta1 binp1"
and Y_def: "Y = Op delta1 inp1 binp1" and st: "stOf delta1 = asSort xs"
hence rec_Op_def:
"rec MOD (Op delta1 inp1 binp1) =
gOp MOD delta1 inp1 (lift ?h inp1) binp1 (lift ?hA binp1)"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq ** unfolding Y_def rec_Op_def gConsInj_def gVarGOpInj_def
using inp1s st 1 by (simp add: termFSwMorph_defs)
}
ultimately show ?thesis using wls_nchotomy[of "asSort xs" Y] Y by blast
qed
qed
next
case (Op delta inp binp)
show ?case proof clarify
fix Y assume Y: "wls (stOf delta) Y"
and "rec MOD (Op delta inp binp) = rec MOD Y"
hence eq: "gOp MOD delta inp (lift ?h inp) binp (lift ?hA binp) = ?h Y"
using 1 Op by (simp add: termFSwMorph_defs)
show "Op delta inp binp = Y"
proof-
{fix ys y assume Y_def: "Y = Var ys y" and st: "asSort ys = stOf delta"
have rec_y_def: "rec MOD (Var ys y) = gVar MOD ys y"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq[THEN sym] ** unfolding Y_def rec_y_def gConsInj_def gVarGOpInj_def
using Op st by simp
}
moreover
{fix delta1 inp1 binp1 assume inp1s: "wlsInp delta1 inp1" "wlsBinp delta1 binp1"
and Y_def: "Y = Op delta1 inp1 binp1" and st: "stOf delta1 = stOf delta"
hence rec_Op_def:
"rec MOD (Op delta1 inp1 binp1) =
gOp MOD delta1 inp1 (lift ?h inp1) binp1 (lift ?hA binp1)"
using pCons unfolding presCons_defs by simp
have 0: "delta = delta1 \<and> lift ?h inp = lift ?h inp1 \<and> lift ?hA binp = lift ?hA binp1"
using eq ** unfolding Y_def rec_Op_def gConsInj_def gOpInj_def
using Op inp1s st apply clarify
apply(erule allE[of _ delta]) apply(erule allE[of _ delta1]) by force
hence delta1_def: "delta1 = delta" by simp
have 1: "inp = inp1"
proof(rule ext)
fix i
show "inp i = inp1 i"
proof(cases "inp i")
case None
hence "lift ?h inp i = None" by(simp add: lift_None)
hence "lift ?h inp1 i = None" using 0 by simp
thus ?thesis unfolding None by(simp add: lift_None)
next
case (Some X)
hence "lift ?h inp i = Some (?h X)" unfolding lift_def by simp
hence "lift ?h inp1 i = Some (?h X)" using 0 by simp
then obtain Y where inp1_i: "inp1 i = Some Y" and hXY: "?h X = ?h Y"
unfolding lift_def by(cases "inp1 i") auto
then obtain s where ar_i: "arOf delta i = Some s"
using inp1s unfolding delta1_def wlsInp_iff sameDom_def
by (cases "arOf delta i") auto
hence Y: "wls s Y"
using inp1s inp1_i unfolding delta1_def wlsInp_iff liftAll2_def by auto
thus ?thesis
unfolding Some inp1_i using ar_i Some hXY Op.IH unfolding liftAll2_def by auto
qed
qed
have 2: "binp = binp1"
proof(rule ext)
fix i
show "binp i = binp1 i"
proof(cases "binp i")
case None
hence "lift ?hA binp i = None" by(simp add: lift_None)
hence "lift ?hA binp1 i = None" using 0 by simp
thus ?thesis unfolding None by (simp add: lift_None)
next
case (Some A)
hence "lift ?hA binp i = Some (?hA A)" unfolding lift_def by simp
hence "lift ?hA binp1 i = Some (?hA A)" using 0 by simp
then obtain B where binp1_i: "binp1 i = Some B" and hAB: "?hA A = ?hA B"
unfolding lift_def by (cases "binp1 i") auto
then obtain us s where bar_i: "barOf delta i = Some (us,s)"
using inp1s unfolding delta1_def wlsBinp_iff sameDom_def
by(cases "barOf delta i") auto
hence B: "wlsAbs (us,s) B"
using inp1s binp1_i unfolding delta1_def wlsBinp_iff liftAll2_def by auto
thus ?thesis unfolding Some binp1_i
using bar_i Some hAB Op.IH unfolding liftAll2_def by fastforce
qed
qed
have ?thesis unfolding Y_def delta1_def 1 2 by simp
}
ultimately show ?thesis using wls_nchotomy[of "stOf delta" Y] Y by blast
qed
qed
next
case (Abs s xs x X)
show ?case proof clarify
fix B
assume B: "wlsAbs (xs,s) B" and "recAbs MOD (Abs xs x X) = recAbs MOD B"
hence eq: "gAbs MOD xs x X (rec MOD X) = ?hA B" using 1 Abs by (simp add: termFSwMorph_defs)
hence hX: "gWls MOD s (?h X)" using pWls Abs unfolding presWlsAll_defs by simp
show "Abs xs x X = B"
proof-
let ?P = "ParS
(\<lambda> xs'. [])
(\<lambda> s'. if s' = s then [X] else [])
(\<lambda> us_s. [])
[]"
have P: "wlsPar ?P" using Abs unfolding wlsPar_def by simp
{fix y Y assume Y: "wls s Y" and B_def: "B = Abs xs y Y"
hence hY: "gWls MOD s (?h Y)" using pWls unfolding presWlsAll_defs by simp
let ?Xsw = "X #[y \<and> x]_xs" let ?hXsw = "gSwap MOD xs y x X (?h X)"
have hXsw: "gWls MOD s ?hXsw"
using Abs hX using * unfolding wlsFSw_def gSwapAllPresGWlsAll_defs by simp
assume "\<forall> s. \<forall> Y \<in> termsOfS ?P s. fresh xs y Y"
hence y_fresh: "fresh xs y X" by simp
hence "gFresh MOD xs y X (?h X)"
using Abs pFresh unfolding presFreshAll_defs by simp
hence "gAbs MOD xs y (?Xsw) ?hXsw = gAbs MOD xs x X (?h X)"
using Abs hX y_fresh * unfolding gAbsRenS_def by fastforce
also have "\<dots> = ?hA B" using eq .
also have "recAbs MOD B = gAbs MOD xs y Y (?h Y)"
unfolding B_def using pCons Abs Y unfolding presCons_defs by blast
finally have "gAbs MOD xs y ?Xsw ?hXsw = gAbs MOD xs y Y (?h Y)" .
hence "?hXsw = ?h Y"
using ** Abs hX hXsw Y hY unfolding gConsInj_def gAbsInj_def
apply clarify apply(erule allE[of _ xs]) apply(erule allE[of _ s])
apply(erule allE[of _ y]) apply(erule allE[of _ ?Xsw]) by fastforce
moreover have "?hXsw = ?h ?Xsw"
using Abs pSwap unfolding presSwapAll_defs by simp
ultimately have "?h ?Xsw = ?h Y" by simp
moreover have "(X,?Xsw) \<in> swapped" using swap_swapped .
ultimately have Y_def: "Y = ?Xsw" using Y Abs.IH by auto
have ?thesis unfolding B_def Y_def
using Abs y_fresh by simp
}
thus ?thesis using B P wlsAbs_fresh_nchotomy[of xs s B] by blast
qed
qed
qed
}
thus ?thesis unfolding isInjAll_defs by blast
qed
-text{* For fresh-subst models *}
+text\<open>For fresh-subst models\<close>
theorem wlsFSb_recAll_isInjAll:
assumes *: "wlsFSb MOD" and **: "gConsInj MOD"
shows "isInjAll (rec MOD) (recAbs MOD)"
proof-
let ?h = "rec MOD" let ?hA = "recAbs MOD"
have 1: "termFSbMorph ?h ?hA MOD" using * wlsFSb_recAll_termFSbMorph by auto
hence pWls: "presWlsAll ?h ?hA MOD"
and pCons: "presCons ?h ?hA MOD"
and pFresh: "presFreshAll ?h ?hA MOD"
and pSubst: "presSubstAll ?h ?hA MOD" unfolding termFSbMorph_def by auto
hence pWlsInps[simp]:
"\<And> delta inp. wlsInp delta inp \<Longrightarrow> gWlsInp MOD delta (lift ?h inp)"
"\<And> delta binp. wlsBinp delta binp \<Longrightarrow> gWlsBinp MOD delta (lift ?hA binp)"
using presWls_wlsInp presWls_wlsBinp unfolding presWlsAll_def by auto
{fix s X us s' A
have
"(wls s X \<longrightarrow> (\<forall> Y. wls s Y \<and> rec MOD X = rec MOD Y \<longrightarrow> X = Y)) \<and>
(wlsAbs (us,s') A \<longrightarrow> (\<forall> B. wlsAbs (us,s') B \<and> recAbs MOD A = recAbs MOD B \<longrightarrow> A = B))"
proof(induction rule: wls_induct)
case (Var xs x)
show ?case proof clarify
fix Y
assume "rec MOD (Var xs x) = rec MOD Y" and Y: "wls (asSort xs) Y"
hence eq: "gVar MOD xs x = rec MOD Y" using 1 by (simp add: termFSbMorph_defs)
show "Var xs x = Y"
proof-
{fix ys y assume Y_def: "Y = Var ys y" and "asSort ys = asSort xs"
hence ys_def: "ys = xs" by simp
have rec_y_def: "rec MOD (Var ys y) = gVar MOD ys y"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq ** unfolding Y_def rec_y_def gConsInj_def gVarInj_def
unfolding ys_def by simp
}
moreover
{fix delta1 inp1 binp1 assume inp1s: "wlsInp delta1 inp1" "wlsBinp delta1 binp1"
and Y_def: "Y = Op delta1 inp1 binp1" and st: "stOf delta1 = asSort xs"
hence rec_Op_def:
"rec MOD (Op delta1 inp1 binp1) =
gOp MOD delta1 inp1 (lift ?h inp1) binp1 (lift ?hA binp1)"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq ** unfolding Y_def rec_Op_def gConsInj_def gVarGOpInj_def
using inp1s st by simp
}
ultimately show ?thesis using wls_nchotomy[of "asSort xs" Y] Y by blast
qed
qed
next
case (Op delta inp binp)
show ?case proof clarify
fix Y
assume "rec MOD (Op delta inp binp) = rec MOD Y" and Y: "wls (stOf delta) Y"
hence eq: "gOp MOD delta inp (lift ?h inp) binp (lift ?hA binp) = ?h Y"
using Op 1 by (simp add: termFSbMorph_defs)
show "Op delta inp binp = Y"
proof-
{fix ys y assume Y_def: "Y = Var ys y" and st: "asSort ys = stOf delta"
have rec_y_def: "rec MOD (Var ys y) = gVar MOD ys y"
using pCons unfolding presCons_defs by simp
have ?thesis
using eq[THEN sym] ** unfolding Y_def rec_y_def gConsInj_def gVarGOpInj_def
using Op st by simp
}
moreover
{fix delta1 inp1 binp1 assume inp1s: "wlsInp delta1 inp1" "wlsBinp delta1 binp1"
and Y_def: "Y = Op delta1 inp1 binp1" and st: "stOf delta1 = stOf delta"
hence rec_Op_def:
"rec MOD (Op delta1 inp1 binp1) =
gOp MOD delta1 inp1 (lift ?h inp1) binp1 (lift ?hA binp1)"
using pCons unfolding presCons_defs by simp
have 0: "delta = delta1 \<and> lift ?h inp = lift ?h inp1 \<and> lift ?hA binp = lift ?hA binp1"
using eq ** unfolding Y_def rec_Op_def gConsInj_def gOpInj_def
using Op inp1s st apply clarify
apply(erule allE[of _ delta]) apply(erule allE[of _ delta1]) by force
hence delta1_def: "delta1 = delta" by simp
have 1: "inp = inp1"
proof(rule ext)
fix i
show "inp i = inp1 i"
proof(cases "inp i")
case None
hence "lift ?h inp i = None" by(simp add: lift_None)
hence "lift ?h inp1 i = None" using 0 by simp
thus ?thesis unfolding None by(simp add: lift_None)
next
case (Some X)
hence "lift ?h inp i = Some (?h X)" unfolding lift_def by simp
hence "lift ?h inp1 i = Some (?h X)" using 0 by simp
then obtain Y where inp1_i: "inp1 i = Some Y" and hXY: "?h X = ?h Y"
unfolding lift_def by (cases "inp1 i") auto
then obtain s where ar_i: "arOf delta i = Some s"
using inp1s unfolding delta1_def wlsInp_iff sameDom_def
by (cases "arOf delta i") auto
hence Y: "wls s Y"
using inp1s inp1_i unfolding delta1_def wlsInp_iff liftAll2_def by auto
thus ?thesis unfolding Some inp1_i
using ar_i Some hXY Op.IH unfolding liftAll2_def by auto
qed
qed
have 2: "binp = binp1"
proof(rule ext)
fix i
show "binp i = binp1 i"
proof(cases "binp i")
case None
hence "lift ?hA binp i = None" by(simp add: lift_None)
hence "lift ?hA binp1 i = None" using 0 by simp
thus ?thesis unfolding None by(simp add: lift_None)
next
case (Some A)
hence "lift ?hA binp i = Some (?hA A)" unfolding lift_def by simp
hence "lift ?hA binp1 i = Some (?hA A)" using 0 by simp
then obtain B where binp1_i: "binp1 i = Some B" and hAB: "?hA A = ?hA B"
unfolding lift_def by(cases "binp1 i", auto)
then obtain us s where bar_i: "barOf delta i = Some (us,s)"
using inp1s unfolding delta1_def wlsBinp_iff sameDom_def
by(cases "barOf delta i") auto
hence B: "wlsAbs (us,s) B"
using inp1s binp1_i unfolding delta1_def wlsBinp_iff liftAll2_def by auto
thus ?thesis unfolding Some binp1_i
using bar_i Some hAB Op.IH
unfolding liftAll2_def by fastforce
qed
qed
have ?thesis unfolding Y_def delta1_def 1 2 by simp
}
ultimately show ?thesis using wls_nchotomy[of "stOf delta" Y] Y by blast
qed
qed
next
case (Abs s xs x X)
show ?case proof clarify
fix B
assume B: "wlsAbs (xs,s) B" and "recAbs MOD (Abs xs x X) = recAbs MOD B"
hence eq: "gAbs MOD xs x X (rec MOD X) = ?hA B"
using 1 Abs by (simp add: termFSbMorph_defs)
hence hX: "gWls MOD s (?h X)" using pWls Abs unfolding presWlsAll_defs by simp
show "Abs xs x X = B"
proof-
let ?P = "ParS
(\<lambda> xs'. [])
(\<lambda> s'. if s' = s then [X] else [])
(\<lambda> us_s. [])
[]"
have P: "wlsPar ?P" using Abs unfolding wlsPar_def by simp
{fix y Y assume Y: "wls s Y" and B_def: "B = Abs xs y Y"
hence hY: "gWls MOD s (?h Y)" using pWls unfolding presWlsAll_defs by simp
let ?Xsb = "X #[y // x]_xs"
let ?hXsb = "gSubst MOD xs (Var xs y) (gVar MOD xs y) x X (?h X)"
have 1: "wls (asSort xs) (Var xs y) \<and> gWls MOD (asSort xs) (gVar MOD xs y)"
using * unfolding wlsFSb_def gConsPresGWls_defs by simp
hence hXsb: "gWls MOD s ?hXsb"
using Abs hX using * unfolding wlsFSb_def gSubstAllPresGWlsAll_defs by simp
assume "\<forall> s. \<forall> Y \<in> termsOfS ?P s. fresh xs y Y"
hence y_fresh: "fresh xs y X" by simp
hence "gFresh MOD xs y X (?h X)"
using Abs pFresh unfolding presFreshAll_defs by simp
hence "gAbs MOD xs y (?Xsb) ?hXsb = gAbs MOD xs x X (?h X)"
using Abs hX y_fresh * unfolding wlsFSb_def gAbsRen_def by fastforce
also have "\<dots> = ?hA B" using eq .
also have "\<dots> = gAbs MOD xs y Y (?h Y)"
unfolding B_def using pCons Abs Y unfolding presCons_defs by blast
finally have
"gAbs MOD xs y ?Xsb ?hXsb = gAbs MOD xs y Y (?h Y)" .
hence "?hXsb = ?h Y"
using ** Abs hX hXsb Y hY unfolding gConsInj_def gAbsInj_def
apply clarify apply(erule allE[of _ xs]) apply(erule allE[of _ s])
apply(erule allE[of _ y]) apply(erule allE[of _ ?Xsb]) by fastforce
moreover have "?hXsb = ?h ?Xsb"
using Abs pSubst 1 pCons unfolding presSubstAll_defs vsubst_def presCons_defs by simp
ultimately have "?h ?Xsb = ?h Y" by simp
hence Y_def: "Y = ?Xsb" using Y Abs.IH by (fastforce simp add: termFSbMorph_defs)
have ?thesis unfolding B_def Y_def
using Abs y_fresh by simp
}
thus ?thesis using B P wlsAbs_fresh_nchotomy[of xs s B] by blast
qed
qed
qed
}
thus ?thesis unfolding isInjAll_defs by blast
qed
-subsubsection{* Criterion for the surjectiveness of the recursive map *}
+subsubsection\<open>Criterion for the surjectiveness of the recursive map\<close>
-text{* First an auxiliary fact, independent of the type of model: *}
+text\<open>First an auxiliary fact, independent of the type of model:\<close>
lemma gInduct_gConsIndif_recAll_isSurjAll:
assumes pWls: "presWlsAll (rec MOD) (recAbs MOD) MOD"
and pCons: "presCons (rec MOD) (recAbs MOD) MOD"
and "gConsIndif MOD" and *: "gInduct MOD"
shows "isSurjAll (rec MOD) (recAbs MOD) MOD"
proof-
let ?h = "rec MOD" let ?hA = "recAbs MOD"
{fix s X us s' A
from * have
"(gWls MOD s X \<longrightarrow> (\<exists> X'. wls s X' \<and> rec MOD X' = X)) \<and>
(gWlsAbs MOD (us,s') A \<longrightarrow> (\<exists> A'. wlsAbs (us,s') A' \<and> recAbs MOD A' = A))"
proof (elim gInduct_elim, safe)
fix xs x
show "\<exists>X'. wls (asSort xs) X' \<and> rec MOD X' = gVar MOD xs x"
using pWls pCons
by (auto simp: presWlsAll_defs presCons_defs intro: exI[of _ "Var xs x"])
next
fix delta inp' inp binp' binp
let ?ar = "arOf delta" let ?bar = "barOf delta" let ?st = "stOf delta"
assume inp': "wlsInp delta inp'" and binp': "wlsBinp delta binp'"
and inp: "gWlsInp MOD delta inp" and binp: "gWlsBinp MOD delta binp"
and IH: "liftAll2 (\<lambda>s X. \<exists>X'. wls s X' \<and> ?h X' = X) ?ar inp"
and BIH: "liftAll2 (\<lambda>us_s A. \<exists>A'. wlsAbs us_s A' \<and> ?hA A' = A) ?bar binp"
(* *)
let ?phi = "\<lambda> s X X'. wls s X' \<and> ?h X' = X"
obtain inp1' where inp1'_def:
"inp1' =
(\<lambda> i.
case (?ar i, inp i) of
(None, None) \<Rightarrow> None
|(Some s, Some X) \<Rightarrow> Some (SOME X'. ?phi s X X'))" by blast
hence [simp]:
"\<And> i. ?ar i = None \<and> inp i = None \<Longrightarrow> inp1' i = None"
"\<And> i s X. ?ar i = Some s \<and> inp i = Some X \<Longrightarrow> inp1' i = Some (SOME X'. ?phi s X X')"
unfolding inp1'_def by auto
have inp1': "wlsInp delta inp1'"
unfolding wlsInp_iff proof safe
show "sameDom ?ar inp1'"
unfolding sameDom_def proof clarify
fix i
have "(?ar i = None) = (inp i = None)"
using inp unfolding gWlsInp_def sameDom_def by simp
thus "(?ar i = None) = (inp1' i = None)"
unfolding inp1'_def by auto
qed
next
show "liftAll2 wls ?ar inp1'"
unfolding liftAll2_def proof auto
fix i s X1'
assume ari: "?ar i = Some s" and inp1'i: "inp1' i = Some X1'"
have "sameDom inp ?ar"
using inp unfolding gWlsInp_def using sameDom_sym by blast
then obtain X where inpi: "inp i = Some X"
using ari unfolding sameDom_def by(cases "inp i") auto
hence X1'_def: "X1' = (SOME X1'. ?phi s X X1')"
using ari inp1'i unfolding inp1'_def by simp
obtain X' where X': "?phi s X X'"
using inpi ari IH unfolding liftAll2_def by blast
hence "?phi s X X1'"
unfolding X1'_def by(rule someI[of "?phi s X"])
thus "wls s X1'" by simp
qed
qed(insert binp' wlsBinp.cases, blast)
(* *)
have lift_inp1': "lift ?h inp1' = inp"
proof(rule ext)
fix i let ?linp1' = "lift ?h inp1'"
show "?linp1' i = inp i"
proof(cases "inp i")
case None
hence "?ar i = None" using inp unfolding gWlsInp_def sameDom_def by simp
hence "inp1' i = None" using None by simp
thus "lift (rec MOD) inp1' i = inp i" using None by (auto simp: lift_def)
next
case (Some X)
then obtain s where ari: "?ar i = Some s"
using inp unfolding gWlsInp_def sameDom_def by(cases "?ar i") auto
let ?X1' = "SOME X1'. ?phi s X X1'"
have inp1'i: "inp1' i = Some ?X1'" using ari Some by simp
hence linp1'i: "?linp1' i = Some (?h ?X1')" unfolding lift_def by simp
obtain X' where X': "?phi s X X'"
using Some ari IH unfolding liftAll2_def by blast
hence "?phi s X ?X1'" by(rule someI[of "?phi s X"])
thus "lift (rec MOD) inp1' i = inp i" using Some linp1'i by (auto simp: lift_def)
qed
qed
(* *)
let ?bphi = "\<lambda> (us,s) A A'. wlsAbs (us,s) A' \<and> ?hA A' = A"
obtain binp1' where binp1'_def:
"binp1' =
(\<lambda> i.
case (?bar i, binp i) of
(None, None) \<Rightarrow> None
|(Some (us,s), Some A) \<Rightarrow> Some (SOME A'. ?bphi (us,s) A A'))" by blast
hence [simp]:
"\<And> i. ?bar i = None \<and> binp i = None \<Longrightarrow> binp1' i = None"
and *:
"\<And> i us s A. ?bar i = Some (us,s) \<and> binp i = Some A \<Longrightarrow>
binp1' i = Some (SOME A'. ?bphi (us,s) A A')"
unfolding binp1'_def by auto
have binp1': "wlsBinp delta binp1'"
unfolding wlsBinp_iff proof safe
show "sameDom ?bar binp1'"
unfolding sameDom_def proof clarify
fix i
have "(?bar i = None) = (binp i = None)"
using binp unfolding gWlsBinp_def sameDom_def by simp
thus "(?bar i = None) = (binp1' i = None)"
unfolding binp1'_def by auto
qed
next
show "liftAll2 wlsAbs ?bar binp1'"
unfolding liftAll2_def proof auto
fix i us s A1'
assume bari: "?bar i = Some (us,s)" and binp1'i: "binp1' i = Some A1'"
have "sameDom binp ?bar"
using binp unfolding gWlsBinp_def using sameDom_sym by blast
then obtain A where binpi: "binp i = Some A"
using bari unfolding sameDom_def by(cases "binp i", auto)
hence A1'_def: "A1' = (SOME A1'. ?bphi (us,s) A A1')"
using bari binp1'i unfolding binp1'_def by simp
obtain A' where A': "?bphi (us,s) A A'"
using binpi bari BIH unfolding liftAll2_def by fastforce
hence "?bphi (us,s) A A1'"
unfolding A1'_def by(rule someI[of "?bphi (us,s) A"])
thus "wlsAbs (us,s) A1'" by simp
qed
qed(insert binp' wlsBinp.cases, blast)
(* *)
have lift_binp1': "lift ?hA binp1' = binp"
proof(rule ext)
fix i let ?lbinp1' = "lift ?hA binp1'"
show "?lbinp1' i = binp i"
proof(cases "binp i")
case None
hence "?bar i = None" using binp unfolding gWlsBinp_def sameDom_def by simp
hence "binp1' i = None" using None by simp
thus "lift (recAbs MOD) binp1' i = binp i" using None by (simp add: lift_def)
next
case (Some A)
then obtain us s where bari: "?bar i = Some (us,s)"
using binp unfolding gWlsBinp_def sameDom_def by(cases "?bar i", auto)
let ?A1' = "SOME A1'. ?bphi (us,s) A A1'"
have binp1'i: "binp1' i = Some ?A1'" using bari Some *[of i us s A] by simp
hence lbinp1'i: "?lbinp1' i = Some (?hA ?A1')" unfolding lift_def by simp
obtain A' where A': "?bphi (us,s) A A'"
using Some bari BIH unfolding liftAll2_def by fastforce
hence "?bphi (us,s) A ?A1'" by(rule someI[of "?bphi (us,s) A"])
thus "lift (recAbs MOD) binp1' i = binp i" using Some lbinp1'i by simp
qed
qed
(* *)
let ?X' = "Op delta inp1' binp1'"
have X': "wls ?st ?X'" using inp1' binp1' by simp
have "?h ?X' = gOp MOD delta inp1' inp binp1' binp"
using inp1' binp1' pCons lift_inp1' lift_binp1'
unfolding presCons_defs by simp
hence "?h ?X' = gOp MOD delta inp' inp binp' binp"
using inp' inp1' inp binp' binp1' binp assms
unfolding gConsIndif_defs by metis
thus "\<exists>X'. wls (stOf delta) X' \<and> ?h X' = gOp MOD delta inp' inp binp' binp"
using X' by blast
next
fix xs s x X' X1'
assume xs_s: "isInBar (xs,s)" and X': "wls s X'" and
hX1': "gWls MOD s (?h X1')" and X1': "wls s X1'"
thus "\<exists>A'. wlsAbs (xs,s) A' \<and> ?hA A' = gAbs MOD xs x X' (?h X1')"
apply(intro exI[of _ "Abs xs x X1'"])
using pCons unfolding presCons_def presAbs_def apply safe
apply(elim allE[of _ xs]) apply(elim allE[of _ x]) apply(elim allE[of _ s])
apply simp_all
using assms unfolding gConsIndif_defs by blast
qed
}
thus ?thesis unfolding isSurjAll_defs by blast
qed
-text{* For fresh-swap models *}
+text\<open>For fresh-swap models\<close>
theorem wlsFSw_recAll_isSurjAll:
"wlsFSw MOD \<Longrightarrow> gConsIndif MOD \<Longrightarrow> gInduct MOD
\<Longrightarrow> isSurjAll (rec MOD) (recAbs MOD) MOD"
using wlsFSw_recAll_termFSwMorph
by (auto simp: termFSwMorph_def intro: gInduct_gConsIndif_recAll_isSurjAll)
-text{* For fresh-subst models *}
+text\<open>For fresh-subst models\<close>
theorem wlsFSb_recAll_isSurjAll:
"wlsFSb MOD \<Longrightarrow> gConsIndif MOD \<Longrightarrow> gInduct MOD
\<Longrightarrow> isSurjAll (rec MOD) (recAbs MOD) MOD"
using wlsFSb_recAll_termFSbMorph
by (auto simp: termFSbMorph_def intro: gInduct_gConsIndif_recAll_isSurjAll)
(********************************************)
lemmas recursion_simps =
fromMOD_simps ipresCons_fromMOD_fst_all_simps fromIMor_simps
declare recursion_simps [simp del]
end (* context FixSyn *)
(************************************************)
end
diff --git a/thys/Binding_Syntax_Theory/Semantic_Domains.thy b/thys/Binding_Syntax_Theory/Semantic_Domains.thy
--- a/thys/Binding_Syntax_Theory/Semantic_Domains.thy
+++ b/thys/Binding_Syntax_Theory/Semantic_Domains.thy
@@ -1,1029 +1,1029 @@
-section {* Interpretation of syntax in semantic domains *}
+section \<open>Interpretation of syntax in semantic domains\<close>
theory Semantic_Domains imports Iteration
begin
-text {* In this section, we employ our iteration principle
+text \<open>In this section, we employ our iteration principle
to obtain interpretation of syntax in semantic domains via valuations.
A bonus from our Horn-theoretic approach is the built-in
commutation of the interpretation with substitution versus valuation update,
-a property known in the literature as the ``substitution lemma". *}
+a property known in the literature as the ``substitution lemma".\<close>
-subsection {* Semantic domains and valuations *}
+subsection \<open>Semantic domains and valuations\<close>
-text{*
+text\<open>
Semantic domains are for binding signatures
what algebras are for standard algebraic signatures. They fix carrier sets for each sort,
and interpret each operation symbol as an operation on these sets
%
\footnote{
To match the Isabelle type system, we model (as usual) the family of carrier sets as a
``well-sortedness" predicate taking sorts and semantic items from a given
(initially unsorted) universe into booleans,
and require the operations, considered on the unsorted universe, to preserve well-sortedness.
}
%
of corresponding arity, where:
%
\\- non-binding arguments
are treated as usual (first-order) arguments;
%
\\- binding arguments are treated as second-order (functional) arguments.
%
\footnote{
In other words, syntactic bindings are captured semantically as functional bindings.}
%
In particular, for the untyped and simply-typed $\lambda$-calculi,
the semantic domains become the well-known (set-theoretic) Henkin models.
We use terminology and notation according to our general methodology employed so far:
the inhabitants of semantic domains are referred to as ``semantic items";
we prefix the reference to semantic items with an ``s": sX, sA, etc.
This convention also applies to the operations on semantic domains: ``sAbs", ``sOp", etc.
We eventually show that the function spaces consisting of maps
from valuations to semantic items form models;
in other words,
these maps can be viewed as ``generalized items"; we use for them
term-like notations ``X", ``A", etc.
(as we did in the theory that dealt with iteration).
-*}
+\<close>
-subsubsection {* Definitions: *}
+subsubsection \<open>Definitions:\<close>
datatype ('varSort,'sTerm)sAbs = sAbs 'varSort "'sTerm \<Rightarrow> 'sTerm"
record ('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom =
sWls :: "'sort \<Rightarrow> 'sTerm \<Rightarrow> bool"
sDummy :: "'sort \<Rightarrow> 'sTerm"
sOp :: "'opSym \<Rightarrow> ('index,'sTerm)input \<Rightarrow> ('bindex,('varSort,'sTerm)sAbs)input \<Rightarrow>'sTerm"
-text{* The type of valuations: *}
+text\<open>The type of valuations:\<close>
type_synonym ('varSort,'var,'sTerm)val = "'varSort \<Rightarrow> 'var \<Rightarrow> 'sTerm"
(* *************************************************** *)
context FixSyn
begin
(* A new type variable, corresponding to 'sTerm, is introduced in the context by
the following definitions and facts. *)
fun sWlsAbs where
"sWlsAbs SEM (xs,s) (sAbs xs' sF) =
(isInBar (xs,s) \<and> xs = xs' \<and>
(\<forall> sX. if sWls SEM (asSort xs) sX
then sWls SEM s (sF sX)
else sF sX = sDummy SEM s))"
definition sWlsInp where
"sWlsInp SEM delta sinp \<equiv>
wlsOpS delta \<and> sameDom (arOf delta) sinp \<and> liftAll2 (sWls SEM) (arOf delta) sinp"
definition sWlsBinp where
"sWlsBinp SEM delta sbinp \<equiv>
wlsOpS delta \<and> sameDom (barOf delta) sbinp \<and> liftAll2 (sWlsAbs SEM) (barOf delta) sbinp"
definition sWlsNE where
"sWlsNE SEM \<equiv>
\<forall> s. \<exists> sX. sWls SEM s sX"
definition sWlsDisj where
"sWlsDisj SEM \<equiv>
\<forall> s s' sX. sWls SEM s sX \<and> sWls SEM s' sX \<longrightarrow> s = s'"
definition sOpPrSWls where
"sOpPrSWls SEM \<equiv>
\<forall> delta sinp sbinp.
sWlsInp SEM delta sinp \<and> sWlsBinp SEM delta sbinp
\<longrightarrow> sWls SEM (stOf delta) (sOp SEM delta sinp sbinp)"
-text{* The notion of a ``well-sorted" (better read as ``well-structured")
+text\<open>The notion of a ``well-sorted" (better read as ``well-structured")
semantic domain:
%
\footnote{
As usual in Isabelle, we first define the ``raw" version,
and then ``fix" it with a well-structuredness predicate.
}
%
-*}
+\<close>
definition wlsSEM where
"wlsSEM SEM \<equiv>
sWlsNE SEM \<and> sWlsDisj SEM \<and> sOpPrSWls SEM"
-text{* The preperties described in the next 4 definitions turn out to be
-consequences of the well-structuredness of the semantic domain: *}
+text\<open>The preperties described in the next 4 definitions turn out to be
+consequences of the well-structuredness of the semantic domain:\<close>
definition sWlsAbsNE where
"sWlsAbsNE SEM \<equiv>
\<forall> us s. isInBar (us,s) \<longrightarrow> (\<exists> sA. sWlsAbs SEM (us,s) sA)"
definition sWlsAbsDisj where
"sWlsAbsDisj SEM \<equiv>
\<forall> us s us' s' sA.
isInBar (us,s) \<and> isInBar (us',s') \<and> sWlsAbs SEM (us,s) sA \<and> sWlsAbs SEM (us',s') sA
\<longrightarrow> us = us' \<and> s = s'"
-text{* The notion of two valuations being equal everywhere but on a given variable: *}
+text\<open>The notion of two valuations being equal everywhere but on a given variable:\<close>
definition eqBut where
"eqBut val val' xs x \<equiv>
\<forall> ys y. (ys = xs \<and> y = x) \<or> val ys y = val' ys y"
definition updVal ::
"('varSort,'var,'sTerm)val \<Rightarrow>
'var \<Rightarrow> 'sTerm \<Rightarrow> 'varSort \<Rightarrow>
('varSort,'var,'sTerm)val" ("_ '(_ := _')'__" 200)
where
"(val (x := sX)_xs) \<equiv>
\<lambda> ys y. (if ys = xs \<and> y = x then sX else val ys y)"
definition swapVal ::
"'varSort \<Rightarrow> 'var \<Rightarrow> 'var \<Rightarrow> ('varSort,'var,'sTerm)val \<Rightarrow>
('varSort,'var,'sTerm)val"
where
"swapVal zs z1 z2 val \<equiv> \<lambda>xs x. val xs (x @xs[z1 \<and> z2]_zs) "
abbreviation swapVal_abbrev ("_ ^[_ \<and> _]'__" 200) where
"val ^[z1 \<and> z2]_zs \<equiv> swapVal zs z1 z2 val"
definition sWlsVal where
"sWlsVal SEM val \<equiv>
\<forall> ys y. sWls SEM (asSort ys) (val ys y)"
(* The last argument is a dummy argument identifying the type 'var --
it may be regarded as a type argument: *)
definition sWlsValNE ::
"('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom \<Rightarrow> 'var \<Rightarrow> bool"
where
"sWlsValNE SEM x \<equiv> \<exists> (val :: ('varSort,'var,'sTerm)val). sWlsVal SEM val"
-subsubsection {* Basic facts *}
+subsubsection \<open>Basic facts\<close>
lemma sWlsNE_imp_sWlsAbsNE:
assumes "sWlsNE SEM"
shows "sWlsAbsNE SEM"
unfolding sWlsAbsNE_def proof clarify
fix xs s
obtain sY where "sWls SEM s sY"
using assms unfolding sWlsNE_def by auto
moreover assume "isInBar (xs,s)"
ultimately
have "sWlsAbs SEM (xs,s) (sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then sY
else sDummy SEM s))" by simp
thus "\<exists>sA. sWlsAbs SEM (xs,s) sA" by blast
qed
lemma sWlsDisj_imp_sWlsAbsDisj:
"sWlsDisj SEM \<Longrightarrow> sWlsNE SEM \<Longrightarrow> sWlsAbsDisj SEM"
by (simp add: sWlsAbsDisj_def sWlsNE_def sWlsDisj_def)
(smt prod.inject sAbs.inject sWlsAbs.elims(2))
lemma sWlsNE_imp_sWlsValNE:
"sWlsNE SEM \<Longrightarrow> sWlsValNE SEM x"
by (auto simp: sWlsNE_def sWlsValNE_def sWlsVal_def
intro!: exI someI_ex[of "(\<lambda> sY. sWls SEM (asSort _) sY)"])
theorem updVal_simp[simp]:
"(val (x := sX)_xs) ys y = (if ys = xs \<and> y = x then sX else val ys y)"
unfolding updVal_def by simp
theorem updVal_over[simp]:
"((val (x := sX)_xs) (x := sX')_xs) = (val (x := sX')_xs)"
unfolding updVal_def by fastforce
theorem updVal_commute:
assumes "xs \<noteq> ys \<or> x \<noteq> y"
shows "((val (x := sX)_xs) (y := sY)_ys) = ((val (y := sY)_ys) (x := sX)_xs)"
using assms unfolding updVal_def by fastforce
theorem updVal_preserves_sWls[simp]:
assumes "sWls SEM (asSort xs) sX" and "sWlsVal SEM val"
shows "sWlsVal SEM (val (x := sX)_xs)"
using assms unfolding sWlsVal_def by auto
lemmas updVal_simps = updVal_simp updVal_over updVal_preserves_sWls
theorem swapVal_ident[simp]: "(val ^[x \<and> x]_xs) = val"
unfolding swapVal_def by auto
theorem swapVal_compose:
"((val ^[x \<and> y]_zs) ^[x' \<and> y']_zs') =
((val ^[x' @zs'[x \<and> y]_zs \<and> y' @zs'[x \<and> y]_zs]_zs') ^[x \<and> y]_zs)"
unfolding swapVal_def by (metis sw_compose)
theorem swapVal_commute:
"zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {} \<Longrightarrow>
((val ^[x \<and> y]_zs) ^[x' \<and> y']_zs') = ((val ^[x' \<and> y']_zs') ^[x \<and> y]_zs)"
using swapVal_compose[of zs' x' y' zs x y val] by(simp add: sw_def)
lemma swapVal_involutive[simp]: "((val ^[x \<and> y]_zs) ^[x \<and> y]_zs) = val"
unfolding swapVal_def by auto
lemma swapVal_sym: "(val ^[x \<and> y]_zs) = (val ^[y \<and> x]_zs)"
unfolding swapVal_def by(auto simp add: sw_sym)
lemma swapVal_preserves_sWls1:
assumes "sWlsVal SEM val"
shows "sWlsVal SEM (val ^[z1 \<and> z2]_zs)"
using assms unfolding sWlsVal_def swapVal_def by simp
theorem swapVal_preserves_sWls[simp]:
"sWlsVal SEM (val ^[z1 \<and> z2]_zs) = sWlsVal SEM val"
using swapVal_preserves_sWls1[of _ _ zs z1 z2] by fastforce
lemmas swapVal_simps = swapVal_ident swapVal_involutive swapVal_preserves_sWls
lemma updVal_swapVal:
"((val (x := sX)_xs) ^[y1 \<and> y2]_ys) =
((val ^[y1 \<and> y2]_ys) ((x @xs[y1 \<and> y2]_ys) := sX)_xs)"
unfolding swapVal_def by fastforce
lemma updVal_preserves_eqBut:
assumes "eqBut val val' ys y"
shows "eqBut (val (x := sX)_xs) (val' (x := sX)_xs) ys y"
using assms unfolding eqBut_def updVal_def by auto
lemma updVal_eqBut_eq:
assumes "eqBut val val' ys y"
shows "(val (y := sY)_ys) = (val' (y := sY)_ys)"
using assms unfolding eqBut_def by fastforce
lemma swapVal_preserves_eqBut:
assumes "eqBut val val' xs x"
shows "eqBut (val ^[z1 \<and> z2]_zs) (val' ^[z1 \<and> z2]_zs) xs (x @xs[z1 \<and> z2]_zs)"
using assms unfolding eqBut_def swapVal_def by force
-subsection {* Interpretation maps *}
+subsection \<open>Interpretation maps\<close>
-text{* An interpretation map, of syntax in a semantic domain,
+text\<open>An interpretation map, of syntax in a semantic domain,
is the usual one w.r.t. valuations. Here we state its compostionality conditions
(including the ``substitution lemma"), and later we prove the existence of a map
-satisfying these conditions. *}
+satisfying these conditions.\<close>
-subsubsection {* Definitions *}
+subsubsection \<open>Definitions\<close>
-text {* Below, prefix ``pr" means ``preserves". *}
+text \<open>Below, prefix ``pr" means ``preserves".\<close>
definition prWls where
"prWls g SEM \<equiv> \<forall> s X val.
wls s X \<and> sWlsVal SEM val
\<longrightarrow> sWls SEM s (g X val)"
definition prWlsAbs where
"prWlsAbs gA SEM \<equiv> \<forall> us s A val.
wlsAbs (us,s) A \<and> sWlsVal SEM val
\<longrightarrow> sWlsAbs SEM (us,s) (gA A val)"
definition prWlsAll where
"prWlsAll g gA SEM \<equiv> prWls g SEM \<and> prWlsAbs gA SEM"
definition prVar where
"prVar g SEM \<equiv> \<forall> xs x val.
sWlsVal SEM val \<longrightarrow> g (Var xs x) val = val xs x"
definition prAbs where
"prAbs g gA SEM \<equiv> \<forall> xs s x X val.
isInBar (xs,s) \<and> wls s X \<and> sWlsVal SEM val
\<longrightarrow>
gA (Abs xs x X) val =
sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX then g X (val (x := sX)_xs)
else sDummy SEM s)"
definition prOp where
"prOp g gA SEM \<equiv> \<forall> delta inp binp val.
wlsInp delta inp \<and> wlsBinp delta binp \<and> sWlsVal SEM val
\<longrightarrow>
g (Op delta inp binp) val =
sOp SEM delta (lift (\<lambda>X. g X val) inp)
(lift (\<lambda>A. gA A val) binp)"
definition prCons where
"prCons g gA SEM \<equiv> prVar g SEM \<and> prAbs g gA SEM \<and> prOp g gA SEM"
definition prFresh where
"prFresh g SEM \<equiv> \<forall> ys y s X val val'.
wls s X \<and> fresh ys y X \<and>
sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> g X val = g X val'"
definition prFreshAbs where
"prFreshAbs gA SEM \<equiv> \<forall> ys y us s A val val'.
wlsAbs (us,s) A \<and> freshAbs ys y A \<and>
sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> gA A val = gA A val'"
definition prFreshAll where
"prFreshAll g gA SEM \<equiv> prFresh g SEM \<and> prFreshAbs gA SEM"
definition prSwap where
"prSwap g SEM \<equiv> \<forall> zs z1 z2 s X val.
wls s X \<and> sWlsVal SEM val
\<longrightarrow>
g (X #[z1 \<and> z2]_zs) val =
g X (val ^[z1 \<and> z2]_zs)"
definition prSwapAbs where
"prSwapAbs gA SEM \<equiv> \<forall> zs z1 z2 us s A val.
wlsAbs (us,s) A \<and> sWlsVal SEM val
\<longrightarrow>
gA (A $[z1 \<and> z2]_zs) val =
gA A (val ^[z1 \<and> z2]_zs)"
definition prSwapAll where
"prSwapAll g gA SEM \<equiv> prSwap g SEM \<and> prSwapAbs gA SEM"
definition prSubst where
"prSubst g SEM \<equiv> \<forall> ys Y y s X val.
wls (asSort ys) Y \<and> wls s X
\<and> sWlsVal SEM val
\<longrightarrow>
g (X #[Y / y]_ys) val =
g X (val (y := g Y val)_ys)"
definition prSubstAbs where
"prSubstAbs g gA SEM \<equiv> \<forall> ys Y y us s A val.
wls (asSort ys) Y \<and> wlsAbs (us,s) A
\<and> sWlsVal SEM val
\<longrightarrow>
gA (A $[Y / y]_ys) val =
gA A (val (y := g Y val)_ys)"
definition prSubstAll where
"prSubstAll g gA SEM \<equiv> prSubst g SEM \<and> prSubstAbs g gA SEM"
definition compInt where
"compInt g gA SEM \<equiv> prWlsAll g gA SEM \<and> prCons g gA SEM \<and>
prFreshAll g gA SEM \<and> prSwapAll g gA SEM \<and> prSubstAll g gA SEM"
-subsubsection {* Extension of domain preservation to inputs *}
+subsubsection \<open>Extension of domain preservation to inputs\<close>
lemma prWls_wlsInp:
assumes "wlsInp delta inp" and "prWls g SEM" and "sWlsVal SEM val"
shows "sWlsInp SEM delta (lift (\<lambda> X. g X val) inp)"
using assms unfolding sWlsInp_def wlsInp_iff liftAll2_def lift_def prWls_def
by (auto simp add: option.case_eq_if sameDom_def)
lemma prWlsAbs_wlsBinp:
assumes "wlsBinp delta binp" and "prWlsAbs gA SEM" and "sWlsVal SEM val"
shows "sWlsBinp SEM delta (lift (\<lambda> A. gA A val) binp)"
using assms unfolding sWlsBinp_def wlsBinp_iff liftAll2_def lift_def prWlsAbs_def
by (auto simp add: option.case_eq_if sameDom_def)
end (* context FixSyn *)
(***************************************)
-subsection {* The iterative model associated to a semantic domain *}
+subsection \<open>The iterative model associated to a semantic domain\<close>
-text{*
+text\<open>
``asIMOD SEM" stands for ``SEM (regarded) as a model".
%
\footnote{
We use the word ``model" as introduced in the theory ``Models-and-Recursion".
}
%
The associated model is built essentially as follows:
%
\\- Its carrier sets consist of functions from valuations to semantic items.
%
\\- The construct operations (i.e., those corresponding to the syntactic constructs
indicated in the given binding signature) are lifted componentwise from those of the semantic domain
``SEM" (also taking into account the higher-order nature of of the semantic counterparts of abstractions).
%
\\- For a map from valuations to items (terms or abstractions), freshness of a variable ``x"
is defined as being oblivious what the argument valuation returns for ``x".
%
\\- Swapping is defined componentwise, by two iterations of the notion of swapping the
returned value of a function.
%
\\- Substitution of a semantic term ``Y" for a variable ``y" is a semantic term ``X"
is defined to map each valuation ``val" to the application of ``X" to
[``val" updated at ``y" with whatever ``Y" returns for ``val"].
Note that:
%
\\- The construct operations definitions are determined by the desired clauses of the standard
notion of interpreting syntax in a semantic domains.
%
\\- Substitution and freshness are defined having in mind the (again standard) facts of
the interpretation commuting with substitution versus valuation update and the interpretation
being oblivious to the valuation of fresh variables.
-*}
+\<close>
-subsubsection {* Definition and basic facts *}
+subsubsection \<open>Definition and basic facts\<close>
-text{*
+text\<open>
The next two types of ``generalized items" are used to build models from semantic domains:
%
\footnote{
Recall that ``generalized items" inhabit models.
}
%
-*}
+\<close>
type_synonym ('varSort,'var,'sTerm) gTerm = "('varSort,'var,'sTerm)val \<Rightarrow> 'sTerm"
type_synonym ('varSort,'var,'sTerm) gAbs = "('varSort,'var,'sTerm)val \<Rightarrow> ('varSort,'sTerm)sAbs"
(* *************************************************** *)
context FixSyn
begin
definition asIMOD ::
"('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom \<Rightarrow>
('index,'bindex,'varSort,'sort,'opSym,'var,
('varSort,'var,'sTerm)gTerm,
('varSort,'var,'sTerm)gAbs)model"
where
"asIMOD SEM \<equiv>
\<lparr>igWls = \<lambda>s X. \<forall> val. (sWlsVal SEM val \<or> X val = undefined) \<and>
(sWlsVal SEM val \<longrightarrow> sWls SEM s (X val)),
igWlsAbs = \<lambda>(xs,s) A. \<forall> val. (sWlsVal SEM val \<or> A val = undefined) \<and>
(sWlsVal SEM val \<longrightarrow> sWlsAbs SEM (xs,s) (A val)),
igVar = \<lambda>ys y. \<lambda>val. if sWlsVal SEM val then val ys y else undefined,
igAbs =
\<lambda>xs x X. \<lambda>val. if sWlsVal SEM val
then sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then X (val (x := sX)_xs)
else sDummy SEM (SOME s. sWls SEM s (X val)))
else undefined,
igOp = \<lambda>delta inp binp. \<lambda>val.
if sWlsVal SEM val then sOp SEM delta (lift (\<lambda>X. X val) inp)
(lift (\<lambda>A. A val) binp)
else undefined,
igFresh =
\<lambda>ys y X. \<forall> val val'. sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> X val = X val',
igFreshAbs =
\<lambda>ys y A. \<forall> val val'. sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> A val = A val',
igSwap = \<lambda>zs z1 z2 X. \<lambda>val. if sWlsVal SEM val then X (val ^[z1 \<and> z2]_zs)
else undefined,
igSwapAbs = \<lambda>zs z1 z2 A. \<lambda>val. if sWlsVal SEM val then A (val ^[z1 \<and> z2]_zs)
else undefined,
igSubst = \<lambda>ys Y y X. \<lambda>val. if sWlsVal SEM val then X (val (y := Y val)_ys)
else undefined,
igSubstAbs = \<lambda>ys Y y A. \<lambda>val. if sWlsVal SEM val then A (val (y := Y val)_ys)
else undefined\<rparr>"
-text{* Next we state, as usual, the direct definitions of the operators and relations
+text\<open>Next we state, as usual, the direct definitions of the operators and relations
of associated model, freeing ourselves from
-having to go through the ``asIMOD" definition each time we reason about them. *}
+having to go through the ``asIMOD" definition each time we reason about them.\<close>
lemma asIMOD_igWls:
"igWls (asIMOD SEM) s X \<longleftrightarrow>
(\<forall> val. (sWlsVal SEM val \<or> X val = undefined) \<and>
(sWlsVal SEM val \<longrightarrow> sWls SEM s (X val)))"
unfolding asIMOD_def by simp
lemma asIMOD_igWlsAbs:
"igWlsAbs (asIMOD SEM) (us,s) A \<longleftrightarrow>
(\<forall> val. (sWlsVal SEM val \<or> A val = undefined) \<and>
(sWlsVal SEM val \<longrightarrow> sWlsAbs SEM (us,s) (A val)))"
unfolding asIMOD_def by simp
lemma asIMOD_igOp:
"igOp (asIMOD SEM) delta inp binp =
(\<lambda>val. if sWlsVal SEM val then sOp SEM delta (lift (\<lambda>X. X val) inp)
(lift (\<lambda>A. A val) binp)
else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igVar:
"igVar (asIMOD SEM) ys y = (\<lambda>val. if sWlsVal SEM val then val ys y else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igAbs:
"igAbs (asIMOD SEM) xs x X =
(\<lambda>val. if sWlsVal SEM val then sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then X (val (x := sX)_xs)
else sDummy SEM (SOME s. sWls SEM s (X val)))
else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igAbs2:
fixes SEM :: "('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom"
assumes *: "sWlsDisj SEM" and **: "igWls (asIMOD SEM) s X"
shows "igAbs (asIMOD SEM) xs x X =
(\<lambda>val. if sWlsVal SEM val then sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then X (val (x := sX)_xs)
else sDummy SEM s)
else undefined)"
proof-
{fix val :: "('varSort,'var,'sTerm)val" assume val: "sWlsVal SEM val"
hence Xval: "sWls SEM s (X val)"
using ** unfolding asIMOD_igWls by simp
hence "(SOME s. sWls SEM s (X val)) = s"
using Xval * unfolding sWlsDisj_def by auto
}
thus ?thesis unfolding asIMOD_igAbs by fastforce
qed
lemma asIMOD_igFresh:
"igFresh (asIMOD SEM) ys y X =
(\<forall> val val'. sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> X val = X val')"
unfolding asIMOD_def by simp
lemma asIMOD_igFreshAbs:
"igFreshAbs (asIMOD SEM) ys y A =
(\<forall> val val'. sWlsVal SEM val \<and> sWlsVal SEM val' \<and> eqBut val val' ys y
\<longrightarrow> A val = A val')"
unfolding asIMOD_def by simp
lemma asIMOD_igSwap:
"igSwap (asIMOD SEM) zs z1 z2 X =
(\<lambda>val. if sWlsVal SEM val then X (val ^[z1 \<and> z2]_zs) else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igSwapAbs:
"igSwapAbs (asIMOD SEM) zs z1 z2 A =
(\<lambda>val. if sWlsVal SEM val then A (val ^[z1 \<and> z2]_zs) else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igSubst:
"igSubst (asIMOD SEM) ys Y y X =
(\<lambda>val. if sWlsVal SEM val then X (val (y := Y val)_ys) else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igSubstAbs:
"igSubstAbs (asIMOD SEM) ys Y y A =
(\<lambda>val. if sWlsVal SEM val then A (val (y := Y val)_ys) else undefined)"
unfolding asIMOD_def by simp
lemma asIMOD_igWlsInp:
assumes "sWlsNE SEM"
shows
"igWlsInp (asIMOD SEM) delta inp \<longleftrightarrow>
((\<forall> val. liftAll (\<lambda>X. sWlsVal SEM val \<or> X val = undefined) inp) \<and>
(\<forall> val. sWlsVal SEM val \<longrightarrow> sWlsInp SEM delta (lift (\<lambda>X. X val) inp)))"
using assms apply safe
subgoal by (simp add: asIMOD_igWls liftAll_def liftAll2_def igWlsInp_def
sameDom_def split: option.splits) (metis option.distinct(1) option.exhaust)
subgoal by (simp add: igWlsInp_def asIMOD_igWls liftAll_def liftAll2_def
lift_def sWlsInp_def sameDom_def split: option.splits)
subgoal by (simp add:igWlsInp_def asIMOD_igWls liftAll_def liftAll2_def
lift_def sWlsInp_def sameDom_def split: option.splits)
(metis (no_types) option.distinct(1) sWlsNE_imp_sWlsValNE sWlsValNE_def) .
lemma asIMOD_igSwapInp:
"sWlsVal SEM val \<Longrightarrow>
lift (\<lambda>X. X val) (igSwapInp (asIMOD SEM) zs z1 z2 inp) =
lift (\<lambda>X. X (swapVal zs z1 z2 val)) inp"
by (auto simp: igSwapInp_def asIMOD_igSwap lift_def split: option.splits)
lemma asIMOD_igSubstInp:
"sWlsVal SEM val \<Longrightarrow>
lift (\<lambda>X. X val) (igSubstInp (asIMOD SEM) ys Y y inp) =
lift (\<lambda>X. X (val (y := Y val)_ys)) inp"
by (auto simp: igSubstInp_def asIMOD_igSubst lift_def split: option.splits)
lemma asIMOD_igWlsBinp:
assumes "sWlsNE SEM"
shows
"igWlsBinp (asIMOD SEM) delta binp =
((\<forall> val. liftAll (\<lambda>X. sWlsVal SEM val \<or> X val = undefined) binp) \<and>
(\<forall> val. sWlsVal SEM val \<longrightarrow> sWlsBinp SEM delta (lift (\<lambda>X. X val) binp)))"
using assms apply safe
subgoal by (simp add: asIMOD_igWlsAbs liftAll_def liftAll2_def igWlsBinp_def
sameDom_def split: option.splits)
(metis option.distinct(1) option.exhaust surj_pair)
subgoal by (simp add: igWlsBinp_def asIMOD_igWlsAbs liftAll_def liftAll2_def
lift_def sWlsBinp_def sameDom_def split: option.splits)
subgoal by (simp add:igWlsBinp_def asIMOD_igWlsAbs liftAll_def liftAll2_def
lift_def sWlsBinp_def sameDom_def split: option.splits)
(metis (no_types) old.prod.exhaust option.distinct(1) option.exhaust
sWlsNE_imp_sWlsValNE sWlsValNE_def) .
lemma asIMOD_igSwapBinp:
"sWlsVal SEM val \<Longrightarrow>
lift (\<lambda>A. A val) (igSwapBinp (asIMOD SEM) zs z1 z2 binp) =
lift (\<lambda>A. A (swapVal zs z1 z2 val)) binp"
by (auto simp: igSwapBinp_def asIMOD_igSwapAbs lift_def split: option.splits)
lemma asIMOD_igSubstBinp:
"sWlsVal SEM val \<Longrightarrow>
lift (\<lambda>A. A val) (igSubstBinp (asIMOD SEM) ys Y y binp) =
lift (\<lambda>A. A (val (y := Y val)_ys)) binp"
by (auto simp: igSubstBinp_def asIMOD_igSubstAbs lift_def split: option.splits)
-subsubsection {* The associated model is well-structured *}
+subsubsection \<open>The associated model is well-structured\<close>
-text{* That is to say: it is a fresh-swap-subst
-and fresh-subst-swap model (hence of course also a fresh-swap and fresh-subst) model. *}
+text\<open>That is to say: it is a fresh-swap-subst
+and fresh-subst-swap model (hence of course also a fresh-swap and fresh-subst) model.\<close>
-text{* Domain disjointness: *}
+text\<open>Domain disjointness:\<close>
lemma asIMOD_igWlsDisj:
"sWlsNE SEM \<Longrightarrow> sWlsDisj SEM \<Longrightarrow> igWlsDisj (asIMOD SEM)"
using sWlsNE_imp_sWlsValNE
by (fastforce simp: igWlsDisj_def asIMOD_igWls sWlsValNE_def sWlsDisj_def)
lemma asIMOD_igWlsAbsDisj:
"sWlsNE SEM \<Longrightarrow> sWlsDisj SEM \<Longrightarrow> igWlsAbsDisj (asIMOD SEM)"
using sWlsNE_imp_sWlsValNE sWlsDisj_imp_sWlsAbsDisj
by (fastforce simp: igWlsAbsDisj_def asIMOD_igWlsAbs sWlsAbsDisj_def sWlsValNE_def)
lemma asIMOD_igWlsAllDisj:
"sWlsNE SEM \<Longrightarrow> sWlsDisj SEM \<Longrightarrow> igWlsAllDisj (asIMOD SEM)"
unfolding igWlsAllDisj_def using asIMOD_igWlsDisj asIMOD_igWlsAbsDisj by auto
-text {* Only ``bound arit" abstraction domains are inhabited: *}
+text \<open>Only ``bound arit" abstraction domains are inhabited:\<close>
lemma asIMOD_igWlsAbsIsInBar:
"sWlsNE SEM \<Longrightarrow> igWlsAbsIsInBar (asIMOD SEM)"
using sWlsNE_imp_sWlsValNE
by (auto simp: sWlsValNE_def igWlsAbsIsInBar_def asIMOD_igWlsAbs
split: option.splits elim: sWlsAbs.elims(2))
-text{* Domain preservation by the operators *}
+text\<open>Domain preservation by the operators\<close>
-text{* The constructs preserve the domains: *}
+text\<open>The constructs preserve the domains:\<close>
lemma asIMOD_igVarIPresIGWls: "igVarIPresIGWls (asIMOD SEM)"
unfolding igVarIPresIGWls_def asIMOD_igWls asIMOD_igVar sWlsVal_def by simp
lemma asIMOD_igAbsIPresIGWls:
"sWlsDisj SEM \<Longrightarrow> igAbsIPresIGWls (asIMOD SEM)"
unfolding igAbsIPresIGWls_def asIMOD_igWlsAbs apply clarify
subgoal for _ _ _ _ val
unfolding asIMOD_igAbs2 by (cases "sWlsVal SEM val") (auto simp: asIMOD_igWls) .
lemma asIMOD_igOpIPresIGWls:
"sOpPrSWls SEM \<Longrightarrow> sWlsNE SEM \<Longrightarrow> igOpIPresIGWls (asIMOD SEM)"
using asIMOD_igWlsInp asIMOD_igWlsBinp
by (fastforce simp: igOpIPresIGWls_def asIMOD_igWls asIMOD_igOp sOpPrSWls_def)
lemma asIMOD_igConsIPresIGWls:
"wlsSEM SEM \<Longrightarrow> igConsIPresIGWls (asIMOD SEM)"
unfolding igConsIPresIGWls_def wlsSEM_def
using asIMOD_igVarIPresIGWls asIMOD_igAbsIPresIGWls asIMOD_igOpIPresIGWls by auto
-text{* Swap preserves the domains: *}
+text\<open>Swap preserves the domains:\<close>
lemma asIMOD_igSwapIPresIGWls: "igSwapIPresIGWls (asIMOD SEM)"
unfolding igSwapIPresIGWls_def asIMOD_igSwap asIMOD_igWls by auto
lemma asIMOD_igSwapAbsIPresIGWlsAbs: "igSwapAbsIPresIGWlsAbs (asIMOD SEM)"
unfolding igSwapAbsIPresIGWlsAbs_def asIMOD_igSwapAbs asIMOD_igWlsAbs by auto
lemma asIMOD_igSwapAllIPresIGWlsAll: "igSwapAllIPresIGWlsAll (asIMOD SEM)"
unfolding igSwapAllIPresIGWlsAll_def
using asIMOD_igSwapIPresIGWls asIMOD_igSwapAbsIPresIGWlsAbs by auto
-text {* Subst preserves the domains: *}
+text \<open>Subst preserves the domains:\<close>
lemma asIMOD_igSubstIPresIGWls: "igSubstIPresIGWls (asIMOD SEM)"
unfolding igSubstIPresIGWls_def asIMOD_igSubst asIMOD_igWls by simp
lemma asIMOD_igSubstAbsIPresIGWlsAbs: "igSubstAbsIPresIGWlsAbs (asIMOD SEM)"
unfolding igSubstAbsIPresIGWlsAbs_def asIMOD_igSubstAbs asIMOD_igWls asIMOD_igWlsAbs by simp
lemma asIMOD_igSubstAllIPresIGWlsAll: "igSubstAllIPresIGWlsAll (asIMOD SEM)"
unfolding igSubstAllIPresIGWlsAll_def
using asIMOD_igSubstIPresIGWls asIMOD_igSubstAbsIPresIGWlsAbs by auto
-text {* The clauses for fresh hold: *}
+text \<open>The clauses for fresh hold:\<close>
lemma asIMOD_igFreshIGVar: "igFreshIGVar (asIMOD SEM)"
unfolding igFreshIGVar_def asIMOD_igFresh asIMOD_igVar eqBut_def by force
lemma asIMOD_igFreshIGAbs1:
"sWlsDisj SEM \<Longrightarrow> igFreshIGAbs1 (asIMOD SEM)"
by(fastforce simp: igFreshIGAbs1_def asIMOD_igFresh asIMOD_igFreshAbs asIMOD_igAbs2 updVal_eqBut_eq)
lemma asIMOD_igFreshIGAbs2:
"sWlsDisj SEM \<Longrightarrow> igFreshIGAbs2 (asIMOD SEM)"
by(fastforce simp: igFreshIGAbs2_def asIMOD_igFresh asIMOD_igFreshAbs asIMOD_igAbs2 updVal_preserves_eqBut)
lemma asIMOD_igFreshIGOp:
fixes SEM :: "('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom"
shows "igFreshIGOp (asIMOD SEM)"
unfolding igFreshIGOp_def proof clarify
fix ys y delta and inp :: "('index, ('varSort,'var,'sTerm)gTerm)input"
and binp :: "('bindex, ('varSort,'var,'sTerm)gAbs)input"
assume inp_fresh: "igFreshInp (asIMOD SEM) ys y inp"
"igFreshBinp (asIMOD SEM) ys y binp"
show "igFresh (asIMOD SEM) ys y (igOp (asIMOD SEM) delta inp binp)"
unfolding asIMOD_igFresh asIMOD_igOp proof safe
fix val val'
let ?sinp = "lift (\<lambda>X. X val) inp" let ?sinp' = "lift (\<lambda>X. X val') inp"
let ?sbinp = "lift (\<lambda>A. A val) binp" let ?sbinp' = "lift (\<lambda>A. A val') binp"
assume wls: "sWlsVal SEM val" "sWlsVal SEM val'" and "eqBut val val' ys y"
hence "?sinp = ?sinp' \<and> ?sbinp = ?sbinp'"
using inp_fresh
by (auto simp: lift_def igFreshInp_def igFreshBinp_def errMOD_def liftAll_def
asIMOD_igFresh asIMOD_igFreshAbs split: option.splits)
then show "(if sWlsVal SEM val then sOp SEM delta (lift (\<lambda>X. X val) inp) (lift (\<lambda>A. A val) binp)
else undefined) =
(if sWlsVal SEM val' then sOp SEM delta (lift (\<lambda>X. X val') inp) (lift (\<lambda>A. A val') binp)
else undefined)" using wls by auto
qed
qed
lemma asIMOD_igFreshCls:
assumes "sWlsDisj SEM"
shows "igFreshCls (asIMOD SEM)"
using assms unfolding igFreshCls_def
using asIMOD_igFreshIGVar asIMOD_igFreshIGAbs1 asIMOD_igFreshIGAbs2 asIMOD_igFreshIGOp by auto
-text {* The clauses for swap hold: *}
+text \<open>The clauses for swap hold:\<close>
lemma asIMOD_igSwapIGVar: "igSwapIGVar (asIMOD SEM)"
unfolding igSwapIGVar_def apply clarsimp apply(rule ext)
unfolding asIMOD_igSwap asIMOD_igVar apply clarsimp
unfolding swapVal_def by simp
lemma asIMOD_igSwapIGAbs: "igSwapIGAbs (asIMOD SEM)"
by (fastforce simp: igSwapIGAbs_def asIMOD_igSwap asIMOD_igSwapAbs asIMOD_igAbs updVal_swapVal)
lemma asIMOD_igSwapIGOp: "igSwapIGOp (asIMOD SEM)"
by (auto simp: igSwapIGOp_def asIMOD_igSwap asIMOD_igOp asIMOD_igSwapInp asIMOD_igSwapBinp)
lemma asIMOD_igSwapCls: "igSwapCls (asIMOD SEM)"
unfolding igSwapCls_def using asIMOD_igSwapIGVar asIMOD_igSwapIGAbs asIMOD_igSwapIGOp by auto
-text{* The clauses for subst hold: *}
+text\<open>The clauses for subst hold:\<close>
lemma asIMOD_igSubstIGVar1: "igSubstIGVar1 (asIMOD SEM)"
by (auto simp: igSubstIGVar1_def asIMOD_igSubst asIMOD_igVar asIMOD_igWls)
lemma asIMOD_igSubstIGVar2: "igSubstIGVar2 (asIMOD SEM)"
by (fastforce simp: igSubstIGVar2_def asIMOD_igSubst asIMOD_igVar asIMOD_igWls)
lemma asIMOD_igSubstIGAbs: "igSubstIGAbs (asIMOD SEM)"
unfolding igSubstIGAbs_def proof(clarify, rule ext)
fix ys y Y xs x s X val
assume Y: "igWls (asIMOD SEM) (asSort ys) Y"
and X: "igWls (asIMOD SEM) s X" and x_diff_y: "xs \<noteq> ys \<or> x \<noteq> y"
and x_fresh_Y: "igFresh (asIMOD SEM) xs x Y"
show "igSubstAbs (asIMOD SEM) ys Y y (igAbs (asIMOD SEM) xs x X) val =
igAbs (asIMOD SEM) xs x (igSubst (asIMOD SEM) ys Y y X) val"
proof(cases "sWlsVal SEM val")
case False
thus ?thesis unfolding asIMOD_igSubst asIMOD_igSubstAbs asIMOD_igAbs by simp
next
case True
hence Yval: "sWls SEM (asSort ys) (Y val)"
using Y unfolding asIMOD_igWls by simp
{fix sX assume sX: "sWls SEM (asSort xs) sX"
let ?val_x = "val (x := sX)_xs"
have "sWlsVal SEM ?val_x" using True sX by simp
moreover have "eqBut ?val_x val xs x"
unfolding eqBut_def updVal_def by simp
ultimately have 1: "Y ?val_x = Y val"
using True x_fresh_Y unfolding asIMOD_igFresh by simp
let ?Left = "X ((val (y := Y val)_ys) (x := sX)_xs)"
let ?Riight = "X (?val_x (y := Y ?val_x)_ys)"
have "?Left = X (?val_x (y := Y val)_ys)"
using x_diff_y by(auto simp add: updVal_commute)
also have "\<dots> = ?Riight" using 1 by simp
finally have "?Left = ?Riight" .
}
thus ?thesis using True Yval by(auto simp: asIMOD_igSubst asIMOD_igSubstAbs asIMOD_igAbs)
qed
qed
lemma asIMOD_igSubstIGOp: "igSubstIGOp (asIMOD SEM)"
unfolding igSubstIGOp_def proof(clarify,rule ext)
fix ys y Y delta inp binp val
assume Y: "igWls (asIMOD SEM) (asSort ys) Y"
and inp: "igWlsInp (asIMOD SEM) delta inp"
and binp: "igWlsBinp (asIMOD SEM) delta binp"
define inpsb binpsb where
inpsb_def: "inpsb \<equiv> igSubstInp (asIMOD SEM) ys Y y inp"
"binpsb \<equiv> igSubstBinp (asIMOD SEM) ys Y y binp"
note inpsb_rev = inpsb_def[symmetric]
let ?sinpsb = "lift (\<lambda>X. X (val (y := Y val)_ys)) inp"
let ?sbinpsb = "lift (\<lambda>A. A (val (y := Y val)_ys)) binp"
show "igSubst (asIMOD SEM) ys Y y (igOp (asIMOD SEM) delta inp binp) val =
igOp (asIMOD SEM) delta (igSubstInp (asIMOD SEM) ys Y y inp)
(igSubstBinp (asIMOD SEM) ys Y y binp) val"
unfolding inpsb_rev unfolding asIMOD_igSubst asIMOD_igOp unfolding inpsb_def
apply(simp add: asIMOD_igSubstInp asIMOD_igSubstBinp)
using Y unfolding asIMOD_def by auto
qed
lemma asIMOD_igSubstCls: "igSubstCls (asIMOD SEM)"
unfolding igSubstCls_def
using asIMOD_igSubstIGVar1 asIMOD_igSubstIGVar2 asIMOD_igSubstIGAbs asIMOD_igSubstIGOp by auto
-text {* The fresh-swap-based congruence clause holds: *}
+text \<open>The fresh-swap-based congruence clause holds:\<close>
lemma updVal_swapVal_eqBut: "eqBut (val (x := sX)_xs) ((val (y := sX)_xs) ^[y \<and> x]_xs) xs y"
by (simp add: updVal_def swapVal_def eqBut_def sw_def)
lemma asIMOD_igAbsCongS: "sWlsDisj SEM \<Longrightarrow> igAbsCongS (asIMOD SEM)"
unfolding igAbsCongS_def asIMOD_igFresh asIMOD_igSwap asIMOD_igAbs2
apply safe apply (simp add: asIMOD_igAbs2)
by (rule ext) (metis (hide_lams) updVal_swapVal_eqBut swapVal_preserves_sWls updVal_preserves_sWls)
-text {* The abstraction-renaming clause holds: *}
+text \<open>The abstraction-renaming clause holds:\<close>
lemma asIMOD_igAbs3:
assumes "sWlsDisj SEM" and "igWls (asIMOD SEM) s X"
shows
"igAbs (asIMOD SEM) xs y (igSubst (asIMOD SEM) xs (igVar (asIMOD SEM) xs y) x X) =
(\<lambda>val. if sWlsVal SEM val
then sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then (igSubst (asIMOD SEM) xs (igVar (asIMOD SEM) xs y) x X) (val (y := sX)_xs)
else sDummy SEM s)
else undefined)"
using assms asIMOD_igVarIPresIGWls asIMOD_igSubstIPresIGWls
unfolding igVarIPresIGWls_def igSubstIPresIGWls_def
by (fastforce intro!: asIMOD_igAbs2)
lemma asIMOD_igAbsRen:
"sWlsDisj SEM \<Longrightarrow> igAbsRen (asIMOD SEM)"
unfolding igAbsRen_def asIMOD_igFresh asIMOD_igSwap apply safe
by (simp add: asIMOD_igAbs2 asIMOD_igAbs3)
(auto intro!: ext simp: asIMOD_igAbs2 asIMOD_igAbs3 eqBut_def asIMOD_igSubst asIMOD_igVar)
-text {* The associated model forms well-structured models of all 4 kinds: *}
+text \<open>The associated model forms well-structured models of all 4 kinds:\<close>
lemma asIMOD_wlsFSw:
assumes "wlsSEM SEM"
shows "iwlsFSw (asIMOD SEM)"
using assms unfolding wlsSEM_def iwlsFSw_def
using assms asIMOD_igWlsAllDisj asIMOD_igWlsAbsIsInBar
asIMOD_igConsIPresIGWls asIMOD_igSwapAllIPresIGWlsAll
asIMOD_igFreshCls asIMOD_igSwapCls asIMOD_igAbsCongS
by auto
lemma asIMOD_wlsFSb:
assumes "wlsSEM SEM"
shows "iwlsFSb (asIMOD SEM)"
using assms unfolding wlsSEM_def iwlsFSb_def
using assms asIMOD_igWlsAllDisj asIMOD_igWlsAbsIsInBar
asIMOD_igConsIPresIGWls[of SEM] asIMOD_igSubstAllIPresIGWlsAll
asIMOD_igFreshCls asIMOD_igSubstCls asIMOD_igAbsRen
by auto
lemma asIMOD_wlsFSwSb: "wlsSEM SEM \<Longrightarrow> iwlsFSwSb (asIMOD SEM)"
unfolding iwlsFSwSb_def
using asIMOD_wlsFSw asIMOD_igSubstAllIPresIGWlsAll asIMOD_igSubstCls by auto
lemma asIMOD_wlsFSbSw: "wlsSEM SEM \<Longrightarrow> iwlsFSbSw (asIMOD SEM)"
unfolding iwlsFSbSw_def
using asIMOD_wlsFSb asIMOD_igSwapAllIPresIGWlsAll asIMOD_igSwapCls by auto
-subsection {* The semantic interpretation *}
+subsection \<open>The semantic interpretation\<close>
-text{* The well-definedness of the semantic interpretation, as well
+text\<open>The well-definedness of the semantic interpretation, as well
as its associated substitution lemma and non-dependence of fresh variables,
are the end products of this theory.
Note that in order to establish these results either fresh-subst-swap or
fresh-swap-subst aligebras would do the job, and, moreover, if we did not care
about swapping, fresh-subst aligebras would do the job. Therefore, our
exhaustive study of the model from previous section had a deigree of redundancy w.r.t. to our main
igoal -- we pursued it however in order to better illustrate the rich structure laying under
the apparent paucity of the notion of a semantic domain. Next, we choose to employ
fresh-subst-swap aligebras to establish the required results. (Recall however that either aligebraic route
-we take, the initial morphism turns out to be the same function.)*}
+we take, the initial morphism turns out to be the same function.)\<close>
definition semInt where "semInt SEM \<equiv> iter (asIMOD SEM)"
definition semIntAbs where "semIntAbs SEM \<equiv> iterAbs (asIMOD SEM)"
lemma semIntAll_termFSwSbImorph:
"wlsSEM SEM \<Longrightarrow>
termFSwSbImorph (semInt SEM) (semIntAbs SEM) (asIMOD SEM)"
unfolding semInt_def semInt_def semIntAbs_def
using asIMOD_wlsFSbSw iwlsFSbSw_iterAll_termFSwSbImorph by auto
lemma semInt_prWls:
"wlsSEM SEM \<Longrightarrow> prWls (semInt SEM) SEM"
unfolding prWls_def using semIntAll_termFSwSbImorph
unfolding termFSwSbImorph_def termFSwImorph_def ipresWlsAll_def ipresWls_def asIMOD_igWls by auto
lemma semIntAbs_prWlsAbs:
"wlsSEM SEM \<Longrightarrow> prWlsAbs (semIntAbs SEM) SEM"
unfolding prWlsAbs_def using semIntAll_termFSwSbImorph
unfolding termFSwSbImorph_def termFSwImorph_def ipresWlsAll_def ipresWlsAbs_def asIMOD_igWlsAbs by blast
lemma semIntAll_prWlsAll:
"wlsSEM SEM \<Longrightarrow> prWlsAll (semInt SEM) (semIntAbs SEM) SEM"
unfolding prWlsAll_def by(simp add: semInt_prWls semIntAbs_prWlsAbs)
lemma semInt_prVar:
"wlsSEM SEM \<Longrightarrow> prVar (semInt SEM) SEM"
using semIntAll_termFSwSbImorph
unfolding prVar_def termFSwSbImorph_def termFSwImorph_def ipresCons_def ipresVar_def asIMOD_igVar
by fastforce
lemma semIntAll_prAbs:
fixes SEM :: "('index,'bindex,'varSort,'sort,'opSym,'sTerm)semDom"
assumes "wlsSEM SEM"
shows "prAbs (semInt SEM) (semIntAbs SEM) SEM"
proof-
{fix xs s x X and val :: "('varSort,'var,'sTerm)val"
assume xs_s: "isInBar (xs,s)" and X: "wls s X"
and val: "sWlsVal SEM val"
let ?L = "semIntAbs SEM (Abs xs x X)"
let ?R = "\<lambda> val. sAbs xs (\<lambda>sX. if sWls SEM (asSort xs) sX
then semInt SEM X (val (x := sX)_xs)
else sDummy SEM s)"
have "?L = igAbs (asIMOD SEM) xs x (semInt SEM X)"
using xs_s X assms semIntAll_termFSwSbImorph[of SEM]
unfolding termFSwSbImorph_def termFSwImorph_def ipresCons_def ipresAbs_def by auto
moreover
{have "prWls (semInt SEM) SEM" using assms semInt_prWls by auto
hence 1: "sWls SEM s (semInt SEM X val)"
using val X unfolding prWls_def by simp
hence "(SOME s. sWls SEM s (semInt SEM X val)) = s"
using 1 assms unfolding wlsSEM_def sWlsDisj_def by auto
hence "igAbs (asIMOD SEM) xs x (semInt SEM X) val = ?R val"
unfolding asIMOD_igAbs using val by fastforce
}
ultimately have "?L val = ?R val" by simp
}
thus ?thesis unfolding prAbs_def by auto
qed
lemma semIntAll_prOp:
assumes "wlsSEM SEM"
shows "prOp (semInt SEM) (semIntAbs SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prOp_def termFSwSbImorph_def termFSwImorph_def ipresCons_def ipresOp_def
asIMOD_igOp lift_comp comp_def by fastforce
lemma semIntAll_prCons:
assumes "wlsSEM SEM"
shows "prCons (semInt SEM) (semIntAbs SEM) SEM"
using assms unfolding prCons_def by(simp add: semInt_prVar semIntAll_prAbs semIntAll_prOp)
lemma semInt_prFresh:
assumes "wlsSEM SEM"
shows "prFresh (semInt SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prFresh_def termFSwSbImorph_def termFSwImorph_def ipresFreshAll_def ipresFresh_def
asIMOD_igFresh by fastforce
lemma semIntAbs_prFreshAbs:
assumes "wlsSEM SEM"
shows "prFreshAbs (semIntAbs SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prFreshAbs_def termFSwSbImorph_def termFSwImorph_def ipresFreshAll_def ipresFreshAbs_def
asIMOD_igFreshAbs by fastforce
lemma semIntAll_prFreshAll:
assumes "wlsSEM SEM"
shows "prFreshAll (semInt SEM) (semIntAbs SEM) SEM"
using assms unfolding prFreshAll_def by(simp add: semInt_prFresh semIntAbs_prFreshAbs)
lemma semInt_prSwap:
assumes "wlsSEM SEM"
shows "prSwap (semInt SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prSwap_def termFSwSbImorph_def termFSwImorph_def ipresSwapAll_def ipresSwap_def
asIMOD_igSwap by fastforce
lemma semIntAbs_prSwapAbs:
assumes "wlsSEM SEM"
shows "prSwapAbs (semIntAbs SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prSwapAbs_def termFSwSbImorph_def termFSwImorph_def ipresSwapAll_def ipresSwapAbs_def
asIMOD_igSwapAbs by fastforce
lemma semIntAll_prSwapAll:
assumes "wlsSEM SEM"
shows "prSwapAll (semInt SEM) (semIntAbs SEM) SEM"
using assms unfolding prSwapAll_def by(simp add: semInt_prSwap semIntAbs_prSwapAbs)
lemma semInt_prSubst:
assumes "wlsSEM SEM"
shows "prSubst (semInt SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prSubst_def termFSwSbImorph_def termFSwImorph_def ipresSubstAll_def ipresSubst_def
asIMOD_igSubst by fastforce
lemma semIntAbs_prSubstAbs:
assumes "wlsSEM SEM"
shows "prSubstAbs (semInt SEM) (semIntAbs SEM) SEM"
using assms semIntAll_termFSwSbImorph
unfolding prSubstAbs_def termFSwSbImorph_def termFSwImorph_def ipresSubstAll_def ipresSubstAbs_def
asIMOD_igSubstAbs by fastforce
lemma semIntAll_prSubstAll:
assumes "wlsSEM SEM"
shows "prSubstAll (semInt SEM) (semIntAbs SEM) SEM"
using assms unfolding prSubstAll_def by(simp add: semInt_prSubst semIntAbs_prSubstAbs)
theorem semIntAll_compInt:
assumes "wlsSEM SEM"
shows "compInt (semInt SEM) (semIntAbs SEM) SEM"
using assms unfolding compInt_def
by(simp add: semIntAll_prWlsAll semIntAll_prCons
semIntAll_prFreshAll semIntAll_prSwapAll semIntAll_prSubstAll)
lemmas semDom_simps = updVal_simps swapVal_simps
end (* context FixSyn *)
end
diff --git a/thys/Binding_Syntax_Theory/Terms.thy b/thys/Binding_Syntax_Theory/Terms.thy
--- a/thys/Binding_Syntax_Theory/Terms.thy
+++ b/thys/Binding_Syntax_Theory/Terms.thy
@@ -1,1274 +1,1274 @@
-section {* More on Terms *}
+section \<open>More on Terms\<close>
theory Terms imports Transition_QuasiTerms_Terms
begin
-text{* In this section, we continue the study of terms, with stating and proving
+text\<open>In this section, we continue the study of terms, with stating and proving
properties specific to terms (while in the previous section we dealt with
lifting properties from quasi-terms).
Consequently, in this theory, not only the theorems, but neither the proofs
should mention quasi-items at all.
Among the properties specific to terms will
be the compositionality properties of substitution (while, by contrast, similar properties
of swapping also held for quasi-tems).
- *}
+\<close>
context FixVars (* scope all throughout the file *)
begin
declare qItem_simps[simp del]
declare qItem_versus_item_simps[simp del]
-subsection {* Identity environment versus other operators *}
+subsection \<open>Identity environment versus other operators\<close>
(* Recall theorem getEnv_idEnv. *)
theorem getEnv_updEnv_idEnv[simp]:
"(idEnv [x \<leftarrow> X]_xs) ys y = (if (ys = xs \<and> y = x) then Some X else None)"
unfolding idEnv_def updEnv_def by simp
theorem subst_psubst_idEnv:
"(X #[Y / y]_ys) = (X #[idEnv [y \<leftarrow> Y]_ys])"
unfolding subst_def idEnv_def updEnv_def psubst_def by simp
theorem vsubst_psubst_idEnv:
"(X #[z // y]_ys) = (X #[idEnv [y \<leftarrow> Var ys z]_ys])"
unfolding vsubst_def by(simp add: subst_psubst_idEnv)
theorem substEnv_psubstEnv_idEnv:
"(rho &[Y / y]_ys) = (rho &[idEnv [y \<leftarrow> Y]_ys])"
unfolding substEnv_def idEnv_def updEnv_def psubstEnv_def by simp
theorem vsubstEnv_psubstEnv_idEnv:
"(rho &[z // y]_ys) = (rho &[idEnv [y \<leftarrow> Var ys z]_ys])"
unfolding vsubstEnv_def by (simp add: substEnv_psubstEnv_idEnv)
theorem freshEnv_idEnv: "freshEnv xs x idEnv"
unfolding idEnv_def freshEnv_def liftAll_def by simp
theorem swapEnv_idEnv[simp]: "(idEnv &[x \<and> y]_xs) = idEnv"
unfolding idEnv_def swapEnv_def comp_def swapEnvDom_def swapEnvIm_def lift_def by simp
theorem psubstEnv_idEnv[simp]: "(idEnv &[rho]) = rho"
unfolding idEnv_def psubstEnv_def lift_def by simp
theorem substEnv_idEnv: "(idEnv &[X / x]_xs) = (idEnv [x \<leftarrow> X]_xs)"
unfolding substEnv_def using psubstEnv_idEnv by auto
theorem vsubstEnv_idEnv: "(idEnv &[y // x]_xs) = (idEnv [x \<leftarrow> (Var xs y)]_xs)"
unfolding vsubstEnv_def using substEnv_idEnv .
lemma psubstAll_idEnv:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs"
shows
"(good X \<longrightarrow> (X #[idEnv]) = X) \<and>
(goodAbs A \<longrightarrow> (A $[idEnv]) = A)"
apply(induct rule: term_rawInduct)
unfolding psubstInp_def psubstBinp_def
using idEnv_preserves_good psubst_Var_simp1
by (simp_all del: getEnv_idEnv add:
liftAll_lift_ext lift_ident freshEnv_idEnv psubstBinp_def psubstInp_def)
fastforce+
lemma psubst_idEnv[simp]:
"good X \<Longrightarrow> (X #[idEnv]) = X"
by(simp add: psubstAll_idEnv)
lemma psubstEnv_idEnv_id[simp]:
assumes "goodEnv rho"
shows "(rho &[idEnv]) = rho"
using assms unfolding psubstEnv_def lift_def goodEnv_def liftAll_def
apply(intro ext) subgoal for xs x by(cases "rho xs x") auto .
-subsection {* Environment update versus other operators *}
+subsection \<open>Environment update versus other operators\<close>
(* Recall theorem getEnv_updEnv. *)
theorem updEnv_overwrite[simp]: "((rho [x \<leftarrow> X]_xs) [x \<leftarrow> X']_xs) = (rho [x \<leftarrow> X']_xs)"
unfolding updEnv_def by fastforce
theorem updEnv_commute:
assumes "xs \<noteq> ys \<or> x \<noteq> y"
shows "((rho [x \<leftarrow> X]_xs) [y \<leftarrow> Y]_ys) = ((rho [y \<leftarrow> Y]_ys) [x \<leftarrow> X]_xs)"
using assms unfolding updEnv_def by fastforce
theorem freshEnv_updEnv_E1:
assumes "freshEnv xs y (rho [x \<leftarrow> X]_xs)"
shows "y \<noteq> x"
using assms unfolding freshEnv_def liftAll_def updEnv_def by auto
theorem freshEnv_updEnv_E2:
assumes "freshEnv ys y (rho [x \<leftarrow> X]_xs)"
shows "fresh ys y X"
using assms unfolding freshEnv_def liftAll_def updEnv_def
by (auto split: if_splits)
theorem freshEnv_updEnv_E3:
assumes "freshEnv ys y (rho [x \<leftarrow> X]_xs)"
shows "rho ys y = None"
using assms freshEnv_updEnv_E1[of ys y] unfolding freshEnv_def
by (metis getEnv_updEnv option.simps(3))
theorem freshEnv_updEnv_E4:
assumes "freshEnv ys y (rho [x \<leftarrow> X]_xs)"
and "zs \<noteq> xs \<or> z \<noteq> x" and "rho zs z = Some Z"
shows "fresh ys y Z"
using assms unfolding freshEnv_def liftAll_def
by (metis getEnv_updEnv1)
theorem freshEnv_updEnv_I:
assumes "ys \<noteq> xs \<or> y \<noteq> x" and "fresh ys y X" and "rho ys y = None"
and "\<And> zs z Z. \<lbrakk>zs \<noteq> xs \<or> z \<noteq> x; rho zs z = Some Z\<rbrakk> \<Longrightarrow> fresh ys y Z"
shows "freshEnv ys y (rho [x \<leftarrow> X]_xs)"
unfolding freshEnv_def liftAll_def
using assms by auto
theorem swapEnv_updEnv:
"((rho [x \<leftarrow> X]_xs) &[y1 \<and> y2]_ys) =
((rho &[y1 \<and> y2]_ys) [(x @xs[y1 \<and> y2]_ys) \<leftarrow> (X #[y1 \<and> y2]_ys)]_xs)"
unfolding swapEnv_defs sw_def lift_def
by(cases "xs = ys") fastforce+
lemma swapEnv_updEnv_fresh:
assumes "ys \<noteq> xs \<or> x \<notin> {y1,y2}" and "good X"
and "fresh ys y1 X" and "fresh ys y2 X"
shows "((rho [x \<leftarrow> X]_xs) &[y1 \<and> y2]_ys) =
((rho &[y1 \<and> y2]_ys) [x \<leftarrow> X]_xs)"
using assms by(simp add: swapEnv_updEnv)
theorem psubstEnv_updEnv:
"((rho [x \<leftarrow> X]_xs) &[rho']) = ((rho &[rho']) [x \<leftarrow> (X #[rho'])]_xs)"
unfolding psubstEnv_def by fastforce
theorem psubstEnv_updEnv_idEnv:
"((idEnv [x \<leftarrow> X]_xs) &[rho]) = (rho [x \<leftarrow> (X #[rho])]_xs)"
by(simp add: psubstEnv_updEnv)
theorem substEnv_updEnv:
"((rho [x \<leftarrow> X]_xs) &[Y / y]_ys) = ((rho &[Y / y]_ys) [x \<leftarrow> (X #[Y / y]_ys)]_xs)"
unfolding substEnv_def subst_def by(rule psubstEnv_updEnv)
theorem vsubstEnv_updEnv:
"((rho [x \<leftarrow> X]_xs) &[y1 // y]_ys) = ((rho &[y1 // y]_ys) [x \<leftarrow> (X #[y1 // y]_ys)]_xs)"
unfolding vsubstEnv_def vsubst_def using substEnv_updEnv .
-subsection {* Environment ``get" versus other operators *}
+subsection \<open>Environment ``get" versus other operators\<close>
-text{* Currently, ``get" is just function application. While the next
+text\<open>Currently, ``get" is just function application. While the next
properties are immediate consequences of the definitions, it is worth stating
them because of their abstract character (since later, concrete terms
inferred from abstract terms by a presumptive package, ``get" will no longer
-be function application). *}
+be function application).\<close>
theorem getEnv_ext:
assumes "\<And> xs x. rho xs x = rho' xs x"
shows "rho = rho'"
using assms by(simp add: ext)
theorem freshEnv_getEnv1[simp]:
"\<lbrakk>freshEnv ys y rho; rho xs x = Some X\<rbrakk> \<Longrightarrow> ys \<noteq> xs \<or> y \<noteq> x"
unfolding freshEnv_def by auto
theorem freshEnv_getEnv2[simp]:
"\<lbrakk>freshEnv ys y rho; rho xs x = Some X\<rbrakk> \<Longrightarrow> fresh ys y X"
unfolding freshEnv_def liftAll_def by simp
theorem freshEnv_getEnv[simp]:
"freshEnv ys y rho \<Longrightarrow> rho ys y = None"
unfolding freshEnv_def by simp
theorem getEnv_swapEnv1[simp]:
assumes "rho xs (x @xs [z1 \<and> z2]_zs) = None"
shows "(rho &[z1 \<and> z2]_zs) xs x = None"
using assms unfolding swapEnv_defs lift_def by simp
theorem getEnv_swapEnv2[simp]:
assumes "rho xs (x @xs [z1 \<and> z2]_zs) = Some X"
shows "(rho &[z1 \<and> z2]_zs) xs x = Some (X #[z1 \<and> z2]_zs)"
using assms unfolding swapEnv_defs lift_def by simp
theorem getEnv_psubstEnv_None[simp]:
assumes "rho xs x = None"
shows "(rho &[rho']) xs x = rho' xs x"
using assms unfolding psubstEnv_def by simp
theorem getEnv_psubstEnv_Some[simp]:
assumes "rho xs x = Some X"
shows "(rho &[rho']) xs x = Some (X #[rho'])"
using assms unfolding psubstEnv_def by simp
theorem getEnv_substEnv1[simp]:
assumes "ys \<noteq> xs \<or> y \<noteq> x" and "rho xs x = None"
shows "(rho &[Y / y]_ys) xs x = None"
using assms unfolding substEnv_def2 by auto
theorem getEnv_substEnv2[simp]:
assumes "ys \<noteq> xs \<or> y \<noteq> x" and "rho xs x = Some X"
shows "(rho &[Y / y]_ys) xs x = Some (X #[Y / y]_ys)"
using assms unfolding substEnv_def2 by auto
theorem getEnv_substEnv3[simp]:
"\<lbrakk>ys \<noteq> xs \<or> y \<noteq> x; freshEnv xs x rho\<rbrakk>
\<Longrightarrow> (rho &[Y / y]_ys) xs x = None"
using getEnv_substEnv1 by auto
theorem getEnv_substEnv4[simp]:
"freshEnv ys y rho \<Longrightarrow> (rho &[Y / y]_ys) ys y = Some Y"
unfolding substEnv_psubstEnv_idEnv by simp
theorem getEnv_vsubstEnv1[simp]:
assumes "ys \<noteq> xs \<or> y \<noteq> x" and "rho xs x = None"
shows "(rho &[y1 // y]_ys) xs x = None"
using assms unfolding vsubstEnv_def by auto
theorem getEnv_vsubstEnv2[simp]:
assumes "ys \<noteq> xs \<or> y \<noteq> x" and "rho xs x = Some X"
shows "(rho &[y1 // y]_ys) xs x = Some (X #[y1 // y]_ys)"
using assms unfolding vsubstEnv_def vsubst_def by auto
theorem getEnv_vsubstEnv3[simp]:
"\<lbrakk>ys \<noteq> xs \<or> y \<noteq> x; freshEnv xs x rho\<rbrakk>
\<Longrightarrow> (rho &[z // y]_ys) xs x = None"
using getEnv_vsubstEnv1 by auto
theorem getEnv_vsubstEnv4[simp]:
"freshEnv ys y rho \<Longrightarrow> (rho &[z // y]_ys) ys y = Some (Var ys z)"
unfolding vsubstEnv_psubstEnv_idEnv by simp
-subsection {* Substitution versus other operators *}
+subsection \<open>Substitution versus other operators\<close>
definition freshImEnvAt ::
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)env \<Rightarrow> 'varSort \<Rightarrow> 'var \<Rightarrow> bool"
where
"freshImEnvAt xs x rho ys y ==
rho ys y = None \<and> (ys \<noteq> xs \<or> y \<noteq> x) \<or>
(\<exists> Y. rho ys y = Some Y \<and> fresh xs x Y)"
lemma freshAll_psubstAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param" and x
assumes goodP: "goodPar P"
shows
"(good X \<longrightarrow> z \<in> varsOf P \<longrightarrow>
(\<forall> rho \<in> envsOf P.
fresh zs z (X #[rho]) =
(\<forall> ys. \<forall> y. fresh ys y X \<or> freshImEnvAt zs z rho ys y)))
\<and>
(goodAbs A \<longrightarrow> z \<in> varsOf P \<longrightarrow>
(\<forall> rho \<in> envsOf P.
freshAbs zs z (A $[rho]) =
(\<forall> ys. \<forall> y. freshAbs ys y A \<or> freshImEnvAt zs z rho ys y)))"
proof(induction rule: term_induct_fresh[of P])
case Par
then show ?case using goodP by simp
next
case (Var ys y)
thus ?case proof clarify
fix rho
assume r: "rho \<in> envsOf P"
hence g: "goodEnv rho" using goodP by simp
thus "fresh zs z (psubst rho (Var ys y)) =
(\<forall>ysa ya. fresh ysa ya (Var ys y) \<or> freshImEnvAt zs z rho ysa ya)"
unfolding freshImEnvAt_def
by(cases "ys = zs \<and> y = z", (cases "rho ys y", auto)+)
qed
next
case (Op delta inp binp)
show ?case proof clarify
fix rho
assume P: "z \<in> varsOf P" "rho \<in> envsOf P"
let ?L1 = "liftAll (fresh zs z \<circ> psubst rho) inp"
let ?L2 = "liftAll (freshAbs zs z \<circ> psubstAbs rho) binp"
let ?R1 = "%ys y. liftAll (fresh ys y) inp"
let ?R2 = "%ys y. liftAll (freshAbs ys y) binp"
let ?R3 = "%ys y. freshImEnvAt zs z rho ys y"
have "(?L1 \<and> ?L2) = (\<forall>ys y. ?R1 ys y \<and> ?R2 ys y \<or> ?R3 ys y)"
using Op.IH P unfolding liftAll_def by simp blast
thus "fresh zs z ((Op delta inp binp) #[rho]) =
(\<forall>ys y. fresh ys y (Op delta inp binp) \<or> freshImEnvAt zs z rho ys y)"
by (metis (no_types, lifting) Op.hyps(1) Op.hyps(2) P(2) envsOf_preserves_good freshBinp_def freshInp_def fresh_Op_simp goodP liftAll_lift_comp psubstBinp_def psubstBinp_preserves_good
psubstInp_def psubstInp_preserves_good psubst_Op_simp)
qed
next
case (Abs xs x X)
thus ?case
using goodP by simp (metis (full_types) freshEnv_def freshImEnvAt_def)
qed
corollary fresh_psubst:
assumes "good X" and "goodEnv rho"
shows
"fresh zs z (X #[rho]) =
(\<forall> ys y. fresh ys y X \<or> freshImEnvAt zs z rho ys y)"
using assms freshAll_psubstAll[of "Par [z] [] [] [rho]"]
unfolding goodPar_def by simp
corollary fresh_psubst_E1:
assumes "good X" and "goodEnv rho"
and "rho ys y = None" and "fresh zs z (X #[rho])"
shows "fresh ys y X \<or> (ys \<noteq> zs \<or> y \<noteq> z)"
using assms fresh_psubst unfolding freshImEnvAt_def by fastforce
corollary fresh_psubst_E2:
assumes "good X" and "goodEnv rho"
and "rho ys y = Some Y" and "fresh zs z (X #[rho])"
shows "fresh ys y X \<or> fresh zs z Y"
using assms fresh_psubst[of X rho] unfolding freshImEnvAt_def by fastforce
corollary fresh_psubst_I1:
assumes "good X" and "goodEnv rho"
and "fresh zs z X" and "freshEnv zs z rho"
shows "fresh zs z (X #[rho])"
using assms apply(simp add: fresh_psubst)
unfolding freshEnv_def liftAll_def freshImEnvAt_def by auto
corollary psubstEnv_preserves_freshEnv:
assumes good: "goodEnv rho" "goodEnv rho'"
and fresh: "freshEnv zs z rho" "freshEnv zs z rho'"
shows "freshEnv zs z (rho &[rho'])"
using assms unfolding freshEnv_def liftAll_def
by simp (smt Var_preserves_good fresh(2) fresh_psubst_I1 option.case_eq_if
option.exhaust_sel option.sel psubstEnv_def psubst_Var_simp2 psubst_preserves_good)
corollary fresh_psubst_I:
assumes "good X" and "goodEnv rho"
and "rho zs z = None \<Longrightarrow> fresh zs z X" and
"\<And> ys y Y. rho ys y = Some Y \<Longrightarrow> fresh ys y X \<or> fresh zs z Y"
shows "fresh zs z (X #[rho])"
using assms unfolding freshImEnvAt_def
by (simp add: fresh_psubst) (metis freshImEnvAt_def not_None_eq)
lemma fresh_subst:
assumes "good X" and "good Y"
shows "fresh zs z (X #[Y / y]_ys) =
(((zs = ys \<and> z = y) \<or> fresh zs z X) \<and> (fresh ys y X \<or> fresh zs z Y))"
using assms unfolding subst_def freshImEnvAt_def
by (simp add: fresh_psubst)
(metis (no_types, lifting) freshImEnvAt_def fresh_psubst fresh_psubst_E2
getEnv_updEnv_idEnv idEnv_preserves_good option.simps(3) updEnv_preserves_good)
lemma fresh_vsubst:
assumes "good X"
shows "fresh zs z (X #[y1 // y]_ys) =
(((zs = ys \<and> z = y) \<or> fresh zs z X) \<and> (fresh ys y X \<or> (zs \<noteq> ys \<or> z \<noteq> y1)))"
unfolding vsubst_def using assms by(auto simp: fresh_subst)
lemma subst_preserves_fresh:
assumes "good X" and "good Y"
and "fresh zs z X" and "fresh zs z Y"
shows "fresh zs z (X #[Y / y]_ys)"
using assms by(simp add: fresh_subst)
lemma substEnv_preserves_freshEnv_aux:
assumes rho: "goodEnv rho" and Y: "good Y"
and fresh_rho: "freshEnv zs z rho" and fresh_Y: "fresh zs z Y" and diff: "zs \<noteq> ys \<or> z \<noteq> y"
shows "freshEnv zs z (rho &[Y / y]_ys)"
using assms unfolding freshEnv_def liftAll_def
by (simp add: option.case_eq_if substEnv_def2 subst_preserves_fresh)
lemma substEnv_preserves_freshEnv:
assumes rho: "goodEnv rho" and Y: "good Y"
and fresh_rho: "freshEnv zs z rho" and fresh_Y: "fresh zs z Y" and diff: "zs \<noteq> ys \<or> z \<noteq> y"
shows "freshEnv zs z (rho &[Y / y]_ys)"
using assms by(simp add: substEnv_preserves_freshEnv_aux)
lemma vsubst_preserves_fresh:
assumes "good X"
and "fresh zs z X" and "zs \<noteq> ys \<or> z \<noteq> y1"
shows "fresh zs z (X #[y1 // y]_ys)"
using assms by(simp add: fresh_vsubst)
lemma vsubstEnv_preserves_freshEnv:
assumes rho: "goodEnv rho"
and fresh_rho: "freshEnv zs z rho" and diff: "zs \<noteq> ys \<or> z \<notin> {y,y1}"
shows "freshEnv zs z (rho &[y1 // y]_ys)"
using assms unfolding vsubstEnv_def
by(simp add: substEnv_preserves_freshEnv)
lemma fresh_fresh_subst[simp]:
assumes "good Y" and "good X"
and "fresh ys y Y"
shows "fresh ys y (X #[Y / y]_ys)"
using assms by(simp add: fresh_subst)
lemma diff_fresh_vsubst[simp]:
assumes "good X"
and "y \<noteq> y1"
shows "fresh ys y (X #[y1 // y]_ys)"
using assms by(simp add: fresh_vsubst)
lemma fresh_subst_E1:
assumes "good X" and "good Y"
and "fresh zs z (X #[Y / y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "fresh zs z X"
using assms by(auto simp add: fresh_subst)
lemma fresh_vsubst_E1:
assumes "good X"
and "fresh zs z (X #[y1 // y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "fresh zs z X"
using assms by(auto simp add: fresh_vsubst)
lemma fresh_subst_E2:
assumes "good X" and "good Y"
and "fresh zs z (X #[Y / y]_ys)"
shows "fresh ys y X \<or> fresh zs z Y"
using assms by(simp add: fresh_subst)
lemma fresh_vsubst_E2:
assumes "good X"
and "fresh zs z (X #[y1 // y]_ys)"
shows "fresh ys y X \<or> zs \<noteq> ys \<or> z \<noteq> y1"
using assms by(simp add: fresh_vsubst)
lemma psubstAll_cong:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param"
assumes goodP: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> rho rho'. {rho, rho'} \<subseteq> envsOf P \<longrightarrow>
(\<forall> ys. \<forall> y. fresh ys y X \<or> rho ys y = rho' ys y) \<longrightarrow>
(X #[rho]) = (X #[rho'])))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> rho rho'. {rho, rho'} \<subseteq> envsOf P \<longrightarrow>
(\<forall> ys. \<forall> y. freshAbs ys y A \<or> rho ys y = rho' ys y) \<longrightarrow>
(A $[rho]) = (A $[rho'])))"
proof(induction rule: term_induct_fresh[of P])
case Par
then show ?case using assms .
next
case (Var xs x)
then show ?case using goodP by (auto simp: psubst_Var)
next
case (Op delta inp binp)
show ?case proof clarify
fix rho rho'
assume envs: "{rho, rho'} \<subseteq> envsOf P"
hence goodEnv: "goodEnv rho \<and> goodEnv rho'" using goodP by simp
assume "\<forall>ys y. fresh ys y (Op delta inp binp) \<or> rho ys y = rho' ys y"
hence 1: "liftAll (\<lambda> X. \<forall>ys y. fresh ys y X \<or> rho ys y = rho' ys y) inp \<and>
liftAll (\<lambda> A. \<forall>ys y. freshAbs ys y A \<or> rho ys y = rho' ys y) binp"
using Op by simp (smt freshBinp_def freshInp_def liftAll_def)
have "liftAll (\<lambda> X. (X #[rho]) = (X #[rho'])) inp \<and>
liftAll (\<lambda> A. (A $[rho]) = (A $[rho'])) binp"
using Op.IH 1 envs by (auto simp: liftAll_def)
thus "(Op delta inp binp) #[rho] = (Op delta inp binp) #[rho']"
using Op.IH 1
by (simp add: Op.hyps goodEnv psubstBinp_def psubstInp_def liftAll_lift_ext)
qed
next
case (Abs xs x X)
thus ?case
using Abs goodP unfolding freshEnv_def liftAll_def
by simp (metis Abs.hyps(5) envsOf_preserves_good psubstAbs_simp)
qed
corollary psubst_cong[fundef_cong]:
assumes "good X" and "goodEnv rho" and "goodEnv rho'"
and "\<And> ys y. fresh ys y X \<or> rho ys y = rho' ys y"
shows "(X #[rho]) = (X #[rho'])"
using assms psubstAll_cong[of "Par [] [] [] [rho,rho']"]
unfolding goodPar_def by simp
(* Note: A congruence principle for ``psubstEnv" would not hold w.r.t. ``freshEnv",
and the one that would hold w.r.t. ``fresh" would be a mere rephrasing of the
definition of ``psubstEnv", not worth stating. *)
lemma fresh_psubst_updEnv:
assumes "good X" and "good Y" and "goodEnv rho"
and "fresh xs x Y"
shows "(Y #[rho [x \<leftarrow> X]_xs]) = (Y #[rho])"
using assms by (auto cong: psubst_cong)
lemma psubstAll_ident:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)term" and
A :: "('index,'bindex,'varSort,'var,'opSym)abs" and
P :: "('index,'bindex,'varSort,'var,'opSym) Transition_QuasiTerms_Terms.param"
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> rho \<in> envsOf P.
(\<forall> zs z. freshEnv zs z rho \<or> fresh zs z X)
\<longrightarrow> (X #[rho]) = X))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> rho \<in> envsOf P.
(\<forall> zs z. freshEnv zs z rho \<or> freshAbs zs z A)
\<longrightarrow> (A $[rho]) = A))"
proof(induction rule: term_induct_fresh)
case (Var xs x)
then show ?case
by (meson assms freshEnv_def fresh_Var_simp goodPar_def psubst_Var_simp1)
next
case (Op delta inp binp)
then show ?case
by (metis (no_types,lifting) Op_preserves_good assms envsOf_preserves_good
freshEnv_getEnv idEnv_def idEnv_preserves_good psubst_cong psubst_idEnv)
qed(insert P, fastforce+)
corollary freshEnv_psubst_ident[simp]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)term"
assumes "good X" and "goodEnv rho"
and "\<And> zs z. freshEnv zs z rho \<or> fresh zs z X"
shows "(X #[rho]) = X"
using assms psubstAll_ident[of "Par [] [] [] [rho]"]
unfolding goodPar_def by simp
lemma fresh_subst_ident[simp]:
assumes "good X" and "good Y" and "fresh xs x Y"
shows "(Y #[X / x]_xs) = Y"
by (simp add: assms fresh_psubst_updEnv subst_def)
corollary substEnv_updEnv_fresh:
assumes "good X" and "good Y" and "fresh ys y X"
shows "((rho [x \<leftarrow> X]_xs) &[Y / y]_ys) = ((rho &[Y / y]_ys) [x \<leftarrow> X]_xs)"
using assms by(simp add: substEnv_updEnv)
lemma fresh_substEnv_updEnv[simp]:
assumes rho: "goodEnv rho" and Y: "good Y"
and *: "freshEnv ys y rho"
shows "(rho &[Y / y]_ys) = (rho [y \<leftarrow> Y]_ys)"
apply (rule getEnv_ext)
subgoal for xs x using assms by (cases "rho xs x") auto .
lemma fresh_vsubst_ident[simp]:
assumes "good Y" and "fresh xs x Y"
shows "(Y #[x1 // x]_xs) = Y"
using assms unfolding vsubst_def by simp
corollary vsubstEnv_updEnv_fresh:
assumes "good X" and "fresh ys y X"
shows "((rho [x \<leftarrow> X]_xs) &[y1 // y]_ys) = ((rho &[y1 // y]_ys) [x \<leftarrow> X]_xs)"
using assms by(simp add: vsubstEnv_updEnv)
lemma fresh_vsubstEnv_updEnv[simp]:
assumes rho: "goodEnv rho"
and *: "freshEnv ys y rho"
shows "(rho &[y1 // y]_ys) = (rho [y \<leftarrow> Var ys y1]_ys)"
using assms unfolding vsubstEnv_def by simp
lemma swapAll_psubstAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param"
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> rho z1 z2. rho \<in> envsOf P \<and> {z1,z2} \<subseteq> varsOf P \<longrightarrow>
((X #[rho]) #[z1 \<and> z2]_zs) = ((X #[z1 \<and> z2]_zs) #[rho &[z1 \<and> z2]_zs])))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> rho z1 z2. rho \<in> envsOf P \<and> {z1,z2} \<subseteq> varsOf P \<longrightarrow>
((A $[rho]) $[z1 \<and> z2]_zs) = ((A $[z1 \<and> z2]_zs) $[rho &[z1 \<and> z2]_zs])))"
proof(induction rule: term_induct_fresh[of P])
case (Var xs x)
then show ?case using assms
by simp (smt Var_preserves_good envsOf_preserves_good getEnv_swapEnv1 getEnv_swapEnv2 option.case_eq_if option.exhaust_sel psubst_Var psubst_Var_simp2 swapEnv_preserves_good
swap_Var_simp swap_involutive2 swap_sym)
next
case (Op delta inp binp)
then show ?case
using assms
unfolding psubstInp_def swapInp_def psubstBinp_def swapBinp_def lift_comp
unfolding liftAll_def lift_def
by simp (auto simp: lift_def psubstInp_def swapInp_def
psubstBinp_def swapBinp_def split: option.splits)
qed(insert assms, auto)
lemma swap_psubst:
assumes "good X" and "goodEnv rho"
shows "((X #[rho]) #[z1 \<and> z2]_zs) = ((X #[z1 \<and> z2]_zs) #[rho &[z1 \<and> z2]_zs])"
using assms swapAll_psubstAll[of "Par [z1,z2] [] [] [rho]"]
unfolding goodPar_def by auto
lemma swap_subst:
assumes "good X" and "good Y"
shows "((X #[Y / y]_ys) #[z1 \<and> z2]_zs) =
((X #[z1 \<and> z2]_zs) #[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
proof-
have 1: "(idEnv [(y @ys[z1 \<and> z2]_zs) \<leftarrow> (Y #[z1 \<and> z2]_zs)]_ys) =
((idEnv [y \<leftarrow> Y]_ys) &[z1 \<and> z2]_zs)"
by(simp add: swapEnv_updEnv)
show ?thesis
using assms unfolding subst_def 1 by (intro swap_psubst) auto
qed
lemma swap_vsubst:
assumes "good X"
shows "((X #[y1 // y]_ys) #[z1 \<and> z2]_zs) =
((X #[z1 \<and> z2]_zs) #[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms unfolding vsubst_def
by(simp add: swap_subst)
lemma swapEnv_psubstEnv:
assumes "goodEnv rho" and "goodEnv rho'"
shows "((rho &[rho']) &[z1 \<and> z2]_zs) = ((rho &[z1 \<and> z2]_zs) &[rho' &[z1 \<and> z2]_zs])"
using assms apply(intro ext)
subgoal for xs x
by (cases "rho xs (x @xs[z1 \<and> z2]_zs)")
(auto simp: lift_def swapEnv_defs swap_psubst) .
lemma swapEnv_substEnv:
assumes "good Y" and "goodEnv rho"
shows "((rho &[Y / y]_ys) &[z1 \<and> z2]_zs) =
((rho &[z1 \<and> z2]_zs) &[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
proof-
have 1: "(idEnv [(y @ys[z1 \<and> z2]_zs) \<leftarrow> (Y #[z1 \<and> z2]_zs)]_ys) =
((idEnv [y \<leftarrow> Y]_ys) &[z1 \<and> z2]_zs)"
by(simp add: swapEnv_updEnv)
show ?thesis
unfolding substEnv_def 1
using assms by (intro swapEnv_psubstEnv) auto
qed
lemma swapEnv_vsubstEnv:
assumes "goodEnv rho"
shows "((rho &[y1 // y]_ys) &[z1 \<and> z2]_zs) =
((rho &[z1 \<and> z2]_zs) &[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms unfolding vsubstEnv_def by(simp add: swapEnv_substEnv)
lemma psubstAll_compose:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param"
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> rho rho'. {rho,rho'} \<subseteq> envsOf P \<longrightarrow> ((X #[rho]) #[rho']) = (X #[(rho &[rho'])])))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> rho rho'. {rho,rho'} \<subseteq> envsOf P \<longrightarrow> ((A $[rho]) $[rho']) = (A $[(rho &[rho'])])))"
proof(induction rule: term_induct_fresh[of P])
case (Var xs x)
then show ?case using assms
by simp (smt envsOf_preserves_good option.case_eq_if option.sel psubstEnv_def
psubstEnv_idEnv_id psubstEnv_preserves_good psubst_Var_simp1 psubst_Var_simp2)
next
case (Op delta inp binp)
then show ?case
using assms
unfolding psubstInp_def swapInp_def psubstBinp_def swapBinp_def lift_comp
unfolding liftAll_def lift_def
by simp (auto simp: lift_def psubstInp_def swapInp_def
psubstBinp_def swapBinp_def split: option.splits)
qed(insert assms, simp_all add: psubstEnv_preserves_freshEnv)
corollary psubst_compose:
assumes "good X" and "goodEnv rho" and "goodEnv rho'"
shows "((X #[rho]) #[rho']) = (X #[(rho &[rho'])])"
using assms psubstAll_compose[of "Par [] [] [] [rho, rho']"]
unfolding goodPar_def by auto
lemma psubstEnv_compose:
assumes "goodEnv rho" and "goodEnv rho'" and "goodEnv rho''"
shows "((rho &[rho']) &[rho'']) = (rho &[(rho' &[rho''])])"
using assms apply(intro ext)
subgoal for xs x
by (cases "rho xs x") (auto simp: lift_def psubstEnv_def psubst_compose) .
lemma psubst_subst_compose:
assumes "good X" and "good Y" and "goodEnv rho"
shows "((X #[Y / y]_ys) #[rho]) = (X #[(rho [y \<leftarrow> (Y #[rho])]_ys)])"
by (simp add: assms psubstEnv_updEnv_idEnv psubst_compose subst_psubst_idEnv)
lemma psubstEnv_substEnv_compose:
assumes "goodEnv rho" and "good Y" and "goodEnv rho'"
shows "((rho &[Y / y]_ys) &[rho']) = (rho &[(rho' [y \<leftarrow> (Y #[rho'])]_ys)])"
by (simp add: assms psubstEnv_compose psubstEnv_updEnv_idEnv substEnv_def)
lemma psubst_vsubst_compose:
assumes "good X" and "goodEnv rho"
shows "((X #[y1 // y]_ys) #[rho]) = (X #[(rho [y \<leftarrow> ((Var ys y1) #[rho])]_ys)])"
using assms unfolding vsubst_def by(simp add: psubst_subst_compose)
lemma psubstEnv_vsubstEnv_compose:
assumes "goodEnv rho" and "goodEnv rho'"
shows "((rho &[y1 // y]_ys) &[rho']) = (rho &[(rho' [y \<leftarrow> ((Var ys y1) #[rho'])]_ys)])"
using assms unfolding vsubstEnv_def by(simp add: psubstEnv_substEnv_compose)
lemma subst_psubst_compose:
assumes "good X" and "good Y" and "goodEnv rho"
shows "((X #[rho]) #[Y / y]_ys) = (X #[(rho &[Y / y]_ys)])"
unfolding subst_def substEnv_def using assms by(simp add: psubst_compose)
lemma substEnv_psubstEnv_compose:
assumes "goodEnv rho" and "good Y" and "goodEnv rho'"
shows "((rho &[rho']) &[Y / y]_ys) = (rho &[(rho' &[Y / y]_ys)])"
unfolding substEnv_def using assms by(simp add: psubstEnv_compose)
lemma psubst_subst_compose_freshEnv:
assumes "goodEnv rho" and "good X" and "good Y"
assumes "freshEnv ys y rho"
shows "((X #[Y / y]_ys) #[rho]) = ((X #[rho]) #[(Y #[rho]) / y]_ys)"
using assms by (simp add: subst_psubst_compose psubst_subst_compose)
lemma psubstEnv_substEnv_compose_freshEnv:
assumes "goodEnv rho" and "goodEnv rho'" and "good Y"
assumes "freshEnv ys y rho'"
shows "((rho &[Y / y]_ys) &[rho']) = ((rho &[rho']) &[(Y #[rho']) / y]_ys)"
using assms by (simp add: substEnv_psubstEnv_compose psubstEnv_substEnv_compose)
lemma vsubst_psubst_compose:
assumes "good X" and "goodEnv rho"
shows "((X #[rho]) #[y1 // y]_ys) = (X #[(rho &[y1 // y]_ys)])"
unfolding vsubst_def vsubstEnv_def using assms by(simp add: subst_psubst_compose)
lemma vsubstEnv_psubstEnv_compose:
assumes "goodEnv rho" and "goodEnv rho'"
shows "((rho &[rho']) &[y1 // y]_ys) = (rho &[(rho' &[y1 // y]_ys)])"
unfolding vsubstEnv_def using assms by(simp add: substEnv_psubstEnv_compose)
lemma subst_compose1:
assumes "good X" and "good Y1" and "good Y2"
shows "((X #[Y1 / y]_ys) #[Y2 / y]_ys) = (X #[(Y1 #[Y2 / y]_ys) / y]_ys)"
proof-
have "goodEnv (idEnv [y \<leftarrow> Y1]_ys) \<and> goodEnv (idEnv [y \<leftarrow> Y2]_ys)" using assms by simp
- thus ?thesis using `good X` unfolding subst_def substEnv_def
+ thus ?thesis using \<open>good X\<close> unfolding subst_def substEnv_def
by(simp add: psubst_compose psubstEnv_updEnv)
qed
lemma substEnv_compose1:
assumes "goodEnv rho" and "good Y1" and "good Y2"
shows "((rho &[Y1 / y]_ys) &[Y2 / y]_ys) = (rho &[(Y1 #[Y2 / y]_ys) / y]_ys)"
by (simp add: assms psubstEnv_compose psubstEnv_updEnv_idEnv substEnv_def subst_psubst_idEnv)
lemma subst_vsubst_compose1:
assumes "good X" and "good Y" and "y \<noteq> y1"
shows "((X #[y1 // y]_ys) #[Y / y]_ys) = (X #[y1 // y]_ys)"
using assms unfolding vsubst_def by(simp add: subst_compose1)
lemma substEnv_vsubstEnv_compose1:
assumes "goodEnv rho" and "good Y" and "y \<noteq> y1"
shows "((rho &[y1 // y]_ys) &[Y / y]_ys) = (rho &[y1 // y]_ys)"
using assms unfolding vsubst_def vsubstEnv_def by(simp add: substEnv_compose1)
lemma vsubst_subst_compose1:
assumes "good X" and "good Y"
shows "((X #[Y / y]_ys) #[y1 // y]_ys) = (X #[(Y #[y1 // y]_ys) / y]_ys)"
using assms unfolding vsubst_def by(simp add: subst_compose1)
lemma vsubstEnv_substEnv_compose1:
assumes "goodEnv rho" and "good Y"
shows "((rho &[Y / y]_ys) &[y1 // y]_ys) = (rho &[(Y #[y1 // y]_ys) / y]_ys)"
using assms unfolding vsubst_def vsubstEnv_def by(simp add: substEnv_compose1)
lemma vsubst_compose1:
assumes "good X"
shows "((X #[y1 // y]_ys) #[y2 // y]_ys) = (X #[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms unfolding vsubst_def
by(cases "y = y1") (auto simp: subst_compose1)
lemma vsubstEnv_compose1:
assumes "goodEnv rho"
shows "((rho &[y1 // y]_ys) &[y2 // y]_ys) = (rho &[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms unfolding vsubstEnv_def
by(cases "y = y1") (auto simp: substEnv_compose1)
lemma subst_compose2:
assumes "good X" and "good Y" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((X #[Y / y]_ys) #[Z / z]_zs) = ((X #[Z / z]_zs) #[(Y #[Z / z]_zs) / y]_ys)"
by (metis assms fresh freshEnv_getEnv freshEnv_getEnv2 freshEnv_idEnv freshEnv_updEnv_I idEnv_preserves_good psubst_subst_compose_freshEnv
subst_psubst_idEnv updEnv_preserves_good)
lemma substEnv_compose2:
assumes "goodEnv rho" and "good Y" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((rho &[Y / y]_ys) &[Z / z]_zs) = ((rho &[Z / z]_zs) &[(Y #[Z / z]_zs) / y]_ys)"
by (metis assms fresh freshEnv_updEnv_I getEnv_idEnv idEnv_preserves_good
option.discI psubstEnv_substEnv_compose_freshEnv substEnv_def
subst_psubst_idEnv updEnv_preserves_good)
lemma subst_vsubst_compose2:
assumes "good X" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((X #[y1 // y]_ys) #[Z / z]_zs) = ((X #[Z / z]_zs) #[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms unfolding vsubst_def by(simp add: subst_compose2)
lemma substEnv_vsubstEnv_compose2:
assumes "goodEnv rho" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((rho &[y1 // y]_ys) &[Z / z]_zs) = ((rho &[Z / z]_zs) &[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms unfolding vsubstEnv_def by(simp add: substEnv_compose2)
lemma vsubst_subst_compose2:
assumes "good X" and "good Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((X #[Y / y]_ys) #[z1 // z]_zs) = ((X #[z1 // z]_zs) #[(Y #[z1 // z]_zs) / y]_ys)"
using assms unfolding vsubst_def by(simp add: subst_compose2)
lemma vsubstEnv_substEnv_compose2:
assumes "goodEnv rho" and "good Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((rho &[Y / y]_ys) &[z1 // z]_zs) = ((rho &[z1 // z]_zs) &[(Y #[z1 // z]_zs) / y]_ys)"
using assms unfolding vsubst_def vsubstEnv_def by(simp add: substEnv_compose2)
lemma vsubst_compose2:
assumes "good X"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((X #[y1 // y]_ys) #[z1 // z]_zs) =
((X #[z1 // z]_zs) #[(y1 @ys[z1 / z]_zs) // y]_ys)"
by (metis vsubst_def Var_preserves_good assms vsubst_Var_simp vsubst_def
vsubst_subst_compose2)
lemma vsubstEnv_compose2:
assumes "goodEnv rho"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((rho &[y1 // y]_ys) &[z1 // z]_zs) =
((rho &[z1 // z]_zs) &[(y1 @ys[z1 / z]_zs) // y]_ys)"
by (metis Var_preserves_good assms
vsubstEnv_def vsubstEnv_substEnv_compose2 vsubst_Var_simp)
-subsection {* Properties specific to variable-for-variable substitution *}
+subsection \<open>Properties specific to variable-for-variable substitution\<close>
(* Note: The results in this section cannot be lifted to environments, and therefore
we don't have ``environment versions" of these. *)
lemma vsubstAll_ident:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param" and zs
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> z. z \<in> varsOf P \<longrightarrow> (X #[z // z]_zs) = X))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> z. z \<in> varsOf P \<longrightarrow> (A $[z // z]_zs) = A))"
proof(induct rule: term_induct_fresh[of P])
case (Op delta inp binp)
then show ?case
using assms
unfolding vsubst_def vsubstAbs_def liftAll_def lift_def
by simp (auto simp: lift_def substInp_def2 substBinp_def2 vsubstInp_def2
split: option.splits)
next
case (Abs xs x X)
then show ?case
by (metis empty_iff insert_iff vsubstAbs_simp)
qed(insert assms, simp_all)
corollary vsubst_ident[simp]:
assumes "good X"
shows "(X #[z // z]_zs) = X"
using assms vsubstAll_ident[of "Par [z] [] [] []" X]
unfolding goodPar_def by simp
corollary subst_ident[simp]:
assumes "good X"
shows "(X #[(Var zs z) / z]_zs) = X"
using assms vsubst_ident unfolding vsubst_def by auto
lemma vsubstAll_swapAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param" and ys
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> y1 y2. {y1,y2} \<subseteq> varsOf P \<and> fresh ys y1 X \<longrightarrow>
(X #[y1 // y2]_ys) = (X #[y1 \<and> y2]_ys)))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> y1 y2. {y1,y2} \<subseteq> varsOf P \<and> freshAbs ys y1 A \<longrightarrow>
(A $[y1 // y2]_ys) = (A $[y1 \<and> y2]_ys)))"
apply(induction rule: term_induct_fresh[OF P])
subgoal by (force simp add: sw_def)
subgoal by simp (auto
simp: vsubstInp_def substInp_def2 vsubst_def swapInp_def
vsubstBinp_def substBinp_def2 vsubstAbs_def swapBinp_def
freshInp_def freshBinp_def lift_def liftAll_def
split: option.splits)
subgoal by simp (metis Var_preserves_good fresh_Var_simp substAbs_simp sw_def
vsubstAbs_def vsubst_def) .
corollary vsubst_eq_swap:
assumes "good X" and "y1 = y2 \<or> fresh ys y1 X"
shows "(X #[y1 // y2]_ys) = (X #[y1 \<and> y2]_ys)"
apply(cases "y1 = y2")
using assms vsubstAll_swapAll[of "Par [y1, y2] [] [] []" X]
unfolding goodPar_def by auto
lemma skelAll_vsubstAll:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
P::"('index,'bindex,'varSort,'var,'opSym)param" and ys
assumes P: "goodPar P"
shows
"(good X \<longrightarrow>
(\<forall> y1 y2. {y1,y2} \<subseteq> varsOf P \<longrightarrow>
skel (X #[y1 // y2]_ys) = skel X))
\<and>
(goodAbs A \<longrightarrow>
(\<forall> y1 y2. {y1,y2} \<subseteq> varsOf P \<longrightarrow>
skelAbs (A $[y1 // y2]_ys) = skelAbs A))"
proof(induction rule: term_induct_fresh[of P])
case (Op delta inp binp)
then show ?case
by (simp add: skelInp_def2 skelBinp_def2)
(auto simp: vsubst_def vsubstInp_def substInp_def2
vsubstAbs_def vsubstBinp_def substBinp_def2 lift_def liftAll_def
split: option.splits)
next
case (Abs xs x X)
then show ?case using assms
by simp (metis not_equals_and_not_equals_not_in
skelAbs_simp vsubstAbs_simp vsubst_preserves_good)
qed(insert assms, simp_all)
corollary skel_vsubst:
assumes "good X"
shows "skel (X #[y1 // y2]_ys) = skel X"
using assms skelAll_vsubstAll[of "Par [y1, y2] [] [] []" X]
unfolding goodPar_def by simp
lemma subst_vsubst_trans:
assumes "good X" and "good Y" and "fresh ys y1 X"
shows "((X #[y1 // y]_ys) #[Y / y1]_ys) = (X #[Y / y]_ys)"
using assms unfolding subst_def vsubst_def
by (cases "y1 = y") (simp_all add: fresh_psubst_updEnv psubstEnv_updEnv_idEnv
psubst_compose updEnv_commute)
lemma vsubst_trans:
assumes "good X" and "fresh ys y1 X"
shows "((X #[y1 // y]_ys) #[y2 // y1]_ys) = (X #[y2 // y]_ys)"
unfolding vsubst_def[of _ y2 y1] vsubst_def[of _ y2 y]
using assms by(simp add: subst_vsubst_trans)
lemma vsubst_commute:
assumes X: "good X"
and "xs \<noteq> xs' \<or> {x,y} \<inter> {x',y'} = {}" and "fresh xs x X" and "fresh xs' x' X"
shows "((X #[x // y]_xs) #[x' // y']_xs') = ((X #[x' // y']_xs') #[x // y]_xs)"
proof-
have "fresh xs' x' (X #[x // y]_xs)"
using assms by (intro vsubst_preserves_fresh) auto
moreover have "fresh xs x (X #[x' // y']_xs')"
using assms by (intro vsubst_preserves_fresh) auto
ultimately show ?thesis using assms
by (auto simp: vsubst_eq_swap intro!: swap_commute)
qed
-subsection {* Abstraction versions of the properties *}
+subsection \<open>Abstraction versions of the properties\<close>
-text{* Environment identity and update versus other operators: *}
+text\<open>Environment identity and update versus other operators:\<close>
lemma psubstAbs_idEnv[simp]:
"goodAbs A \<Longrightarrow> (A $[idEnv]) = A"
by(simp add: psubstAll_idEnv)
-text{* Substitution versus other operators: *}
+text\<open>Substitution versus other operators:\<close>
corollary freshAbs_psubstAbs:
assumes "goodAbs A" and "goodEnv rho"
shows
"freshAbs zs z (A $[rho]) =
(\<forall> ys y. freshAbs ys y A \<or> freshImEnvAt zs z rho ys y)"
using assms freshAll_psubstAll[of "Par [z] [] [] [rho]"]
unfolding goodPar_def by simp
corollary freshAbs_psubstAbs_E1:
assumes "goodAbs A" and "goodEnv rho"
and "rho ys y = None" and "freshAbs zs z (A $[rho])"
shows "freshAbs ys y A \<or> (ys \<noteq> zs \<or> y \<noteq> z)"
using assms freshAbs_psubstAbs unfolding freshImEnvAt_def by fastforce
corollary freshAbs_psubstAbs_E2:
assumes "goodAbs A" and "goodEnv rho"
and "rho ys y = Some Y" and "freshAbs zs z (A $[rho])"
shows "freshAbs ys y A \<or> fresh zs z Y"
using assms freshAbs_psubstAbs[of A rho] unfolding freshImEnvAt_def by fastforce
corollary freshAbs_psubstAbs_I1:
assumes "goodAbs A" and "goodEnv rho"
and "freshAbs zs z A" and "freshEnv zs z rho"
shows "freshAbs zs z (A $[rho])"
using assms apply(simp add: freshAbs_psubstAbs)
unfolding freshEnv_def liftAll_def freshImEnvAt_def by auto
corollary freshAbs_psubstAbs_I:
assumes "goodAbs A" and "goodEnv rho"
and "rho zs z = None \<Longrightarrow> freshAbs zs z A" and
"\<And> ys y Y. rho ys y = Some Y \<Longrightarrow> freshAbs ys y A \<or> fresh zs z Y"
shows "freshAbs zs z (A $[rho])"
using assms using option.exhaust_sel
by (simp add: freshAbs_psubstAbs freshImEnvAt_def) blast
lemma freshAbs_substAbs:
assumes "goodAbs A" and "good Y"
shows "freshAbs zs z (A $[Y / y]_ys) =
(((zs = ys \<and> z = y) \<or> freshAbs zs z A) \<and> (freshAbs ys y A \<or> fresh zs z Y))"
unfolding substAbs_def using assms
by (auto simp: freshAbs_psubstAbs freshImEnvAt_def)
lemma freshAbs_vsubstAbs:
assumes "goodAbs A"
shows "freshAbs zs z (A $[y1 // y]_ys) =
(((zs = ys \<and> z = y) \<or> freshAbs zs z A) \<and>
(freshAbs ys y A \<or> (zs \<noteq> ys \<or> z \<noteq> y1)))"
unfolding vsubstAbs_def using assms by(auto simp: freshAbs_substAbs)
lemma substAbs_preserves_freshAbs:
assumes "goodAbs A" and "good Y"
and "freshAbs zs z A" and "fresh zs z Y"
shows "freshAbs zs z (A $[Y / y]_ys)"
using assms by(simp add: freshAbs_substAbs)
lemma vsubstAbs_preserves_freshAbs:
assumes "goodAbs A"
and "freshAbs zs z A" and "zs \<noteq> ys \<or> z \<noteq> y1"
shows "freshAbs zs z (A $[y1 // y]_ys)"
using assms by(simp add: freshAbs_vsubstAbs)
lemma fresh_freshAbs_substAbs[simp]:
assumes "good Y" and "goodAbs A"
and "fresh ys y Y"
shows "freshAbs ys y (A $[Y / y]_ys)"
using assms by(simp add: freshAbs_substAbs)
lemma diff_freshAbs_vsubstAbs[simp]:
assumes "goodAbs A"
and "y \<noteq> y1"
shows "freshAbs ys y (A $[y1 // y]_ys)"
using assms by(simp add: freshAbs_vsubstAbs)
lemma freshAbs_substAbs_E1:
assumes "goodAbs A" and "good Y"
and "freshAbs zs z (A $[Y / y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "freshAbs zs z A"
using assms by(auto simp: freshAbs_substAbs)
lemma freshAbs_vsubstAbs_E1:
assumes "goodAbs A"
and "freshAbs zs z (A $[y1 // y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "freshAbs zs z A"
using assms by(auto simp: freshAbs_vsubstAbs)
lemma freshAbs_substAbs_E2:
assumes "goodAbs A" and "good Y"
and "freshAbs zs z (A $[Y / y]_ys)"
shows "freshAbs ys y A \<or> fresh zs z Y"
using assms by(simp add: freshAbs_substAbs)
lemma freshAbs_vsubstAbs_E2:
assumes "goodAbs A"
and "freshAbs zs z (A $[y1 // y]_ys)"
shows "freshAbs ys y A \<or> zs \<noteq> ys \<or> z \<noteq> y1"
using assms by(simp add: freshAbs_vsubstAbs)
corollary psubstAbs_cong[fundef_cong]:
assumes "goodAbs A" and "goodEnv rho" and "goodEnv rho'"
and "\<And> ys y. freshAbs ys y A \<or> rho ys y = rho' ys y"
shows "(A $[rho]) = (A $[rho'])"
using assms psubstAll_cong[of "Par [] [] [] [rho,rho']"]
unfolding goodPar_def by simp
lemma freshAbs_psubstAbs_updEnv:
assumes "good X" and "goodAbs A" and "goodEnv rho"
and "freshAbs xs x A"
shows "(A $[rho [x \<leftarrow> X]_xs]) = (A $[rho])"
using assms by (intro psubstAbs_cong) auto
corollary freshEnv_psubstAbs_ident[simp]:
fixes A :: "('index,'bindex,'varSort,'var,'opSym)abs"
assumes "goodAbs A" and "goodEnv rho"
and "\<And> zs z. freshEnv zs z rho \<or> freshAbs zs z A"
shows "(A $[rho]) = A"
using assms psubstAll_ident[of "Par [] [] [] [rho]"]
unfolding goodPar_def by simp
lemma freshAbs_substAbs_ident[simp]:
assumes "good X" and "goodAbs A" and "freshAbs xs x A"
shows "(A $[X / x]_xs) = A"
by (simp add: assms freshAbs_psubstAbs_updEnv substAbs_def)
corollary substAbs_Abs[simp]:
assumes "good X" and "good Y"
shows "((Abs xs x X) $[Y / x]_xs) = Abs xs x X"
using assms by simp
lemma freshAbs_vsubstAbs_ident[simp]:
assumes "goodAbs A" and "freshAbs xs x A"
shows "(A $[x1 // x]_xs) = A"
using assms unfolding vsubstAbs_def by(auto simp: freshAbs_substAbs_ident)
lemma swapAbs_psubstAbs:
assumes "goodAbs A" and "goodEnv rho"
shows "((A $[rho]) $[z1 \<and> z2]_zs) = ((A $[z1 \<and> z2]_zs) $[rho &[z1 \<and> z2]_zs])"
using assms swapAll_psubstAll[of "Par [z1,z2] [] [] [rho]"]
unfolding goodPar_def by auto
lemma swapAbs_substAbs:
assumes "goodAbs A" and "good Y"
shows "((A $[Y / y]_ys) $[z1 \<and> z2]_zs) =
((A $[z1 \<and> z2]_zs) $[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
proof-
have 1: "(idEnv [(y @ys[z1 \<and> z2]_zs) \<leftarrow> (Y #[z1 \<and> z2]_zs)]_ys) =
((idEnv [y \<leftarrow> Y]_ys) &[z1 \<and> z2]_zs)"
by(simp add: swapEnv_updEnv)
show ?thesis
unfolding substAbs_def 1 using assms by (intro swapAbs_psubstAbs) auto
qed
lemma swapAbs_vsubstAbs:
assumes "goodAbs A"
shows "((A $[y1 // y]_ys) $[z1 \<and> z2]_zs) =
((A $[z1 \<and> z2]_zs) $[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms unfolding vsubstAbs_def
by(simp add: swapAbs_substAbs)
lemma psubstAbs_compose:
assumes "goodAbs A" and "goodEnv rho" and "goodEnv rho'"
shows "((A $[rho]) $[rho']) = (A $[(rho &[rho'])])"
using assms psubstAll_compose[of "Par [] [] [] [rho, rho']"]
unfolding goodPar_def by auto
lemma psubstAbs_substAbs_compose:
assumes "goodAbs A" and "good Y" and "goodEnv rho"
shows "((A $[Y / y]_ys) $[rho]) = (A $[(rho [y \<leftarrow> (Y #[rho])]_ys)])"
by (simp add: assms psubstAbs_compose psubstEnv_updEnv_idEnv substAbs_def)
lemma psubstAbs_vsubstAbs_compose:
assumes "goodAbs A" and "goodEnv rho"
shows "((A $[y1 // y]_ys) $[rho]) = (A $[(rho [y \<leftarrow> ((Var ys y1) #[rho])]_ys)])"
using assms unfolding vsubstAbs_def by(simp add: psubstAbs_substAbs_compose)
lemma substAbs_psubstAbs_compose:
assumes "goodAbs A" and "good Y" and "goodEnv rho"
shows "((A $[rho]) $[Y / y]_ys) = (A $[(rho &[Y / y]_ys)])"
unfolding substAbs_def substEnv_def using assms by(simp add: psubstAbs_compose)
lemma psubstAbs_substAbs_compose_freshEnv:
assumes "goodAbs A" and "goodEnv rho" and "good Y"
assumes "freshEnv ys y rho"
shows "((A $[Y / y]_ys) $[rho]) = ((A $[rho]) $[(Y #[rho]) / y]_ys)"
using assms by (simp add: substAbs_psubstAbs_compose psubstAbs_substAbs_compose)
lemma vsubstAbs_psubstAbs_compose:
assumes "goodAbs A" and "goodEnv rho"
shows "((A $[rho]) $[y1 // y]_ys) = (A $[(rho &[y1 // y]_ys)])"
unfolding vsubstAbs_def vsubstEnv_def using assms
by(simp add: substAbs_psubstAbs_compose)
lemma substAbs_compose1:
assumes "goodAbs A" and "good Y1" and "good Y2"
shows "((A $[Y1 / y]_ys) $[Y2 / y]_ys) = (A $[(Y1 #[Y2 / y]_ys) / y]_ys)"
by (metis assms idEnv_preserves_good psubstAbs_substAbs_compose substAbs_def
subst_psubst_idEnv updEnv_overwrite updEnv_preserves_good)
lemma substAbs_vsubstAbs_compose1:
assumes "goodAbs A" and "good Y" and "y \<noteq> y1"
shows "((A $[y1 // y]_ys) $[Y / y]_ys) = (A $[y1 // y]_ys)"
using assms unfolding vsubstAbs_def by(simp add: substAbs_compose1)
lemma vsubstAbs_substAbs_compose1:
assumes "goodAbs A" and "good Y"
shows "((A $[Y / y]_ys) $[y1 // y]_ys) = (A $[(Y #[y1 // y]_ys) / y]_ys)"
using assms unfolding vsubstAbs_def vsubst_def by(simp add: substAbs_compose1)
lemma vsubstAbs_compose1:
assumes "goodAbs A"
shows "((A $[y1 // y]_ys) $[y2 // y]_ys) = (A $[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms unfolding vsubstAbs_def
by(cases "y = y1") (auto simp: substAbs_compose1)
lemma substAbs_compose2:
assumes "goodAbs A" and "good Y" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((A $[Y / y]_ys) $[Z / z]_zs) = ((A $[Z / z]_zs) $[(Y #[Z / z]_zs) / y]_ys)"
by (metis assms fresh freshEnv_idEnv idEnv_preserves_good
psubstAbs_substAbs_compose_freshEnv substAbs_def
substEnv_idEnv substEnv_preserves_freshEnv_aux
subst_psubst_idEnv updEnv_preserves_good)
lemma substAbs_vsubstAbs_compose2:
assumes "goodAbs A" and "good Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((A $[y1 // y]_ys) $[Z / z]_zs) = ((A $[Z / z]_zs) $[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms unfolding vsubstAbs_def by(simp add: substAbs_compose2)
lemma vsubstAbs_substAbs_compose2:
assumes "goodAbs A" and "good Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((A $[Y / y]_ys) $[z1 // z]_zs) = ((A $[z1 // z]_zs) $[(Y #[z1 // z]_zs) / y]_ys)"
using assms unfolding vsubstAbs_def vsubst_def by(simp add: substAbs_compose2)
lemma vsubstAbs_compose2:
assumes "goodAbs A"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((A $[y1 // y]_ys) $[z1 // z]_zs) =
((A $[z1 // z]_zs) $[(y1 @ys[z1 / z]_zs) // y]_ys)"
unfolding vsubstAbs_def
by (smt Var_preserves_good assms fresh_Var_simp insertCI
substAbs_compose2 vsubst_Var_simp vsubst_def)
-text{* Properties specific to variable-for-variable substitution: *}
+text\<open>Properties specific to variable-for-variable substitution:\<close>
corollary vsubstAbs_ident[simp]:
assumes "goodAbs A"
shows "(A $[z // z]_zs) = A"
using assms vsubstAll_ident[of "Par [z] [] [] []" _ _ A]
unfolding goodPar_def by simp
corollary substAbs_ident[simp]:
assumes "goodAbs A"
shows "(A $[(Var zs z) / z]_zs) = A"
using assms vsubstAbs_ident unfolding vsubstAbs_def by auto
corollary vsubstAbs_eq_swapAbs:
assumes "goodAbs A" and "freshAbs ys y1 A"
shows "(A $[y1 // y2]_ys) = (A $[y1 \<and> y2]_ys)"
using assms vsubstAll_swapAll[of "Par [y1, y2] [] [] []" _ _ A]
unfolding goodPar_def by simp
corollary skelAbs_vsubstAbs:
assumes "goodAbs A"
shows "skelAbs (A $[y1 // y2]_ys) = skelAbs A"
using assms skelAll_vsubstAll[of "Par [y1, y2] [] [] []" _ _ A]
unfolding goodPar_def by simp
lemma substAbs_vsubstAbs_trans:
assumes "goodAbs A" and "good Y" and "freshAbs ys y1 A"
shows "((A $[y1 // y]_ys) $[Y / y1]_ys) = (A $[Y / y]_ys)"
using assms unfolding substAbs_def vsubstAbs_def
by (cases "y1 = y") (auto simp: freshAbs_psubstAbs_updEnv psubstAbs_compose
psubstEnv_updEnv_idEnv updEnv_commute)
lemma vsubstAbs_trans:
assumes "goodAbs A" and "freshAbs ys y1 A"
shows "((A $[y1 // y]_ys) $[y2 // y1]_ys) = (A $[y2 // y]_ys)"
unfolding vsubstAbs_def[of _ y2 y1] vsubstAbs_def[of _ y2 y]
using assms by(simp add: substAbs_vsubstAbs_trans)
lemmas good_psubstAll_freshAll_otherSimps =
psubst_idEnv psubstEnv_idEnv_id psubstAbs_idEnv
freshEnv_psubst_ident freshEnv_psubstAbs_ident
lemmas good_substAll_freshAll_otherSimps =
fresh_fresh_subst fresh_subst_ident fresh_substEnv_updEnv subst_ident
fresh_freshAbs_substAbs freshAbs_substAbs_ident substAbs_ident
lemmas good_vsubstAll_freshAll_otherSimps =
diff_fresh_vsubst fresh_vsubst_ident fresh_vsubstEnv_updEnv vsubst_ident
diff_freshAbs_vsubstAbs freshAbs_vsubstAbs_ident vsubstAbs_ident
lemmas good_allOpers_otherSimps =
good_swapAll_freshAll_otherSimps
good_psubstAll_freshAll_otherSimps
good_substAll_freshAll_otherSimps
good_vsubstAll_freshAll_otherSimps
lemmas good_item_simps =
param_simps
all_preserve_good
good_freeCons
good_allOpers_simps
good_allOpers_otherSimps
end (* context FixVars *)
end
diff --git a/thys/Binding_Syntax_Theory/Transition_QuasiTerms_Terms.thy b/thys/Binding_Syntax_Theory/Transition_QuasiTerms_Terms.thy
--- a/thys/Binding_Syntax_Theory/Transition_QuasiTerms_Terms.thy
+++ b/thys/Binding_Syntax_Theory/Transition_QuasiTerms_Terms.thy
@@ -1,2801 +1,2801 @@
-section {* Transition from Quasi-Terms to Terms*}
+section \<open>Transition from Quasi-Terms to Terms\<close>
theory Transition_QuasiTerms_Terms
imports QuasiTerms_Environments_Substitution Equiv_Relation2
begin
-text{* This section transits from quasi-terms to terms: defines terms as alpha-equivalence
+text\<open>This section transits from quasi-terms to terms: defines terms as alpha-equivalence
classes of quasi-terms
(and also abstractions as alpha-equivalence classes of quasi-abstractions),
then defines operators on terms corresponding to those on quasi-terms:
variable injection, binding operation, freshness, swapping, parallel substitution, etc.
Properties previously shown invariant
under alpha-equivalence, including induction principles, are lifted from quasi-terms.
Moreover, a new powerful induction principle, allowing freshness assumptions,
is proved for terms.
As a matter of notation:
Starting from this section, we change the notations for quasi-item meta-variables, prefixing
their names with a "q" -- e.g., qX, qA, qinp, qenv, etc. The old names are now assigned
-to the ``real" items: terms, abstractions, inputs, environments. *}
-
-subsection {* Preparation: Integrating quasi-inputs as first-class citizens *}
+to the ``real" items: terms, abstractions, inputs, environments.\<close>
+
+subsection \<open>Preparation: Integrating quasi-inputs as first-class citizens\<close>
context FixVars
begin
-text{* From now on it will be convenient to
- also define fresh, swap, good and alpha-equivalence for quasi-inpus. *}
+text\<open>From now on it will be convenient to
+ also define fresh, swap, good and alpha-equivalence for quasi-inpus.\<close>
definition qSwapInp where
"qSwapInp xs x y qinp == lift (qSwap xs x y) qinp"
definition qSwapBinp where
"qSwapBinp xs x y qbinp == lift (qSwapAbs xs x y) qbinp"
abbreviation qSwapInp_abbrev ("_ %[[_ \<and> _]]'__" 200) where
"(qinp %[[z1 \<and> z2]]_zs) == qSwapInp zs z1 z2 qinp"
abbreviation qSwapBinp_abbrev ("_ %%[[_ \<and> _]]'__" 200) where
"(qbinp %%[[z1 \<and> z2]]_zs) == qSwapBinp zs z1 z2 qbinp"
lemma qSwap_qSwapInp:
"((qOp delta qinp qbinp) #[[x \<and> y]]_xs) =
qOp delta (qinp %[[x \<and> y]]_xs) (qbinp %%[[x \<and> y]]_xs)"
unfolding qSwapInp_def qSwapBinp_def by simp
(* For the qOp case, qSwap shall henceforth simplify to qSwapInp: *)
declare qSwap.simps(2) [simp del]
declare qSwap_qSwapInp[simp]
(* and qSwap_simps and qSwapAll_simps, rather than qSwap.simps and qSwapAll.simps,
shall refer to the simplification rules for qSwap *)
lemmas qSwapAll_simps = qSwap.simps(1) qSwap_qSwapInp
definition qPsubstInp where
"qPsubstInp qrho qinp == lift (qPsubst qrho) qinp"
definition qPsubstBinp where
"qPsubstBinp qrho qbinp == lift (qPsubstAbs qrho) qbinp"
abbreviation qPsubstInp_abbrev ("_ %[[_]]" 200)
where "(qinp %[[qrho]]) == qPsubstInp qrho qinp"
abbreviation qPsubstBinp_abbrev ("_ %%[[_]]" 200)
where "(qbinp %%[[qrho]]) == qPsubstBinp qrho qbinp"
lemma qPsubst_qPsubstInp:
"((qOp delta qinp qbinp) #[[rho]]) = qOp delta (qinp %[[rho]]) (qbinp %%[[rho]])"
unfolding qPsubstInp_def qPsubstBinp_def by simp
(* For the qOp case, qPsubst shall henceforth simplify to qPsubstInp: *)
declare qPsubst.simps(2) [simp del]
declare qPsubst_qPsubstInp[simp]
(* and qPsubst_simps and qPsubstAll_simps, rather than qPsubst.simps and qPsubstAll.simps,
shall refer to the simplification rules for qPsubst *)
lemmas qPsubstAll_simps = qPsubst.simps(1) qPsubst_qPsubstInp
definition qSkelInp
where "qSkelInp qinp = lift qSkel qinp"
definition qSkelBinp
where "qSkelBinp qbinp = lift qSkelAbs qbinp"
lemma qSkel_qSkelInp:
"qSkel (qOp delta qinp qbinp) =
Branch (qSkelInp qinp) (qSkelBinp qbinp)"
unfolding qSkelInp_def qSkelBinp_def by simp
(* For the qOp case, qSkel shall henceforth simplify to qSkelInp: *)
declare qSkel.simps(2) [simp del]
declare qSkel_qSkelInp[simp]
(* and qSkel_simps and qSkelAll_simps, rather than qSkel.simps and qSkelAll.simps,
shall refer to the simplification rules for qSkel *)
lemmas qSkelAll_simps = qSkel.simps(1) qSkel_qSkelInp
definition qFreshInp ::
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,('index,'bindex,'varSort,'var,'opSym)qTerm)input \<Rightarrow> bool"
where
"qFreshInp xs x qinp == liftAll (qFresh xs x) qinp"
definition qFreshBinp ::
"'varSort \<Rightarrow> 'var \<Rightarrow> ('bindex,('index,'bindex,'varSort,'var,'opSym)qAbs)input \<Rightarrow> bool"
where
"qFreshBinp xs x qbinp == liftAll (qFreshAbs xs x) qbinp"
lemma qFresh_qFreshInp:
"qFresh xs x (qOp delta qinp qbinp) =
(qFreshInp xs x qinp \<and> qFreshBinp xs x qbinp)"
unfolding qFreshInp_def qFreshBinp_def by simp
(* For the qOp case, qFresh shall henceforth simplify to qFreshInp: *)
declare qFresh.simps(2) [simp del]
declare qFresh_qFreshInp[simp]
(* and qFresh_simps and qFreshAll_simps, rather than qFresh.simps and qFreshAll.simps,
shall refer to the simplification rules for qFresh *)
lemmas qFreshAll_simps = qFresh.simps(1) qFresh_qFreshInp
definition qGoodInp where
"qGoodInp qinp ==
liftAll qGood qinp \<and>
|{i. qinp i \<noteq> None}| <o |UNIV :: 'var set|"
definition qGoodBinp where
"qGoodBinp qbinp ==
liftAll qGoodAbs qbinp \<and>
|{i. qbinp i \<noteq> None}| <o |UNIV :: 'var set|"
lemma qGood_qGoodInp:
"qGood (qOp delta qinp qbinp) = (qGoodInp qinp \<and> qGoodBinp qbinp)"
unfolding qGoodInp_def qGoodBinp_def by auto
(* For the qOp case, qGood shall henceforth simplify to qGoodInp: *)
declare qGood.simps(2) [simp del]
declare qGood_qGoodInp [simp]
(* and qGood_simps (and qGoodAll_simps), rather than qGood.simps,
shall refer to the simplification rules for qGood *)
lemmas qGoodAll_simps = qGood.simps(1) qGood_qGoodInp
definition alphaInp where
"alphaInp ==
{(qinp,qinp'). sameDom qinp qinp' \<and> liftAll2 (\<lambda>qX qX'. qX #= qX') qinp qinp'}"
definition alphaBinp where
"alphaBinp ==
{(qbinp,qbinp'). sameDom qbinp qbinp' \<and> liftAll2 (\<lambda>qA qA'. qA $= qA') qbinp qbinp'}"
abbreviation alphaInp_abbrev (infix "%=" 50) where
"qinp %= qinp' == (qinp,qinp') \<in> alphaInp"
abbreviation alphaBinp_abbrev (infix "%%=" 50) where
"qbinp %%= qbinp' == (qbinp,qbinp') \<in> alphaBinp"
lemma alpha_alphaInp:
"(qOp delta qinp qbinp #= qOp delta' qinp' qbinp') =
(delta = delta' \<and> qinp %= qinp' \<and> qbinp %%= qbinp')"
unfolding alphaInp_def alphaBinp_def by auto
(* For the qOp case, alpha shall henceforth simplify to alphaInp: *)
declare alpha.simps(2) [simp del]
declare alpha_alphaInp[simp]
(* and alpha_Simps and alphaAll_Simps, rather than alpha_simps and alphaAll_simps,
shall refer to the simplification rules for alpha *)
lemmas alphaAll_Simps =
alpha.simps(1) alpha_alphaInp
alphaAbs.simps
lemma alphaInp_refl:
"qGoodInp qinp \<Longrightarrow> qinp %= qinp"
using alpha_refl
unfolding alphaInp_def qGoodInp_def liftAll_def liftAll2_def sameDom_def
by fastforce
lemma alphaBinp_refl:
"qGoodBinp qbinp \<Longrightarrow> qbinp %%= qbinp"
using alphaAbs_refl
unfolding alphaBinp_def qGoodBinp_def liftAll_def liftAll2_def sameDom_def
by fastforce
lemma alphaInp_sym:
fixes qinp qinp' :: "('index,('index,'bindex,'varSort,'var,'opSym)qTerm)input"
shows "qinp %= qinp' \<Longrightarrow> qinp' %= qinp"
using alpha_sym unfolding alphaInp_def sameDom_def liftAll2_def by blast
lemma alphaBinp_sym:
fixes qbinp qbinp' :: "('bindex,('index,'bindex,'varSort,'var,'opSym)qAbs)input"
shows "qbinp %%= qbinp' \<Longrightarrow> qbinp' %%= qbinp"
using alphaAbs_sym unfolding alphaBinp_def sameDom_def liftAll2_def by blast
lemma alphaInp_trans:
assumes good: "qGoodInp qinp" and
alpha1: "qinp %= qinp'" and alpha2: "qinp' %= qinp''"
shows "qinp %= qinp''"
proof-
{fix i qX qX'' assume qinp: "qinp i = Some qX" and qinp'': "qinp'' i = Some qX''"
then obtain qX' where qinp': "qinp' i = Some qX'"
using alpha1 unfolding alphaInp_def sameDom_def liftAll2_def by(cases "qinp' i", force)
hence "qX #= qX'"
using alpha1 qinp unfolding alphaInp_def sameDom_def liftAll2_def by auto
moreover have "qX' #= qX''" using alpha2 qinp' qinp''
unfolding alphaInp_def sameDom_def liftAll2_def by auto
moreover have "qGood qX" using good qinp unfolding qGoodInp_def liftAll_def by auto
ultimately have "qX #= qX''" using alpha_trans by blast
}
thus ?thesis using assms unfolding alphaInp_def sameDom_def liftAll2_def by auto
qed
lemma alphaBinp_trans:
assumes good: "qGoodBinp qbinp" and
alpha1: "qbinp %%= qbinp'" and alpha2: "qbinp' %%= qbinp''"
shows "qbinp %%= qbinp''"
proof-
{fix i qA qA'' assume qbinp: "qbinp i = Some qA" and qbinp'': "qbinp'' i = Some qA''"
then obtain qA' where qbinp': "qbinp' i = Some qA'"
using alpha1 unfolding alphaBinp_def sameDom_def liftAll2_def by(cases "qbinp' i", force)
hence "qA $= qA'"
using alpha1 qbinp unfolding alphaBinp_def sameDom_def liftAll2_def by auto
moreover have "qA' $= qA''" using alpha2 qbinp' qbinp''
unfolding alphaBinp_def sameDom_def liftAll2_def by auto
moreover have "qGoodAbs qA" using good qbinp unfolding qGoodBinp_def liftAll_def by auto
ultimately have "qA $= qA''" using alphaAbs_trans by blast
}
thus ?thesis using assms unfolding alphaBinp_def sameDom_def liftAll2_def by auto
qed
lemma qSwapInp_preserves_qGoodInp:
assumes "qGoodInp qinp"
shows "qGoodInp (qinp %[[x1 \<and> x2]]_xs)"
proof-
{let ?qinp' = "lift (qSwap xs x1 x2) qinp"
fix xsa let ?Left = "{i. ?qinp' i \<noteq> None}"
have "?Left = {i. qinp i \<noteq> None}" by(auto simp add: lift_None)
hence "|?Left| <o |UNIV :: 'var set|" using assms unfolding qGoodInp_def by auto
}
thus ?thesis using assms
unfolding qGoodInp_def qSwapInp_def liftAll_lift_comp qGoodInp_def
unfolding comp_def liftAll_def
by (auto simp add: qSwap_preserves_qGood simp del: not_None_eq)
qed
lemma qSwapBinp_preserves_qGoodBinp:
assumes "qGoodBinp qbinp"
shows "qGoodBinp (qbinp %%[[x1 \<and> x2]]_xs)"
proof-
{let ?qbinp' = "lift (qSwapAbs xs x1 x2) qbinp"
fix xsa let ?Left = "{i. ?qbinp' i \<noteq> None}"
have "?Left = {i. qbinp i \<noteq> None}" by(auto simp add: lift_None)
hence "|?Left| <o |UNIV :: 'var set|" using assms unfolding qGoodBinp_def by auto
}
thus ?thesis using assms
unfolding qGoodBinp_def qSwapBinp_def liftAll_lift_comp
unfolding qGoodBinp_def unfolding comp_def liftAll_def
by (auto simp add: qSwapAbs_preserves_qGoodAbs simp del: not_None_eq)
qed
lemma qSwapInp_preserves_alphaInp:
assumes "qGoodInp qinp \<or> qGoodInp qinp'" and "qinp %= qinp'"
shows "(qinp %[[x1 \<and> x2]]_xs) %= (qinp' %[[x1 \<and> x2]]_xs)"
using assms unfolding alphaInp_def qSwapInp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt liftAll_def lift_def option.case_eq_if option.exhaust_sel
option.sel qGoodInp_def qSwap_preserves_alpha)
lemma qSwapBinp_preserves_alphaBinp:
assumes "qGoodBinp qbinp \<or> qGoodBinp qbinp'" and "qbinp %%= qbinp'"
shows "(qbinp %%[[x1 \<and> x2]]_xs) %%= (qbinp' %%[[x1 \<and> x2]]_xs)"
using assms unfolding alphaBinp_def qSwapBinp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt liftAll_def lift_def option.case_eq_if option.exhaust_sel option.sel
qGoodBinp_def qSwapAbs_preserves_alphaAbs)
lemma qPsubstInp_preserves_qGoodInp:
assumes "qGoodInp qinp" and "qGoodEnv qrho"
shows "qGoodInp (qinp %[[qrho]])"
using assms unfolding qGoodInp_def qPsubstInp_def liftAll_def
by simp (smt Collect_cong lift_def option.case_eq_if
option.exhaust_sel option.sel qPsubst_preserves_qGood)
lemma qPsubstBinp_preserves_qGoodBinp:
assumes "qGoodBinp qbinp" and "qGoodEnv qrho"
shows "qGoodBinp (qbinp %%[[qrho]])"
using assms unfolding qGoodBinp_def qPsubstBinp_def liftAll_def
by simp (smt Collect_cong lift_def option.case_eq_if
option.exhaust_sel option.sel qPsubstAbs_preserves_qGoodAbs)
lemma qPsubstInp_preserves_alphaInp:
assumes "qGoodInp qinp \<or> qGoodInp qinp'" and "qGoodEnv qrho" and "qinp %= qinp'"
shows "(qinp %[[qrho]]) %= (qinp' %[[qrho]])"
using assms unfolding alphaInp_def qPsubstInp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt liftAll_def lift_def option.case_eq_if option.exhaust_sel
option.sel qGoodInp_def qPsubst_preserves_alpha1)
lemma qPsubstBinp_preserves_alphaBinp:
assumes "qGoodBinp qbinp \<or> qGoodBinp qbinp'" and "qGoodEnv qrho" and "qbinp %%= qbinp'"
shows "(qbinp %%[[qrho]]) %%= (qbinp' %%[[qrho]])"
using assms unfolding alphaBinp_def qPsubstBinp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt liftAll_def lift_def option.case_eq_if option.exhaust_sel
option.sel qGoodBinp_def qPsubstAbs_preserves_alphaAbs1)
lemma qFreshInp_preserves_alphaInp_aux:
assumes good: "qGoodInp qinp \<or> qGoodInp qinp'" and alpha: "qinp %= qinp'"
and fresh: "qFreshInp xs x qinp"
shows "qFreshInp xs x qinp'"
using assms unfolding qFreshInp_def liftAll_def proof clarify
fix i qX' assume qinp': "qinp' i = Some qX'"
then obtain qX where qinp: "qinp i = Some qX"
using alpha unfolding alphaInp_def sameDom_def liftAll2_def by (cases "qinp i", auto)
hence "qGood qX \<or> qGood qX'"
using qinp' good unfolding qGoodInp_def liftAll_def by auto
moreover have "qX #= qX'"
using qinp qinp' alpha unfolding alphaInp_def sameDom_def liftAll2_def by auto
moreover have "qFresh xs x qX"
using fresh qinp unfolding qFreshInp_def liftAll_def by simp
ultimately show "qFresh xs x qX'"
using qFresh_preserves_alpha by auto
qed
lemma qFreshBinp_preserves_alphaBinp_aux:
assumes good: "qGoodBinp qbinp \<or> qGoodBinp qbinp'" and alpha: "qbinp %%= qbinp'"
and fresh: "qFreshBinp xs x qbinp"
shows "qFreshBinp xs x qbinp'"
using assms unfolding qFreshBinp_def liftAll_def proof clarify
fix i qA' assume qbinp': "qbinp' i = Some qA'"
then obtain qA where qbinp: "qbinp i = Some qA"
using alpha unfolding alphaBinp_def sameDom_def liftAll2_def by (cases "qbinp i", auto)
hence "qGoodAbs qA \<or> qGoodAbs qA'"
using qbinp' good unfolding qGoodBinp_def liftAll_def by auto
moreover have "qA $= qA'"
using qbinp qbinp' alpha unfolding alphaBinp_def sameDom_def liftAll2_def by auto
moreover have "qFreshAbs xs x qA"
using fresh qbinp unfolding qFreshBinp_def liftAll_def by simp
ultimately show "qFreshAbs xs x qA'"
using qFreshAbs_preserves_alphaAbs by auto
qed
lemma qFreshInp_preserves_alphaInp:
assumes "qGoodInp qinp \<or> qGoodInp qinp'" and "qinp %= qinp'"
shows "qFreshInp xs x qinp \<longleftrightarrow> qFreshInp xs x qinp'"
using alphaInp_sym assms qFreshInp_preserves_alphaInp_aux by blast
lemma qFreshBinp_preserves_alphaBinp:
assumes "qGoodBinp qbinp \<or> qGoodBinp qbinp'" and "qbinp %%= qbinp'"
shows "qFreshBinp xs x qbinp \<longleftrightarrow> qFreshBinp xs x qbinp'"
using alphaBinp_sym assms qFreshBinp_preserves_alphaBinp_aux by blast
(****************************************************)
lemmas qItem_simps =
qSkelAll_simps qFreshAll_simps qSwapAll_simps qPsubstAll_simps qGoodAll_simps alphaAll_Simps
qSwap_qAFresh_otherSimps qAFresh.simps qGoodItem.simps
(****************************************************)
end (* context FixVars *)
-subsection {* Definitions of terms and their operators *}
+subsection \<open>Definitions of terms and their operators\<close>
type_synonym ('index,'bindex,'varSort,'var,'opSym)"term" =
"('index,'bindex,'varSort,'var,'opSym)qTerm set"
type_synonym ('index,'bindex,'varSort,'var,'opSym)abs =
"('index,'bindex,'varSort,'var,'opSym)qAbs set"
type_synonym ('index,'bindex,'varSort,'var,'opSym)env =
"'varSort \<Rightarrow> 'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term option"
-text{* A ``parameter" will be something for which
+text\<open>A ``parameter" will be something for which
freshness makes sense. Here is the most typical case of a parameter in proofs, putting
-together (as lists) finite collections of variables, terms, abstractions and environments: *}
+together (as lists) finite collections of variables, terms, abstractions and environments:\<close>
datatype ('index,'bindex,'varSort,'var,'opSym)param =
Par "'var list"
"('index,'bindex,'varSort,'var,'opSym)term list"
"('index,'bindex,'varSort,'var,'opSym)abs list"
"('index,'bindex,'varSort,'var,'opSym)env list"
fun varsOf where
"varsOf (Par xL _ _ _) = set xL"
fun termsOf where
"termsOf (Par _ XL _ _) = set XL"
fun absOf where
"absOf (Par _ _ AL _) = set AL"
fun envsOf where
"envsOf (Par _ _ _ rhoL) = set rhoL"
context FixVars (* scope all throughout the file *)
begin
(* Recall the abbreviation "Restr r qA" for "r Int (qA <*> qA)" *)
definition "alphaGood \<equiv> \<lambda> qX qY. qGood qX \<and> qGood qY \<and> qX #= qY"
definition "alphaAbsGood \<equiv> \<lambda> qA qB. qGoodAbs qA \<and> qGoodAbs qB \<and> qA $= qB"
definition "good \<equiv> qGood /// alphaGood"
definition "goodAbs \<equiv> qGoodAbs /// alphaAbsGood"
definition goodInp where
"goodInp inp ==
liftAll good inp \<and>
|{i. inp i \<noteq> None}| <o |UNIV :: 'var set|"
definition goodBinp where
"goodBinp binp ==
liftAll goodAbs binp \<and>
|{i. binp i \<noteq> None}| <o |UNIV :: 'var set|"
definition goodEnv where
"goodEnv rho ==
(\<forall> ys. liftAll good (rho ys)) \<and>
(\<forall> ys. |{y. rho ys y \<noteq> None}| <o |UNIV :: 'var set| )"
definition asTerm where
"asTerm qX \<equiv> proj alphaGood qX"
definition asAbs where
"asAbs qA \<equiv> proj alphaAbsGood qA"
definition pickInp where
"pickInp inp \<equiv> lift pick inp"
definition pickBinp where
"pickBinp binp \<equiv> lift pick binp"
(* Note: pickInp and pickBinp are the same (polymorphically), but
I keep distinct notations for uniformity with the rest of the notations. *)
definition asInp where
"asInp qinp \<equiv> lift asTerm qinp"
definition asBinp where
"asBinp qbinp \<equiv> lift asAbs qbinp"
definition pickE where
"pickE rho \<equiv> \<lambda> xs. lift pick (rho xs)"
definition asEnv where
"asEnv qrho \<equiv> \<lambda> xs. lift asTerm (qrho xs)"
definition Var where
"Var xs x \<equiv> asTerm(qVar xs x)"
definition Op where
"Op delta inp binp \<equiv> asTerm (qOp delta (pickInp inp) (pickBinp binp))"
definition Abs where
"Abs xs x X \<equiv> asAbs (qAbs xs x (pick X))"
definition skel where
"skel X \<equiv> qSkel (pick X)"
definition skelAbs where
"skelAbs A \<equiv> qSkelAbs (pick A)"
definition skelInp where
"skelInp inp = qSkelInp (pickInp inp)"
definition skelBinp where
"skelBinp binp = qSkelBinp (pickBinp binp)"
lemma skelInp_def2:
assumes "goodInp inp"
shows "skelInp inp = lift skel inp"
unfolding skelInp_def
unfolding qSkelInp_def pickInp_def skel_def[abs_def]
unfolding lift_comp comp_def by simp
lemma skelBinp_def2:
assumes "goodBinp binp"
shows "skelBinp binp = lift skelAbs binp"
unfolding skelBinp_def
unfolding qSkelBinp_def pickBinp_def skelAbs_def[abs_def]
unfolding lift_comp comp_def by simp
definition swap where
"swap xs x y X = asTerm (qSwap xs x y (pick X))"
abbreviation swap_abbrev ("_ #[_ \<and> _]'__" 200) where
"(X #[z1 \<and> z2]_zs) \<equiv> swap zs z1 z2 X"
definition swapAbs where
"swapAbs xs x y A = asAbs (qSwapAbs xs x y (pick A))"
abbreviation swapAbs_abbrev ("_ $[_ \<and> _]'__" 200) where
"(A $[z1 \<and> z2]_zs) \<equiv> swapAbs zs z1 z2 A"
definition swapInp where
"swapInp xs x y inp \<equiv> lift (swap xs x y) inp"
definition swapBinp where
"swapBinp xs x y binp \<equiv> lift (swapAbs xs x y) binp"
abbreviation swapInp_abbrev ("_ %[_ \<and> _]'__" 200) where
"(inp %[z1 \<and> z2]_zs) \<equiv> swapInp zs z1 z2 inp"
abbreviation swapBinp_abbrev ("_ %%[_ \<and> _]'__" 200) where
"(binp %%[z1 \<and> z2]_zs) \<equiv> swapBinp zs z1 z2 binp"
definition swapEnvDom where
"swapEnvDom xs x y rho \<equiv> \<lambda>zs z. rho zs (z @zs[x \<and> y]_xs)"
definition swapEnvIm where
"swapEnvIm xs x y rho \<equiv> \<lambda>zs. lift (swap xs x y) (rho zs)"
definition swapEnv where
"swapEnv xs x y \<equiv> swapEnvIm xs x y o swapEnvDom xs x y"
abbreviation swapEnv_abbrev ("_ &[_ \<and> _]'__" 200) where
"(rho &[z1 \<and> z2]_zs) \<equiv> swapEnv zs z1 z2 rho"
lemmas swapEnv_defs = swapEnv_def comp_def swapEnvDom_def swapEnvIm_def
inductive_set swapped where
Refl: "(X,X) \<in> swapped"
|
Trans: "\<lbrakk>(X,Y) \<in> swapped; (Y,Z) \<in> swapped\<rbrakk> \<Longrightarrow> (X,Z) \<in> swapped"
|
Swap: "(X,Y) \<in> swapped \<Longrightarrow> (X, Y #[x \<and> y]_zs) \<in> swapped"
lemmas swapped_Clauses = swapped.Refl swapped.Trans swapped.Swap
definition fresh where
"fresh xs x X \<equiv> qFresh xs x (pick X)"
definition freshAbs where
"freshAbs xs x A \<equiv> qFreshAbs xs x (pick A)"
definition freshInp where
"freshInp xs x inp \<equiv> liftAll (fresh xs x) inp"
definition freshBinp where
"freshBinp xs x binp \<equiv> liftAll (freshAbs xs x) binp"
definition freshEnv where
"freshEnv xs x rho ==
rho xs x = None \<and> (\<forall> ys. liftAll (fresh xs x) (rho ys))"
definition psubst where
"psubst rho X \<equiv> asTerm(qPsubst (pickE rho) (pick X))"
abbreviation psubst_abbrev ("_ #[_]") where
"(X #[rho]) \<equiv> psubst rho X"
definition psubstAbs where
"psubstAbs rho A \<equiv> asAbs(qPsubstAbs (pickE rho) (pick A))"
abbreviation psubstAbs_abbrev ("_ $[_]") where
"A $[rho] \<equiv> psubstAbs rho A"
definition psubstInp where
"psubstInp rho inp \<equiv> lift (psubst rho) inp"
definition psubstBinp where
"psubstBinp rho binp \<equiv> lift (psubstAbs rho) binp"
abbreviation psubstInp_abbrev ("_ %[_]") where
"inp %[rho] \<equiv> psubstInp rho inp"
abbreviation psubstBinp_abbrev ("_ %%[_]") where
"binp %%[rho] \<equiv> psubstBinp rho binp"
definition psubstEnv where
"psubstEnv rho rho' \<equiv>
\<lambda> xs x. case rho' xs x of None \<Rightarrow> rho xs x
|Some X \<Rightarrow> Some (X #[rho])"
abbreviation psubstEnv_abbrev ("_ &[_]") where
"rho &[rho'] \<equiv> psubstEnv rho' rho"
definition idEnv where
"idEnv \<equiv> \<lambda>xs. Map.empty"
definition updEnv ::
"('index,'bindex,'varSort,'var,'opSym)env \<Rightarrow>
'var \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> 'varSort \<Rightarrow>
('index,'bindex,'varSort,'var,'opSym)env"
("_ [_ \<leftarrow> _]'__" 200) where
"(rho [x \<leftarrow> X]_xs) \<equiv> \<lambda> ys y. (if ys = xs \<and> y = x then Some X else rho ys y)"
-text{* (Unary) substitution: *}
+text\<open>(Unary) substitution:\<close>
definition subst where
"subst xs X x \<equiv> psubst (idEnv [x \<leftarrow> X]_xs)"
abbreviation subst_abbrev ("_ #[_ '/ _]'__" 200) where
"(Y #[X / x]_xs) \<equiv> subst xs X x Y"
definition substAbs where
"substAbs xs X x \<equiv> psubstAbs (idEnv [x \<leftarrow> X]_xs)"
abbreviation substAbs_abbrev ("_ $[_ '/ _]'__" 200) where
"(A $[X / x]_xs) \<equiv> substAbs xs X x A"
definition substInp where
"substInp xs X x \<equiv> psubstInp (idEnv [x \<leftarrow> X]_xs)"
definition substBinp where
"substBinp xs X x \<equiv> psubstBinp (idEnv [x \<leftarrow> X]_xs)"
abbreviation substInp_abbrev ("_ %[_ '/ _]'__" 200) where
"(inp %[X / x]_xs) \<equiv> substInp xs X x inp"
abbreviation substBinp_abbrev ("_ %%[_ '/ _]'__" 200) where
"(binp %%[X / x]_xs) \<equiv> substBinp xs X x binp"
theorem substInp_def2:
"substInp ys Y y = lift (subst ys Y y)"
unfolding substInp_def[abs_def] subst_def psubstInp_def[abs_def] by simp
theorem substBinp_def2:
"substBinp ys Y y = lift (substAbs ys Y y)"
unfolding substBinp_def[abs_def] substAbs_def psubstBinp_def[abs_def] by simp
definition substEnv where
"substEnv xs X x \<equiv> psubstEnv (idEnv [x \<leftarrow> X]_xs)"
abbreviation substEnv_abbrev ("_ &[_ '/ _]'__" 200) where
"(Y &[X / x]_xs) \<equiv> substEnv xs X x Y"
theorem substEnv_def2:
"(rho &[Y / y]_ys) =
(\<lambda>xs x. case rho xs x of
None \<Rightarrow> if (xs = ys \<and> x = y) then Some Y else None
|Some X \<Rightarrow> Some (X #[Y / y]_ys))"
unfolding substEnv_def psubstEnv_def subst_def idEnv_def updEnv_def
apply(rule ext)+ by(case_tac "rho xs x", simp_all)
-text{* Variable-for-variable substitution: *}
+text\<open>Variable-for-variable substitution:\<close>
definition vsubst where
"vsubst ys y1 y2 \<equiv> subst ys (Var ys y1) y2"
abbreviation vsubst_abbrev ("_ #[_ '/'/ _]'__" 200) where
"(X #[y1 // y2]_ys) \<equiv> vsubst ys y1 y2 X"
definition vsubstAbs where
"vsubstAbs ys y1 y2 \<equiv> substAbs ys (Var ys y1) y2"
abbreviation vsubstAbs_abbrev ("_ $[_ '/'/ _]'__" 200) where
"(A $[y1 // y2]_ys) \<equiv> vsubstAbs ys y1 y2 A"
definition vsubstInp where
"vsubstInp ys y1 y2 \<equiv> substInp ys (Var ys y1) y2"
definition vsubstBinp where
"vsubstBinp ys y1 y2 \<equiv> substBinp ys (Var ys y1) y2"
abbreviation vsubstInp_abbrev ("_ %[_ '/'/ _]'__" 200) where
"(inp %[y1 // y2]_ys) \<equiv> vsubstInp ys y1 y2 inp"
abbreviation vsubstBinp_abbrev ("_ %%[_ '/'/ _]'__" 200) where
"(binp %%[y1 // y2]_ys) \<equiv> vsubstBinp ys y1 y2 binp"
lemma vsubstInp_def2:
"(inp %[y1 // y2]_ys) = lift (vsubst ys y1 y2) inp"
unfolding vsubstInp_def vsubst_def
by(auto simp add: substInp_def2)
lemma vsubstBinp_def2:
"(binp %%[y1 // y2]_ys) = lift (vsubstAbs ys y1 y2) binp"
unfolding vsubstBinp_def vsubstAbs_def
by(auto simp add: substBinp_def2)
definition vsubstEnv where
"vsubstEnv ys y1 y2 \<equiv> substEnv ys (Var ys y1) y2"
abbreviation vsubstEnv_abbrev ("_ &[_ '/'/ _]'__" 200) where
"(rho &[y1 // y2]_ys) \<equiv> vsubstEnv ys y1 y2 rho"
theorem vsubstEnv_def2:
"(rho &[y1 // y]_ys) =
(\<lambda>xs x. case rho xs x of
None \<Rightarrow> if (xs = ys \<and> x = y) then Some (Var ys y1) else None
|Some X \<Rightarrow> Some (X #[y1 // y]_ys))"
unfolding vsubstEnv_def vsubst_def by(auto simp add: substEnv_def2)
definition goodPar where
"goodPar P \<equiv> (\<forall> X \<in> termsOf P. good X) \<and>
(\<forall> A \<in> absOf P. goodAbs A) \<and>
(\<forall> rho \<in> envsOf P. goodEnv rho)"
lemma Par_preserves_good[simp]:
assumes "!! X. X \<in> set XL \<Longrightarrow> good X"
and "!! A. A \<in> set AL \<Longrightarrow> goodAbs A"
and "!! rho. rho \<in> set rhoL \<Longrightarrow> goodEnv rho"
shows "goodPar (Par xL XL AL rhoL)"
using assms unfolding goodPar_def by auto
lemma termsOf_preserves_good[simp]:
assumes "goodPar P" and "X : termsOf P"
shows "good X"
using assms unfolding goodPar_def by auto
lemma absOf_preserves_good[simp]:
assumes "goodPar P" and "A : absOf P"
shows "goodAbs A"
using assms unfolding goodPar_def by auto
lemma envsOf_preserves_good[simp]:
assumes "goodPar P" and "rho : envsOf P"
shows "goodEnv rho"
using assms unfolding goodPar_def by blast
lemmas param_simps =
termsOf.simps absOf.simps envsOf.simps
Par_preserves_good
termsOf_preserves_good absOf_preserves_good envsOf_preserves_good
-subsection {* Items versus quasi-items modulo alpha *}
-
-text{* Here we ``close the accounts" (for a while) with quasi-items --
+subsection \<open>Items versus quasi-items modulo alpha\<close>
+
+text\<open>Here we ``close the accounts" (for a while) with quasi-items --
beyond this subsection, there will not be any theorem that mentions
quasi-items, except much later when we deal with iteration principles
(and need to briefly switch back to quasi-terms in order to define the needed
- iterative map by the universality of the alpha-quotient). *}
-
-subsubsection {* For terms *}
+ iterative map by the universality of the alpha-quotient).\<close>
+
+subsubsection \<open>For terms\<close>
lemma alphaGood_equivP: "equivP qGood alphaGood"
unfolding equivP_def reflP_def symP_def transP_def alphaGood_def
using alpha_refl alpha_sym alpha_trans by blast
lemma univ_asTerm_alphaGood[simp]:
assumes *: "congruentP alphaGood f" and **: "qGood X"
shows "univ f (asTerm X) = f X"
by (metis assms alphaGood_equivP asTerm_def univ_commute)
corollary univ_asTerm_alpha[simp]:
assumes *: "congruentP alpha f" and **: "qGood X"
shows "univ f (asTerm X) = f X"
apply(rule univ_asTerm_alphaGood)
using assms unfolding alphaGood_def congruentP_def by auto
lemma pick_inj_on_good: "inj_on pick (Collect good)"
unfolding good_def using alphaGood_equivP equivP_pick_inj_on by auto
lemma pick_injective_good[simp]:
"\<lbrakk>good X; good Y\<rbrakk> \<Longrightarrow> (pick X = pick Y) = (X = Y)"
using pick_inj_on_good unfolding inj_on_def by auto
lemma good_imp_qGood_pick:
"good X \<Longrightarrow> qGood (pick X)"
unfolding good_def
by (metis alphaGood_equivP equivP_pick_preserves)
lemma qGood_iff_good_asTerm:
"good (asTerm qX) = qGood qX"
unfolding good_def asTerm_def
using alphaGood_equivP proj_in_iff by fastforce
lemma pick_asTerm:
assumes "qGood qX"
shows "pick (asTerm qX) #= qX"
by (metis (full_types) alphaGood_def alphaGood_equivP asTerm_def assms pick_proj)
lemma asTerm_pick:
assumes "good X"
shows "asTerm (pick X) = X"
by (metis alphaGood_equivP asTerm_def assms good_def proj_pick)
lemma pick_alpha: "good X \<Longrightarrow> pick X #= pick X"
using good_imp_qGood_pick alpha_refl by auto
lemma alpha_imp_asTerm_equal:
assumes "qGood qX" and "qX #= qY"
shows "asTerm qX = asTerm qY"
proof-
have "alphaGood qX qY" unfolding alphaGood_def using assms
by (metis alpha_preserves_qGood)
thus ?thesis unfolding asTerm_def using alphaGood_equivP proj_iff
by (metis alpha_preserves_qGood1 assms)
qed
lemma asTerm_equal_imp_alpha:
assumes "qGood qX" and "asTerm qX = asTerm qY"
shows "qX #= qY"
by (metis alphaAll_sym alphaAll_trans assms pick_asTerm qGood_iff_good_asTerm)
lemma asTerm_equal_iff_alpha:
assumes "qGood qX \<or> qGood qY"
shows "(asTerm qX = asTerm qY) = (qX #= qY)"
by (metis alpha_imp_asTerm_equal alpha_sym asTerm_equal_imp_alpha assms)
lemma pick_alpha_iff_equal:
assumes "good X" and "good Y"
shows "pick X #= pick Y \<longleftrightarrow> X = Y"
by (metis asTerm_equal_iff_alpha asTerm_pick assms good_imp_qGood_pick)
lemma pick_swap_qSwap:
assumes "good X"
shows "pick (X #[x1 \<and> x2]_xs) #= ((pick X) #[[x1 \<and> x2]]_xs)"
by (metis assms good_imp_qGood_pick pick_asTerm qSwap_preserves_qGood1 swap_def)
lemma asTerm_qSwap_swap:
assumes "qGood qX"
shows "asTerm (qX #[[x1 \<and> x2]]_xs) = ((asTerm qX) #[x1 \<and> x2]_xs)"
by (simp add: alpha_imp_asTerm_equal alpha_sym assms local.swap_def
pick_asTerm qSwap_preserves_alpha qSwap_preserves_qGood1)
lemma fresh_asTerm_qFresh:
assumes "qGood qX"
shows "fresh xs x (asTerm qX) = qFresh xs x qX"
by (simp add: assms fresh_def pick_asTerm qFresh_preserves_alpha)
(* Note that fresh and skel commute with pick by definition, so we only need
to prove they commute with asTerm. *)
lemma skel_asTerm_qSkel:
assumes "qGood qX"
shows "skel (asTerm qX) = qSkel qX"
by (simp add: alpha_qSkel assms pick_asTerm skel_def)
lemma double_swap_qSwap:
assumes "good X"
shows "qGood (((pick X) #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs') \<and>
((X #[x \<and> y]_zs) #[x' \<and> y']_zs') = asTerm (((pick X) #[[x \<and> y]]_zs) #[[x' \<and> y']]_zs')"
by (simp add: asTerm_qSwap_swap assms
good_imp_qGood_pick local.swap_def qSwap_preserves_qGood1)
lemma fresh_swap_qFresh_qSwap:
assumes "good X"
shows "fresh xs x (X #[y1 \<and> y2]_ys) = qFresh xs x ((pick X) #[[y1 \<and> y2]]_ys)"
by (simp add: assms
fresh_asTerm_qFresh good_imp_qGood_pick local.swap_def qSwap_preserves_qGood)
-subsubsection {* For abstractions *}
+subsubsection \<open>For abstractions\<close>
lemma alphaAbsGood_equivP: "equivP qGoodAbs alphaAbsGood"
unfolding equivP_def reflP_def symP_def transP_def alphaAbsGood_def
using alphaAbs_refl alphaAbs_sym alphaAbs_trans by blast
lemma univ_asAbs_alphaAbsGood[simp]:
assumes "fAbs respectsP alphaAbsGood" and "qGoodAbs A"
shows "univ fAbs (asAbs A) = fAbs A"
by (metis assms alphaAbsGood_equivP asAbs_def univ_commute)
corollary univ_asAbs_alphaAbs[simp]:
assumes *: "fAbs respectsP alphaAbs" and **: "qGoodAbs A"
shows "univ fAbs (asAbs A) = fAbs A"
apply(rule univ_asAbs_alphaAbsGood)
using assms unfolding alphaAbsGood_def congruentP_def by auto
lemma pick_inj_on_goodAbs: "inj_on pick (Collect goodAbs)"
unfolding goodAbs_def using alphaAbsGood_equivP equivP_pick_inj_on by auto
lemma pick_injective_goodAbs[simp]:
"\<lbrakk>goodAbs A; goodAbs B\<rbrakk> \<Longrightarrow> pick A = pick B \<longleftrightarrow> A = B"
using pick_inj_on_goodAbs unfolding inj_on_def by auto
lemma goodAbs_imp_qGoodAbs_pick:
"goodAbs A \<Longrightarrow> qGoodAbs (pick A)"
unfolding goodAbs_def
using alphaAbsGood_equivP equivP_pick_preserves by fastforce
lemma qGoodAbs_iff_goodAbs_asAbs:
"goodAbs(asAbs qA) = qGoodAbs qA"
unfolding goodAbs_def asAbs_def
using alphaAbsGood_equivP proj_in_iff by fastforce
lemma pick_asAbs:
assumes "qGoodAbs qA"
shows "pick (asAbs qA) $= qA"
by (metis (full_types) alphaAbsGood_def alphaAbsGood_equivP asAbs_def assms pick_proj)
lemma asAbs_pick:
assumes "goodAbs A"
shows "asAbs (pick A) = A"
by (metis alphaAbsGood_equivP asAbs_def assms goodAbs_def proj_pick)
lemma pick_alphaAbs: "goodAbs A \<Longrightarrow> pick A $= pick A"
using goodAbs_imp_qGoodAbs_pick alphaAbs_refl by auto
lemma alphaAbs_imp_asAbs_equal:
assumes "qGoodAbs qA" and "qA $= qB"
shows "asAbs qA = asAbs qB"
by (metis (no_types, hide_lams) proj_iff alphaAbsGood_def alphaAbsGood_equivP
alphaAbs_preserves_qGoodAbs asAbs_def assms)
lemma asAbs_equal_imp_alphaAbs:
assumes "qGoodAbs qA" and "asAbs qA = asAbs qB"
shows "qA $= qB"
by (metis alphaAbs_refl
alphaAbs_sym alphaAbs_trans_twice assms pick_asAbs qGoodAbs_iff_goodAbs_asAbs)
lemma asAbs_equal_iff_alphaAbs:
assumes "qGoodAbs qA \<or> qGoodAbs qB"
shows "(asAbs qA = asAbs qB) = (qA $= qB)"
by (metis alphaAbs_imp_asAbs_equal alphaAbs_preserves_qGoodAbs
asAbs_equal_imp_alphaAbs assms)
lemma pick_alphaAbs_iff_equal:
assumes "goodAbs A" and "goodAbs B"
shows "(pick A $= pick B) = (A = B)"
using asAbs_equal_iff_alphaAbs asAbs_pick assms goodAbs_imp_qGoodAbs_pick by blast
lemma pick_swapAbs_qSwapAbs:
assumes "goodAbs A"
shows "pick (A $[x1 \<and> x2]_xs) $= ((pick A) $[[x1 \<and> x2]]_xs)"
by (simp add: assms goodAbs_imp_qGoodAbs_pick
pick_asAbs qSwapAbs_preserves_qGoodAbs swapAbs_def)
lemma asAbs_qSwapAbs_swapAbs:
assumes "qGoodAbs qA"
shows "asAbs (qA $[[x1 \<and> x2]]_xs) = ((asAbs qA) $[x1 \<and> x2]_xs)"
by (simp add: alphaAbs_imp_asAbs_equal alphaAbs_sym assms pick_asAbs
qSwapAbs_preserves_alphaAbs
qSwapAbs_preserves_qGoodAbs1 swapAbs_def)
lemma freshAbs_asAbs_qFreshAbs:
assumes "qGoodAbs qA"
shows "freshAbs xs x (asAbs qA) = qFreshAbs xs x qA"
by (simp add: assms freshAbs_def pick_asAbs qFreshAbs_preserves_alphaAbs)
lemma skelAbs_asAbs_qSkelAbs:
assumes "qGoodAbs qA"
shows "skelAbs (asAbs qA) = qSkelAbs qA"
by (simp add: alphaAll_qSkelAll assms pick_asAbs skelAbs_def)
-subsubsection {* For inputs *}
-
-text {* For unbound inputs: *}
+subsubsection \<open>For inputs\<close>
+
+text \<open>For unbound inputs:\<close>
lemma pickInp_inj_on_goodInp: "inj_on pickInp (Collect goodInp)"
unfolding pickInp_def[abs_def] inj_on_def
proof(safe, rule ext)
fix inp inp' i
assume good: "goodInp inp" "goodInp inp'" and *: "lift pick inp = lift pick inp'"
show "inp i = inp' i"
proof(cases "inp i")
assume inp: "inp i = None"
hence "lift pick inp i = None" by (auto simp add: lift_None)
hence "lift pick inp' i = None" using * by simp
hence "inp' i = None" by (auto simp add: lift_None)
thus ?thesis using inp by simp
next
fix X assume inp: "inp i = Some X"
hence "lift pick inp i = Some (pick X)" unfolding lift_def by simp
hence "lift pick inp' i = Some (pick X)" using * by simp
then obtain X' where inp': "inp' i = Some X'" and XX': "pick X = pick X'"
unfolding lift_def by(cases "inp' i", auto)
hence "good X \<and> good X'"
using inp good goodInp_def liftAll_def by (metis (hide_lams, full_types))
hence "X = X'" using XX' by auto
thus ?thesis unfolding inp inp' by simp
qed
qed
lemma goodInp_imp_qGoodInp_pickInp:
assumes "goodInp inp"
shows "qGoodInp (pickInp inp)"
unfolding pickInp_def qGoodInp_def liftAll_def
proof safe
fix i qX assume "lift pick inp i = Some qX"
then obtain X where inp: "inp i = Some X" and qX: "qX = pick X"
unfolding lift_def by(cases "inp i", auto)
hence "good X" using assms
unfolding goodInp_def liftAll_def by simp
thus "qGood qX" unfolding qX using good_imp_qGood_pick by auto
next
fix xs let ?Left = "{i. lift pick inp i \<noteq> None}"
have "?Left = {i. inp i \<noteq> None}" by(force simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using assms unfolding goodInp_def by auto
qed
lemma qGoodInp_iff_goodInp_asInp:
"goodInp (asInp qinp) = qGoodInp qinp"
proof(unfold asInp_def)
let ?inp = "lift asTerm qinp"
{assume qgood_qinp: "qGoodInp qinp"
have "goodInp ?inp"
unfolding goodInp_def liftAll_def proof safe
fix i X assume inp: "?inp i = Some X"
then obtain qX where qinp: "qinp i = Some qX" and X: "X = asTerm qX"
unfolding lift_def by(cases "qinp i", auto)
hence "qGood qX"
using qgood_qinp unfolding qGoodInp_def liftAll_def by auto
thus "good X" using X qGood_iff_good_asTerm by auto
next
fix xs let ?Left = "{i. lift asTerm qinp i \<noteq> None}"
have "?Left = {i. qinp i \<noteq> None}" by(auto simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using qgood_qinp unfolding qGoodInp_def by auto
qed
}
moreover
{assume good_inp: "goodInp ?inp"
have "qGoodInp qinp"
unfolding qGoodInp_def liftAll_def proof safe
fix i qX assume qinp: "qinp i = Some qX" let ?X = "asTerm qX"
have inp: "?inp i = Some ?X" unfolding lift_def using qinp by simp
hence "good ?X"
using good_inp unfolding goodInp_def liftAll_def by auto
thus "qGood qX" using qGood_iff_good_asTerm by auto
next
fix xs let ?Left = "{i. qinp i \<noteq> None}"
have "?Left = {i. lift asTerm qinp i \<noteq> None}" by(auto simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using good_inp unfolding goodInp_def by auto
qed
}
ultimately show "goodInp ?inp = qGoodInp qinp" by blast
qed
lemma pickInp_asInp:
assumes "qGoodInp qinp"
shows "pickInp (asInp qinp) %= qinp"
using assms unfolding pickInp_def asInp_def lift_comp
by (smt CollectI alphaInp_def asTerm_equal_iff_alpha asTerm_pick case_prodI comp_apply liftAll2_def liftAll_def lift_def option.case(2) option.sel qGoodInp_def qGood_iff_good_asTerm
sameDom_lift2)
lemma asInp_pickInp:
assumes "goodInp inp"
shows "asInp (pickInp inp) = inp"
unfolding asInp_def pickInp_def lift_comp
proof(rule ext)
fix i show "lift (asTerm \<circ> pick) inp i = inp i"
unfolding lift_def proof(cases "inp i", simp+)
fix X assume "inp i = Some X"
hence "good X" using assms unfolding goodInp_def liftAll_def by simp
thus "asTerm (pick X) = X" using asTerm_pick by auto
qed
qed
lemma pickInp_alphaInp:
assumes goodInp: "goodInp inp"
shows "pickInp inp %= pickInp inp"
using assms goodInp_imp_qGoodInp_pickInp alphaInp_refl by auto
lemma alphaInp_imp_asInp_equal:
assumes "qGoodInp qinp" and "qinp %= qinp'"
shows "asInp qinp = asInp qinp'"
unfolding asInp_def proof(rule ext)
fix i show "lift asTerm qinp i = lift asTerm qinp' i"
proof(cases "qinp i")
assume Case1: "qinp i = None"
hence "qinp' i = None"
using assms unfolding alphaInp_def sameDom_def liftAll2_def by auto
thus ?thesis using Case1 unfolding lift_def by simp
next
fix qX assume Case2: "qinp i = Some qX"
then obtain qX' where qinp': "qinp' i = Some qX'"
using assms unfolding alphaInp_def sameDom_def liftAll2_def by (cases "qinp' i", force)
hence "qX #= qX'"
using assms Case2 unfolding alphaInp_def sameDom_def liftAll2_def by auto
moreover have "qGood qX" using assms Case2 unfolding qGoodInp_def liftAll_def by auto
ultimately show ?thesis
using Case2 qinp' alpha_imp_asTerm_equal unfolding lift_def by auto
qed
qed
lemma asInp_equal_imp_alphaInp:
assumes "qGoodInp qinp" and "asInp qinp = asInp qinp'"
shows "qinp %= qinp'"
using assms unfolding alphaInp_def liftAll2_def sameDom_def
by simp (smt asInp_def asTerm_equal_iff_alpha liftAll_def lift_def option.case(2)
option.sel qGoodInp_def sameDom_def sameDom_lift2)
lemma asInp_equal_iff_alphaInp:
"qGoodInp qinp \<Longrightarrow> (asInp qinp = asInp qinp') = (qinp %= qinp')"
using asInp_equal_imp_alphaInp alphaInp_imp_asInp_equal by blast
lemma pickInp_alphaInp_iff_equal:
assumes "goodInp inp" and "goodInp inp'"
shows "(pickInp inp %= pickInp inp') = (inp = inp')"
by (metis alphaInp_imp_asInp_equal asInp_equal_imp_alphaInp
asInp_pickInp assms goodInp_imp_qGoodInp_pickInp)
lemma pickInp_swapInp_qSwapInp:
assumes "goodInp inp"
shows "pickInp (inp %[x1 \<and> x2]_xs) %= ((pickInp inp) %[[x1 \<and> x2]]_xs)"
using assms unfolding alphaInp_def sameDom_def liftAll2_def
pickInp_def swapInp_def qSwapInp_def lift_comp
by (simp add: lift_None)
(smt assms comp_apply goodInp_imp_qGoodInp_pickInp liftAll_def lift_def local.swap_def option.case_eq_if option.sel option.simps(3) pickInp_def
pick_asTerm qGoodInp_def qSwap_preserves_qGood1)
lemma asInp_qSwapInp_swapInp:
assumes "qGoodInp qinp"
shows "asInp (qinp %[[x1 \<and> x2]]_xs) = ((asInp qinp) %[x1 \<and> x2]_xs)"
proof-
{fix i qX assume "qinp i = Some qX"
hence "qGood qX" using assms unfolding qGoodInp_def liftAll_def by auto
hence "asTerm (qX #[[x1 \<and> x2]]_xs) = swap xs x1 x2 (asTerm qX)"
by(auto simp add: asTerm_qSwap_swap)
}
thus ?thesis
using assms
by (smt asInp_def comp_apply lift_comp lift_cong qSwapInp_def swapInp_def)
qed
lemma swapInp_def2:
"(inp %[x1 \<and> x2]_xs) = asInp ((pickInp inp) %[[x1 \<and> x2]]_xs)"
unfolding swapInp_def asInp_def pickInp_def qSwapInp_def lift_def swap_def
apply(rule ext) subgoal for i by (cases "inp i") auto .
lemma freshInp_def2:
"freshInp xs x inp = qFreshInp xs x (pickInp inp)"
unfolding freshInp_def qFreshInp_def pickInp_def lift_def fresh_def liftAll_def
apply(rule iff_allI) subgoal for i by (cases "inp i") auto .
-text {* For bound inputs: *}
+text \<open>For bound inputs:\<close>
lemma pickBinp_inj_on_goodBinp: "inj_on pickBinp (Collect goodBinp)"
unfolding pickBinp_def[abs_def] inj_on_def
proof(safe, rule ext)
fix binp binp' i
assume good: "goodBinp binp" "goodBinp binp'" and *: "lift pick binp = lift pick binp'"
show "binp i = binp' i"
proof(cases "binp i")
assume binp: "binp i = None"
hence "lift pick binp i = None" by (auto simp add: lift_None)
hence "lift pick binp' i = None" using * by simp
hence "binp' i = None" by (auto simp add: lift_None)
thus ?thesis using binp by simp
next
fix A assume binp: "binp i = Some A"
hence "lift pick binp i = Some (pick A)" unfolding lift_def by simp
hence "lift pick binp' i = Some (pick A)" using * by simp
then obtain A' where binp': "binp' i = Some A'" and AA': "pick A = pick A'"
unfolding lift_def by(cases "binp' i", auto)
hence "goodAbs A \<and> goodAbs A'"
using binp good goodBinp_def liftAll_def by (metis (hide_lams, full_types))
hence "A = A'" using AA' by auto
thus ?thesis unfolding binp binp' by simp
qed
qed
lemma goodBinp_imp_qGoodBinp_pickBinp:
assumes "goodBinp binp"
shows "qGoodBinp (pickBinp binp)"
unfolding pickBinp_def qGoodBinp_def liftAll_def proof safe
fix i qA assume "lift pick binp i = Some qA"
then obtain A where binp: "binp i = Some A" and qA: "qA = pick A"
unfolding lift_def by(cases "binp i", auto)
hence "goodAbs A" using assms
unfolding goodBinp_def liftAll_def by simp
thus "qGoodAbs qA" unfolding qA using goodAbs_imp_qGoodAbs_pick by auto
next
fix xs let ?Left = "{i. lift pick binp i \<noteq> None}"
have "?Left = {i. binp i \<noteq> None}" by(force simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using assms unfolding goodBinp_def by auto
qed
lemma qGoodBinp_iff_goodBinp_asBinp:
"goodBinp (asBinp qbinp) = qGoodBinp qbinp"
proof(unfold asBinp_def)
let ?binp = "lift asAbs qbinp"
{assume qgood_qbinp: "qGoodBinp qbinp"
have "goodBinp ?binp"
unfolding goodBinp_def liftAll_def proof safe
fix i A assume binp: "?binp i = Some A"
then obtain qA where qbinp: "qbinp i = Some qA" and A: "A = asAbs qA"
unfolding lift_def by(cases "qbinp i", auto)
hence "qGoodAbs qA"
using qgood_qbinp unfolding qGoodBinp_def liftAll_def by auto
thus "goodAbs A" using A qGoodAbs_iff_goodAbs_asAbs by auto
next
fix xs let ?Left = "{i. lift asAbs qbinp i \<noteq> None}"
have "?Left = {i. qbinp i \<noteq> None}" by(auto simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using qgood_qbinp unfolding qGoodBinp_def by auto
qed
}
moreover
{assume good_binp: "goodBinp ?binp"
have "qGoodBinp qbinp"
unfolding qGoodBinp_def liftAll_def proof safe
fix i qA assume qbinp: "qbinp i = Some qA" let ?A = "asAbs qA"
have binp: "?binp i = Some ?A" unfolding lift_def using qbinp by simp
hence "goodAbs ?A"
using good_binp unfolding goodBinp_def liftAll_def by auto
thus "qGoodAbs qA" using qGoodAbs_iff_goodAbs_asAbs by auto
next
fix xs let ?Left = "{i. qbinp i \<noteq> None}"
have "?Left = {i. lift asAbs qbinp i \<noteq> None}" by(auto simp add: lift_None)
thus "|?Left| <o |UNIV :: 'var set|" using good_binp unfolding goodBinp_def by auto
qed
}
ultimately show "goodBinp ?binp = qGoodBinp qbinp" by blast
qed
lemma pickBinp_asBinp:
assumes "qGoodBinp qbinp"
shows "pickBinp (asBinp qbinp) %%= qbinp"
unfolding pickBinp_def asBinp_def lift_comp alphaBinp_def using sameDom_lift2
by auto (smt assms comp_apply liftAll2_def liftAll_def
lift_def option.sel option.simps(5) pick_asAbs qGoodBinp_def)
lemma asBinp_pickBinp:
assumes "goodBinp binp"
shows "asBinp (pickBinp binp) = binp"
unfolding asBinp_def pickBinp_def lift_comp
apply(rule ext)
subgoal for i apply(cases "binp i")
using assms asAbs_pick unfolding goodBinp_def liftAll_def lift_def by auto .
lemma pickBinp_alphaBinp:
assumes goodBinp: "goodBinp binp"
shows "pickBinp binp %%= pickBinp binp"
using assms goodBinp_imp_qGoodBinp_pickBinp alphaBinp_refl by auto
lemma alphaBinp_imp_asBinp_equal:
assumes "qGoodBinp qbinp" and "qbinp %%= qbinp'"
shows "asBinp qbinp = asBinp qbinp'"
unfolding asBinp_def proof(rule ext)
fix i show "lift asAbs qbinp i = lift asAbs qbinp' i"
proof(cases "qbinp i")
case None
hence "qbinp' i = None"
using assms unfolding alphaBinp_def sameDom_def liftAll2_def by auto
thus ?thesis using None unfolding lift_def by simp
next
case (Some qA)
then obtain qA' where qbinp': "qbinp' i = Some qA'"
using assms unfolding alphaBinp_def sameDom_def liftAll2_def by (cases "qbinp' i", force)
hence "qA $= qA'"
using assms Some unfolding alphaBinp_def sameDom_def liftAll2_def by auto
moreover have "qGoodAbs qA" using assms Some unfolding qGoodBinp_def liftAll_def by auto
ultimately show ?thesis
using Some qbinp' alphaAbs_imp_asAbs_equal unfolding lift_def by auto
qed
qed
lemma asBinp_equal_imp_alphaBinp:
assumes "qGoodBinp qbinp" and "asBinp qbinp = asBinp qbinp'"
shows "qbinp %%= qbinp'"
using assms unfolding alphaBinp_def liftAll2_def sameDom_def
by simp (smt asAbs_equal_imp_alphaAbs asBinp_def liftAll_def
lift_None lift_def option.inject option.simps(5) qGoodBinp_def)
lemma asBinp_equal_iff_alphaBinp:
"qGoodBinp qbinp \<Longrightarrow> (asBinp qbinp = asBinp qbinp') = (qbinp %%= qbinp')"
using asBinp_equal_imp_alphaBinp alphaBinp_imp_asBinp_equal by blast
lemma pickBinp_alphaBinp_iff_equal:
assumes "goodBinp binp" and "goodBinp binp'"
shows "(pickBinp binp %%= pickBinp binp') = (binp = binp')"
using assms goodBinp_imp_qGoodBinp_pickBinp asBinp_pickBinp pickBinp_alphaBinp
by (metis asBinp_equal_iff_alphaBinp)
lemma pickBinp_swapBinp_qSwapBinp:
assumes "goodBinp binp"
shows "pickBinp (binp %%[x1 \<and> x2]_xs) %%= ((pickBinp binp) %%[[x1 \<and> x2]]_xs)"
using assms unfolding pickBinp_def swapBinp_def qSwapBinp_def lift_comp
alphaBinp_def sameDom_def liftAll2_def
by (simp add: goodBinp_def liftAll_def lift_def option.case_eq_if pick_swapAbs_qSwapAbs)
lemma asBinp_qSwapBinp_swapBinp:
assumes "qGoodBinp qbinp"
shows "asBinp (qbinp %%[[x1 \<and> x2]]_xs) = ((asBinp qbinp) %%[x1 \<and> x2]_xs)"
unfolding asBinp_def swapBinp_def qSwapBinp_def lift_comp alphaBinp_def lift_def
apply(rule ext) subgoal for i apply(cases "qbinp i")
using assms asAbs_qSwapAbs_swapAbs by (fastforce simp add: liftAll_def qGoodBinp_def)+ .
lemma swapBinp_def2:
"(binp %%[x1 \<and> x2]_xs) = asBinp ((pickBinp binp) %%[[x1 \<and> x2]]_xs)"
unfolding swapBinp_def asBinp_def pickBinp_def qSwapBinp_def lift_def swapAbs_def
apply (rule ext) subgoal for i by (cases "binp i") simp_all .
lemma freshBinp_def2:
"freshBinp xs x binp = qFreshBinp xs x (pickBinp binp)"
unfolding freshBinp_def qFreshBinp_def pickBinp_def lift_def freshAbs_def liftAll_def
apply (rule iff_allI) subgoal for i by (cases "binp i") simp_all .
(* Note that psubstInp and psubstBinp are discussed in the next subsubsection,
about environments. *)
-subsubsection {* For environments *}
+subsubsection \<open>For environments\<close>
(* Remember we do not have any "quasi-swap" for environments --
we plan to prove most of the things concerning parallel substitution
and environments for equivPalence classes directly. *)
lemma goodEnv_imp_qGoodEnv_pickE:
assumes "goodEnv rho"
shows "qGoodEnv (pickE rho)"
unfolding qGoodEnv_def pickE_def
apply(auto simp del: "not_None_eq")
using assms good_imp_qGood_pick unfolding liftAll_lift_comp comp_def
by (auto simp: goodEnv_def liftAll_def lift_None)
lemma qGoodEnv_iff_goodEnv_asEnv:
"goodEnv (asEnv qrho) = qGoodEnv qrho"
unfolding asEnv_def unfolding goodEnv_def liftAll_lift_comp comp_def
by (auto simp: qGoodEnv_def lift_None liftAll_def qGood_iff_good_asTerm)
lemma pickE_asEnv:
assumes "qGoodEnv qrho"
shows "pickE (asEnv qrho) &= qrho"
using assms
by (auto simp: lift_None liftAll_def lift_def alphaEnv_def sameDom_def liftAll2_def
pick_asTerm qGoodEnv_def pickE_def asEnv_def split: option.splits)
lemma asEnv_pickE:
assumes "goodEnv rho" shows "asEnv (pickE rho) xs x = rho xs x"
using assms asTerm_pick
by (cases "rho xs x") (auto simp: goodEnv_def liftAll_def asEnv_def pickE_def lift_comp lift_def)
lemma pickE_alphaEnv:
assumes goodEnv: "goodEnv rho" shows "pickE rho &= pickE rho"
using assms goodEnv_imp_qGoodEnv_pickE alphaEnv_refl by auto
lemma alphaEnv_imp_asEnv_equal:
assumes "qGoodEnv qrho" and "qrho &= qrho'"
shows "asEnv qrho = asEnv qrho'"
apply (rule ext)+ subgoal for xs x apply(cases "qrho xs x")
using assms asTerm_equal_iff_alpha alpha_imp_asTerm_equal
by (auto simp add: alphaEnv_def sameDom_def asEnv_def lift_def
qGoodEnv_def liftAll_def liftAll2_def option.case_eq_if split: option.splits)
blast+ .
lemma asEnv_equal_imp_alphaEnv:
assumes "qGoodEnv qrho" and "asEnv qrho = asEnv qrho'"
shows "qrho &= qrho'"
using assms unfolding alphaEnv_def sameDom_def liftAll2_def
apply (simp add: asEnv_def lift_None lift_def qGoodEnv_def liftAll_def)
by (smt asTerm_equal_imp_alpha option.sel option.simps(5) option.case_eq_if option.distinct(1))
lemma asEnv_equal_iff_alphaEnv:
"qGoodEnv qrho \<Longrightarrow> (asEnv qrho = asEnv qrho') = (qrho &= qrho')"
using asEnv_equal_imp_alphaEnv alphaEnv_imp_asEnv_equal by blast
lemma pickE_alphaEnv_iff_equal:
assumes "goodEnv rho" and "goodEnv rho'"
shows "(pickE rho &= pickE rho') = (rho = rho')"
proof(rule iffI, safe, (rule ext)+)
fix xs x
assume alpha: "pickE rho &= pickE rho'"
have qgood_rho: "qGoodEnv (pickE rho)" using assms goodEnv_imp_qGoodEnv_pickE by auto
have "rho xs x = asEnv (pickE rho) xs x" using assms asEnv_pickE by fastforce
also have "\<dots> = asEnv (pickE rho') xs x"
using qgood_rho alpha alphaEnv_imp_asEnv_equal by fastforce
also have "\<dots> = rho' xs x" using assms asEnv_pickE by fastforce
finally show "rho xs x = rho' xs x" .
next
have "qGoodEnv(pickE rho')" using assms goodEnv_imp_qGoodEnv_pickE by auto
thus "pickE rho' &= pickE rho'" using alphaEnv_refl by auto
qed
lemma freshEnv_def2:
"freshEnv xs x rho = qFreshEnv xs x (pickE rho)"
unfolding freshEnv_def qFreshEnv_def pickE_def lift_def fresh_def liftAll_def
apply(cases "rho xs x")
by (auto intro!: iff_allI) (metis map_option_case map_option_eq_Some)
lemma pick_psubst_qPsubst:
assumes "good X" and "goodEnv rho"
shows "pick (X #[rho]) #= ((pick X) #[[pickE rho]])"
by (simp add: assms goodEnv_imp_qGoodEnv_pickE good_imp_qGood_pick
pick_asTerm psubst_def qPsubst_preserves_qGood)
lemma pick_psubstAbs_qPsubstAbs:
assumes "goodAbs A" and "goodEnv rho"
shows "pick (A $[rho]) $= ((pick A) $[[pickE rho]])"
by (simp add: assms goodAbs_imp_qGoodAbs_pick goodEnv_imp_qGoodEnv_pickE pick_asAbs
psubstAbs_def qPsubstAbs_preserves_qGoodAbs)
lemma pickInp_psubstInp_qPsubstInp:
assumes good: "goodInp inp" and good_rho: "goodEnv rho"
shows "pickInp (inp %[rho]) %= ((pickInp inp) %[[pickE rho]])"
using assms unfolding pickInp_def psubstInp_def qPsubstInp_def lift_comp
unfolding alphaInp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt comp_apply goodEnv_imp_qGoodEnv_pickE goodInp_imp_qGoodInp_pickInp liftAll_def lift_def map_option_case map_option_eq_Some option.sel pickInp_def
pick_asTerm psubst_def qGoodInp_def qPsubst_preserves_qGood)
lemma pickBinp_psubstBinp_qPsubstBinp:
assumes good: "goodBinp binp" and good_rho: "goodEnv rho"
shows "pickBinp (binp %%[rho]) %%= ((pickBinp binp) %%[[pickE rho]])"
using assms unfolding pickBinp_def psubstBinp_def qPsubstBinp_def lift_comp
unfolding alphaBinp_def sameDom_def liftAll2_def
by (simp add: lift_None)
(smt comp_apply goodBinp_def liftAll_def lift_def map_option_case map_option_eq_Some
option.sel pick_psubstAbs_qPsubstAbs)
-subsubsection{* The structural alpha-equivPalence maps commute with the syntactic constructs *}
+subsubsection\<open>The structural alpha-equivPalence maps commute with the syntactic constructs\<close>
lemma pick_Var_qVar:
"pick (Var xs x) #= qVar xs x"
unfolding Var_def using pick_asTerm by force
lemma Op_asInp_asTerm_qOp:
assumes "qGoodInp qinp" and "qGoodBinp qbinp"
shows "Op delta (asInp qinp) (asBinp qbinp) = asTerm (qOp delta qinp qbinp)"
using assms pickInp_asInp pickBinp_asBinp unfolding Op_def
by(auto simp add: asTerm_equal_iff_alpha)
lemma qOp_pickInp_pick_Op:
assumes "goodInp inp" and "goodBinp binp"
shows "qOp delta (pickInp inp) (pickBinp binp) #= pick (Op delta inp binp)"
using assms goodInp_imp_qGoodInp_pickInp goodBinp_imp_qGoodBinp_pickBinp
unfolding Op_def using pick_asTerm alpha_sym by force
lemma Abs_asTerm_asAbs_qAbs:
assumes "qGood qX"
shows "Abs xs x (asTerm qX) = asAbs (qAbs xs x qX)"
using assms pick_asTerm qAbs_preserves_alpha unfolding Abs_def
by(force simp add: asAbs_equal_iff_alphaAbs)
lemma qAbs_pick_Abs:
assumes "good X"
shows "qAbs xs x (pick X) $= pick (Abs xs x X)"
using assms good_imp_qGood_pick pick_asAbs alphaAbs_sym unfolding Abs_def by force
lemmas qItem_versus_item_simps =
univ_asTerm_alphaGood univ_asAbs_alphaAbsGood
univ_asTerm_alpha univ_asAbs_alphaAbs
pick_injective_good pick_injective_goodAbs
-subsection {* All operators preserve the ``good'' predicate *}
+subsection \<open>All operators preserve the ``good'' predicate\<close>
(* Note: some facts here simply do not hold as ``iff"s. *)
lemma Var_preserves_good[simp]:
"good(Var xs x::('index,'bindex,'varSort,'var,'opSym)term)"
by (metis Var_def qGood.simps(1) qGood_iff_good_asTerm)
lemma Op_preserves_good[simp]:
assumes "goodInp inp" and "goodBinp binp"
shows "good(Op delta inp binp)"
using assms goodInp_imp_qGoodInp_pickInp goodBinp_imp_qGoodBinp_pickBinp
qGood_iff_good_asTerm unfolding Op_def by fastforce
lemma Abs_preserves_good[simp]:
assumes good: "good X"
shows "goodAbs(Abs xs x X)"
using assms good_imp_qGood_pick qGoodAbs_iff_goodAbs_asAbs
unfolding Abs_def by fastforce
lemmas Cons_preserve_good =
Var_preserves_good Op_preserves_good Abs_preserves_good
lemma swap_preserves_good[simp]:
assumes "good X"
shows "good (X #[x \<and> y]_xs)"
using assms good_imp_qGood_pick qSwap_preserves_qGood qGood_iff_good_asTerm
unfolding swap_def by fastforce
lemma swapAbs_preserves_good[simp]:
assumes "goodAbs A"
shows "goodAbs (A $[x \<and> y]_xs)"
using assms goodAbs_imp_qGoodAbs_pick qSwapAbs_preserves_qGoodAbs qGoodAbs_iff_goodAbs_asAbs
unfolding swapAbs_def by fastforce
lemma swapInp_preserves_good[simp]:
assumes "goodInp inp"
shows "goodInp (inp %[x \<and> y]_xs)"
using assms
by (auto simp: goodInp_def lift_def swapInp_def liftAll_def split: option.splits)
lemma swapBinp_preserves_good[simp]:
assumes "goodBinp binp"
shows "goodBinp (binp %%[x \<and> y]_xs)"
using assms
by (auto simp: goodBinp_def lift_def swapBinp_def liftAll_def split: option.splits)
lemma swapEnvDom_preserves_good:
assumes "goodEnv rho"
shows "goodEnv (swapEnvDom xs x y rho)" (is "goodEnv ?rho'")
unfolding goodEnv_def liftAll_def proof safe
fix zs z X' assume rho': "?rho' zs z = Some X'"
hence "rho zs (z @zs[x \<and> y]_xs) = Some X'" unfolding swapEnvDom_def by simp
thus "good X'" using assms unfolding goodEnv_def liftAll_def by simp
next
fix xsa ys let ?Left = "{ya. ?rho' ys ya \<noteq> None}"
have "|{y} \<union> {ya. rho ys ya \<noteq> None}| <o |UNIV :: 'var set|"
using assms var_infinite_INNER card_of_Un_singl_ordLess_infinite
unfolding goodEnv_def by fastforce
hence "|{x,y} \<union> {ya. rho ys ya \<noteq> None}| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_singl_ordLess_infinite by fastforce
moreover
{have "?Left \<subseteq> {x,y} \<union> {ya. rho ys ya \<noteq> None}"
unfolding swapEnvDom_def sw_def[abs_def] by auto
hence "|?Left| \<le>o |{x,y} \<union> {ya. rho ys ya \<noteq> None}|"
using card_of_mono1 by auto
}
ultimately show "|?Left| <o |UNIV :: 'var set|" using ordLeq_ordLess_trans by blast
qed
lemma swapEnvIm_preserves_good:
assumes "goodEnv rho"
shows "goodEnv (swapEnvIm xs x y rho)"
using assms unfolding goodEnv_def swapEnvIm_def liftAll_def
by (auto simp: lift_def split: option.splits)
lemma swapEnv_preserves_good[simp]:
assumes "goodEnv rho"
shows "goodEnv (rho &[x \<and> y]_xs)"
unfolding swapEnv_def comp_def
using assms by(auto simp add: swapEnvDom_preserves_good swapEnvIm_preserves_good)
lemmas swapAll_preserve_good =
swap_preserves_good swapAbs_preserves_good
swapInp_preserves_good swapBinp_preserves_good
swapEnv_preserves_good
lemma psubst_preserves_good[simp]:
assumes "goodEnv rho" and "good X"
shows "good (X #[rho])"
using assms good_imp_qGood_pick goodEnv_imp_qGoodEnv_pickE
qPsubst_preserves_qGood qGood_iff_good_asTerm unfolding psubst_def by fastforce
lemma psubstAbs_preserves_good[simp]:
assumes good_rho: "goodEnv rho" and goodAbs_A: "goodAbs A"
shows "goodAbs (A $[rho])"
using assms goodAbs_A goodAbs_imp_qGoodAbs_pick goodEnv_imp_qGoodEnv_pickE
qPsubstAbs_preserves_qGoodAbs qGoodAbs_iff_goodAbs_asAbs unfolding psubstAbs_def by fastforce
lemma psubstInp_preserves_good[simp]:
assumes good_rho: "goodEnv rho" and good: "goodInp inp"
shows "goodInp (inp %[rho])"
using assms unfolding goodInp_def psubstInp_def liftAll_def
by (auto simp add: lift_def split: option.splits)
lemma psubstBinp_preserves_good[simp]:
assumes good_rho: "goodEnv rho" and good: "goodBinp binp"
shows "goodBinp (binp %%[rho])"
using assms unfolding goodBinp_def psubstBinp_def liftAll_def
by (auto simp add: lift_def split: option.splits)
lemma psubstEnv_preserves_good[simp]:
assumes good: "goodEnv rho" and good': "goodEnv rho'"
shows "goodEnv (rho &[rho'])"
unfolding goodEnv_def liftAll_def
proof safe
fix zs z X'
assume *: "(rho &[rho']) zs z = Some X'"
show "good X'"
proof(cases "rho zs z")
case None
hence "rho' zs z = Some X'" using * unfolding psubstEnv_def by auto
thus ?thesis using good' unfolding goodEnv_def liftAll_def by auto
next
case (Some X)
hence "X' = (X #[rho'])" using * unfolding psubstEnv_def by auto
moreover have "good X" using Some good unfolding goodEnv_def liftAll_def by auto
ultimately show ?thesis using good' psubst_preserves_good by auto
qed
next
fix xs ys let ?Left = "{y. (rho &[rho']) ys y \<noteq> None}"
let ?Left1 = "{y. rho ys y \<noteq> None}" let ?Left2 = "{y. rho' ys y \<noteq> None}"
have "|?Left1| <o |UNIV :: 'var set| \<and> |?Left2| <o |UNIV :: 'var set|"
using good good' unfolding goodEnv_def by simp
hence "|?Left1 \<union> ?Left2| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_ordLess_infinite by auto
moreover
{have "?Left \<subseteq> ?Left1 \<union> ?Left2" unfolding psubstEnv_def by auto
hence "|?Left| \<le>o |?Left1 \<union> ?Left2|" using card_of_mono1 by auto
}
ultimately show "|?Left| <o |UNIV :: 'var set|" using ordLeq_ordLess_trans by blast
qed
lemmas psubstAll_preserve_good =
psubst_preserves_good psubstAbs_preserves_good
psubstInp_preserves_good psubstBinp_preserves_good
psubstEnv_preserves_good
lemma idEnv_preserves_good[simp]: "goodEnv idEnv"
unfolding goodEnv_def idEnv_def liftAll_def
using var_infinite_INNER finite_ordLess_infinite2 by auto
lemma updEnv_preserves_good[simp]:
assumes good_X: "good X" and good_rho: "goodEnv rho"
shows "goodEnv (rho [x \<leftarrow> X]_xs)"
using assms unfolding updEnv_def goodEnv_def liftAll_def
proof safe
fix ys y Y
assume "good X" and "\<forall>ys y Y. rho ys y = Some Y \<longrightarrow> good Y"
and "(if ys = xs \<and> y = x then Some X else rho ys y) = Some Y"
thus "good Y"
by(cases "ys = xs \<and> y = x") auto
next
fix ys
let ?V' = "{y. (if ys = xs \<and> y = x then Some X else rho ys y) \<noteq> None}"
let ?V = "\<lambda> ys. {y. rho ys y \<noteq> None}"
assume "\<forall> ys. |?V ys| <o |UNIV :: 'var set|"
hence "|{x} \<union> ?V ys| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_singl_ordLess_infinite by fastforce
moreover
{have "?V' \<subseteq> {x} \<union> ?V ys" by auto
hence "|?V'| \<le>o |{x} \<union> ?V ys|" using card_of_mono1 by auto
}
ultimately show "|?V'| <o |UNIV :: 'var set|" using ordLeq_ordLess_trans by blast
qed
lemma getEnv_preserves_good[simp]:
assumes "goodEnv rho" and "rho xs x = Some X"
shows "good X"
using assms unfolding goodEnv_def liftAll_def by simp
lemmas envOps_preserve_good =
idEnv_preserves_good updEnv_preserves_good
getEnv_preserves_good
lemma subst_preserves_good[simp]:
assumes "good X" and "good Y"
shows "good (Y #[X / x]_xs)"
unfolding subst_def
using assms by simp
lemma substAbs_preserves_good[simp]:
assumes "good X" and "goodAbs A"
shows "goodAbs (A $[X / x]_xs)"
unfolding substAbs_def
using assms by simp
lemma substInp_preserves_good[simp]:
assumes "good X" and "goodInp inp"
shows "goodInp (inp %[X / x]_xs)"
unfolding substInp_def using assms by simp
lemma substBinp_preserves_good[simp]:
assumes "good X" and "goodBinp binp"
shows "goodBinp (binp %%[X / x]_xs)"
unfolding substBinp_def using assms by simp
lemma substEnv_preserves_good[simp]:
assumes "good X" and "goodEnv rho"
shows "goodEnv (rho &[X / x]_xs)"
unfolding substEnv_def using assms by simp
lemmas substAll_preserve_good =
subst_preserves_good substAbs_preserves_good
substInp_preserves_good substBinp_preserves_good
substEnv_preserves_good
lemma vsubst_preserves_good[simp]:
assumes "good Y"
shows "good (Y #[x1 // x]_xs)"
unfolding vsubst_def using assms by simp
lemma vsubstAbs_preserves_good[simp]:
assumes "goodAbs A"
shows "goodAbs (A $[x1 // x]_xs)"
unfolding vsubstAbs_def using assms by simp
lemma vsubstInp_preserves_good[simp]:
assumes "goodInp inp"
shows "goodInp (inp %[x1 // x]_xs)"
unfolding vsubstInp_def using assms by simp
lemma vsubstBinp_preserves_good[simp]:
assumes "goodBinp binp"
shows "goodBinp (binp %%[x1 // x]_xs)"
unfolding vsubstBinp_def using assms by simp
lemma vsubstEnv_preserves_good[simp]:
assumes "goodEnv rho"
shows "goodEnv (rho &[x1 // x]_xs)"
unfolding vsubstEnv_def using assms by simp
lemmas vsubstAll_preserve_good =
vsubst_preserves_good vsubstAbs_preserves_good
vsubstInp_preserves_good vsubstBinp_preserves_good
vsubstEnv_preserves_good
lemmas all_preserve_good =
Cons_preserve_good
swapAll_preserve_good
psubstAll_preserve_good
envOps_preserve_good
substAll_preserve_good
vsubstAll_preserve_good
-subsubsection {* The syntactic operators are almost constructors *}
-
-text{* The only one that does not act precisely like a constructor is ``Abs". *}
+subsubsection \<open>The syntactic operators are almost constructors\<close>
+
+text\<open>The only one that does not act precisely like a constructor is ``Abs".\<close>
theorem Var_inj[simp]:
"(((Var xs x)::('index,'bindex,'varSort,'var,'opSym)term) = Var ys y) =
(xs = ys \<and> x = y)"
by (metis alpha_qVar_iff pick_Var_qVar qTerm.inject)
lemma Op_inj[simp]:
assumes "goodInp inp" and "goodBinp binp"
and "goodInp inp'" and "goodBinp binp'"
shows
"(Op delta inp binp = Op delta' inp' binp') =
(delta = delta' \<and> inp = inp' \<and> binp = binp')"
using assms pickInp_alphaInp_iff_equal pickBinp_alphaBinp_iff_equal
goodInp_imp_qGoodInp_pickInp goodBinp_imp_qGoodBinp_pickBinp
unfolding Op_def by (fastforce simp: asTerm_equal_iff_alpha)
-text{* ``Abs" is almost injective (``ainj"), with almost injectivity expressed
+text\<open>``Abs" is almost injective (``ainj"), with almost injectivity expressed
in two ways:
\\- maximally, using "forall" -- this is suitable for elimination of ``Abs" equalities;
\\- minimally, using "exists" -- this is suitable for introduction of ``Abs" equalities.
- *}
+\<close>
lemma Abs_ainj_all:
assumes good: "good X" and good': "good X'"
shows
"(Abs xs x X = Abs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X') \<longrightarrow>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
proof-
let ?qX = "pick X" let ?qX' = "pick X'"
have qgood: "qGood ?qX \<and> qGood ?qX'" using good good' good_imp_qGood_pick by auto
hence qgood_qXyx: "\<forall> y. qGood (?qX #[[y \<and> x]]_xs)"
using qSwap_preserves_qGood by auto
have "qGoodAbs(qAbs xs x ?qX)" using qgood by simp
hence "(Abs xs x X = Abs xs' x' X') = (qAbs xs x ?qX $= qAbs xs' x' ?qX')"
unfolding Abs_def by (auto simp add: asAbs_equal_iff_alphaAbs)
also
have "\<dots> = (xs = xs' \<and>
(\<forall> y. (y = x \<or> qFresh xs y ?qX) \<and> (y = x' \<or> qFresh xs y ?qX') \<longrightarrow>
(?qX #[[y \<and> x]]_xs) #= (?qX' #[[y \<and> x']]_xs)))"
using qgood alphaAbs_qAbs_iff_all_equal_or_qFresh[of ?qX ?qX'] by blast
also
have "\<dots> = (xs = xs' \<and>
(\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X') \<longrightarrow>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
unfolding fresh_def swap_def using qgood_qXyx by (auto simp add: asTerm_equal_iff_alpha)
finally show ?thesis .
qed
lemma Abs_ainj_ex:
assumes good: "good X" and good': "good X'"
shows
"(Abs xs x X = Abs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X' \<and>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
proof-
let ?qX = "pick X" let ?qX' = "pick X'"
have qgood: "qGood ?qX \<and> qGood ?qX'" using good good' good_imp_qGood_pick by auto
hence qgood_qXyx: "\<forall> y. qGood (?qX #[[y \<and> x]]_xs)"
using qSwap_preserves_qGood by auto
have "qGoodAbs(qAbs xs x ?qX)" using qgood by simp
hence "(Abs xs x X = Abs xs' x' X') = (qAbs xs x ?qX $= qAbs xs' x' ?qX')"
unfolding Abs_def by (auto simp add: asAbs_equal_iff_alphaAbs)
also
have "\<dots> = (xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> qFresh xs y ?qX \<and> qFresh xs y ?qX' \<and>
(?qX #[[y \<and> x]]_xs) #= (?qX' #[[y \<and> x']]_xs)))"
using qgood alphaAbs_qAbs_iff_ex_distinct_qFresh[of ?qX xs x xs' x' ?qX'] by blast
also
have "\<dots> = (xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X' \<and>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
unfolding fresh_def swap_def using qgood_qXyx asTerm_equal_iff_alpha by auto
finally show ?thesis .
qed
lemma Abs_cong[fundef_cong]:
assumes good: "good X" and good': "good X'"
and y: "fresh xs y X" and y': "fresh xs y X'"
and eq: "(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)"
shows "Abs xs x X = Abs xs x' X'"
proof-
let ?qX = "pick X" let ?qX' = "pick X'"
have qgood: "qGood ?qX \<and> qGood ?qX'" using good good' good_imp_qGood_pick by auto
hence qgood_qXyx: "\<forall> y. qGood (?qX #[[y \<and> x]]_xs)"
using qSwap_preserves_qGood by auto
have qEq: "(?qX #[[y \<and> x]]_xs) #= (?qX' #[[y \<and> x']]_xs)"
using eq unfolding fresh_def swap_def
using qgood_qXyx asTerm_equal_iff_alpha by auto
have "(qAbs xs x ?qX $= qAbs xs x' ?qX')"
apply(rule alphaAbs_ex_equal_or_qFresh_imp_alphaAbs_qAbs)
using qgood apply simp
unfolding alphaAbs_ex_equal_or_qFresh_def using y y' qEq
unfolding fresh_def by auto
moreover have "qGoodAbs(qAbs xs x ?qX)" using qgood by simp
ultimately show "Abs xs x X = Abs xs x' X'"
unfolding Abs_def by (auto simp add: asAbs_equal_iff_alphaAbs)
qed
lemma Abs_swap_fresh:
assumes good_X: "good X" and fresh: "fresh xs x' X"
shows "Abs xs x X = Abs xs x' (X #[x' \<and> x]_xs)"
proof-
let ?x'x = "swap xs x' x" let ?qx'x = "qSwap xs x' x"
have good_pickX: "qGood (pick X)" using good_X good_imp_qGood_pick by auto
hence good_qAbs_pickX: "qGoodAbs (qAbs xs x (pick X))" by simp
have good_x'x_pickX: "qGood (?qx'x (pick X))"
using good_pickX qSwap_preserves_qGood by auto
(* *)
have "Abs xs x X = asAbs (qAbs xs x (pick X))" unfolding Abs_def by simp
also
{have "qAbs xs x (pick X) $= qAbs xs x' (?qx'x (pick X))"
using good_pickX fresh unfolding fresh_def using qAbs_alphaAbs_qSwap_qFresh by fastforce
moreover
{have "?qx'x (pick X) #= pick (?x'x X)"
using good_X by (auto simp add: pick_swap_qSwap alpha_sym)
hence "qAbs xs x' (?qx'x (pick X)) $= qAbs xs x' (pick (?x'x X))"
using good_x'x_pickX qAbs_preserves_alpha by fastforce
}
ultimately have "qAbs xs x (pick X) $= qAbs xs x' (pick (?x'x X))"
using good_qAbs_pickX alphaAbs_trans by blast
hence "asAbs (qAbs xs x (pick X)) = asAbs (qAbs xs x' (pick (?x'x X)))"
using good_qAbs_pickX by (auto simp add: asAbs_equal_iff_alphaAbs)
}
also have "asAbs (qAbs xs x' (pick (?x'x X))) = Abs xs x' (?x'x X)"
unfolding Abs_def by auto
finally show ?thesis .
qed
lemma Var_diff_Op[simp]:
"Var xs x \<noteq> Op delta inp binp"
by (simp add: Op_def Var_def asTerm_equal_iff_alpha)
lemma Op_diff_Var[simp]:
"Op delta inp binp \<noteq> Var xs x"
using Var_diff_Op[of _ _ _ inp] by blast
theorem term_nchotomy:
assumes "good X"
shows
"(\<exists> xs x. X = Var xs x) \<or>
(\<exists> delta inp binp. goodInp inp \<and> goodBinp binp \<and> X = Op delta inp binp)"
proof-
let ?qX = "pick X"
have good_qX: "qGood ?qX" using assms good_imp_qGood_pick by auto
have X: "X = asTerm ?qX" using assms asTerm_pick by auto
show ?thesis
proof(cases "?qX")
fix xs x assume Case1: "?qX = qVar xs x"
have "X = Var xs x" unfolding Var_def using X Case1 by simp
thus ?thesis by blast
next
fix delta qinp qbinp assume Case2: "?qX = qOp delta qinp qbinp"
hence good_qinp: "qGoodInp qinp \<and> qGoodBinp qbinp" using good_qX by simp
let ?inp = "asInp qinp" let ?binp = "asBinp qbinp"
have "qinp %= pickInp ?inp \<and> qbinp %%= pickBinp ?binp"
using good_qinp pickInp_asInp alphaInp_sym pickBinp_asBinp alphaBinp_sym by blast
hence "qOp delta qinp qbinp #= qOp delta (pickInp ?inp) (pickBinp ?binp)" by simp
hence "asTerm (qOp delta qinp qbinp) = Op delta ?inp ?binp"
unfolding Op_def using Case2 good_qX by (auto simp add: asTerm_equal_iff_alpha)
hence "X = Op delta ?inp ?binp" using X Case2 by auto
moreover have "goodInp ?inp \<and> goodBinp ?binp"
using good_qinp qGoodInp_iff_goodInp_asInp qGoodBinp_iff_goodBinp_asBinp by auto
ultimately show ?thesis by blast
qed
qed
theorem abs_nchotomy:
assumes "goodAbs A"
shows "\<exists> xs x X. good X \<and> A = Abs xs x X"
by (metis Abs_asTerm_asAbs_qAbs asAbs_pick assms
goodAbs_imp_qGoodAbs_pick qGoodAbs.elims(2) qGood_iff_good_asTerm)
lemmas good_freeCons =
Op_inj Var_diff_Op Op_diff_Var
-subsection {* Properties lifted from quasi-terms to terms *}
-
-subsubsection {* Simplification rules *}
+subsection \<open>Properties lifted from quasi-terms to terms\<close>
+
+subsubsection \<open>Simplification rules\<close>
theorem swap_Var_simp[simp]:
"((Var xs x) #[y1 \<and> y2]_ys) = Var xs (x @xs[y1 \<and> y2]_ys)"
by (metis QuasiTerms_Swap_Fresh.qSwapAll_simps(1) Var_def asTerm_qSwap_swap qItem_simps(9))
lemma swap_Op_simp[simp]:
assumes "goodInp inp" "goodBinp binp"
shows "((Op delta inp binp) #[x1 \<and> x2]_xs) =
Op delta (inp %[x1 \<and> x2]_xs) (binp %%[x1 \<and> x2]_xs)"
by (metis Op_asInp_asTerm_qOp Op_def asTerm_qSwap_swap assms(1) assms(2) goodBinp_imp_qGoodBinp_pickBinp goodInp_imp_qGoodInp_pickInp qGood_qGoodInp qSwapBinp_preserves_qGoodBinp
qSwapInp_preserves_qGoodInp qSwap_qSwapInp swapBinp_def2 swapInp_def2)
lemma swapAbs_simp[simp]:
assumes "good X"
shows "((Abs xs x X) $[y1 \<and> y2]_ys) = Abs xs (x @xs[y1 \<and> y2]_ys) (X #[y1 \<and> y2]_ys)"
by (metis Abs_asTerm_asAbs_qAbs Abs_preserves_good alphaAbs_preserves_qGoodAbs2 asAbs_qSwapAbs_swapAbs assms goodAbs_imp_qGoodAbs_pick good_imp_qGood_pick local.Abs_def
local.swap_def qAbs_pick_Abs qSwapAbs.simps qSwap_preserves_qGood1)
lemmas good_swapAll_simps =
swap_Op_simp swapAbs_simp
theorem fresh_Var_simp[simp]:
"fresh ys y (Var xs x :: ('index,'bindex,'varSort,'var,'opSym)term) \<longleftrightarrow>
(ys \<noteq> xs \<or> y \<noteq> x)"
by (simp add: Var_def fresh_asTerm_qFresh)
lemma fresh_Op_simp[simp]:
assumes "goodInp inp" "goodBinp binp"
shows
"fresh xs x (Op delta inp binp) \<longleftrightarrow>
(freshInp xs x inp \<and> freshBinp xs x binp)"
by (metis Op_def Op_preserves_good assms(1) assms(2) freshBinp_def2
freshInp_def2 fresh_asTerm_qFresh qFresh_qFreshInp qGood_iff_good_asTerm)
lemma freshAbs_simp[simp]:
assumes "good X"
shows "freshAbs ys y (Abs xs x X) \<longleftrightarrow> (ys = xs \<and> y = x \<or> fresh ys y X)"
proof-
let ?fr = "fresh ys y" let ?qfr = "qFresh ys y"
let ?frA = "freshAbs ys y" let ?qfrA = "qFreshAbs ys y"
have "qGood (pick X)" using assms by(auto simp add: good_imp_qGood_pick)
hence good_qAbs_pick_X: "qGoodAbs (qAbs xs x (pick X))"
using assms good_imp_qGood_pick by auto
(* *)
have "?frA (Abs xs x X) = ?qfrA ((pick o asAbs) (qAbs xs x (pick X)))"
unfolding freshAbs_def Abs_def by simp
also
{have "(pick o asAbs) (qAbs xs x (pick X)) $= qAbs xs x (pick X)"
using good_qAbs_pick_X pick_asAbs by fastforce
hence "?qfrA ((pick o asAbs) (qAbs xs x (pick X))) = ?qfrA (qAbs xs x (pick X))"
using good_qAbs_pick_X qFreshAbs_preserves_alphaAbs by blast
}
also have "?qfrA(qAbs xs x (pick X)) = (ys = xs \<and> y = x \<or> ?qfr (pick X))" by simp
also have "\<dots> = (ys = xs \<and> y = x \<or> ?fr X)" unfolding fresh_def by simp
finally show ?thesis .
qed
lemmas good_freshAll_simps =
fresh_Op_simp freshAbs_simp
theorem skel_Var_simp[simp]:
"skel (Var xs x) = Branch Map.empty Map.empty"
by (metis alpha_qSkel pick_Var_qVar qSkel.simps(1) skel_def)
lemma skel_Op_simp[simp]:
assumes "goodInp inp" and "goodBinp binp"
shows "skel (Op delta inp binp) = Branch (skelInp inp) (skelBinp binp)"
by (metis (no_types, lifting) alpha_qSkel assms
qOp_pickInp_pick_Op qSkel_qSkelInp skelBinp_def skelInp_def skel_def)
lemma skelAbs_simp[simp]:
assumes "good X"
shows "skelAbs (Abs xs x X) = Branch (\<lambda>i. Some (skel X)) Map.empty"
by (metis alphaAll_qSkelAll assms qAbs_pick_Abs qSkelAbs.simps skelAbs_def skel_def)
lemmas good_skelAll_simps =
skel_Op_simp skelAbs_simp
lemma psubst_Var:
assumes "goodEnv rho"
shows "((Var xs x) #[rho]) =
(case rho xs x of None \<Rightarrow> Var xs x
|Some X \<Rightarrow> X)"
proof-
let ?X = "Var xs x" let ?qX = "qVar xs x"
let ?qrho = "pickE rho"
have good_qX: "qGood ?qX" using assms by simp
moreover have good_qrho: "qGoodEnv ?qrho" using assms goodEnv_imp_qGoodEnv_pickE by auto
ultimately have good_qXrho: "qGood (?qX #[[?qrho]])"
using assms qPsubst_preserves_qGood by(auto simp del: qGoodAll_simps qPsubst.simps)
(* *)
have "(?X #[rho]) = asTerm ((pick (asTerm ?qX)) #[[?qrho]])"
unfolding Var_def psubst_def by simp
also
{have "?qX #= pick (asTerm ?qX)" using good_qX pick_asTerm alpha_sym by fastforce
hence "(?qX #[[?qrho]]) #= ((pick (asTerm ?qX)) #[[?qrho]])"
using good_qrho good_qX qPsubst_preserves_alpha1[of _ ?qX] by fastforce
hence "asTerm ((pick (asTerm ?qX)) #[[?qrho]]) = asTerm (?qX #[[?qrho]])"
using good_qXrho asTerm_equal_iff_alpha[of "?qX #[[?qrho]]"] by blast
}
also have "asTerm (?qX #[[?qrho]]) =
asTerm (case ?qrho xs x of None \<Rightarrow> qVar xs x
|Some qY \<Rightarrow> qY)" unfolding Var_def by simp
finally have 1: "(?X #[rho]) = asTerm (case ?qrho xs x of None \<Rightarrow> qVar xs x
|Some qY \<Rightarrow> qY)" .
show ?thesis
proof(cases "rho xs x")
assume Case1: "rho xs x = None"
hence "?qrho xs x = None" unfolding pickE_def lift_def by simp
thus ?thesis using 1 Case1 unfolding Var_def by simp
next
fix X assume Case2: "rho xs x = Some X"
hence "good X" using assms unfolding goodEnv_def liftAll_def by auto
hence "asTerm (pick X) = X" using asTerm_pick by auto
moreover have qrho: "?qrho xs x = Some (pick X)"
using Case2 unfolding pickE_def lift_def by simp
ultimately show ?thesis using 1 Case2 unfolding Var_def by simp
qed
qed
corollary psubst_Var_simp1[simp]:
assumes "goodEnv rho" and "rho xs x = None"
shows "((Var xs x) #[rho]) = Var xs x"
using assms by(simp add: psubst_Var)
corollary psubst_Var_simp2[simp]:
assumes "goodEnv rho" and "rho xs x = Some X"
shows "((Var xs x) #[rho]) = X"
using assms by(simp add: psubst_Var)
lemma psubst_Op_simp[simp]:
assumes good_inp: "goodInp inp" "goodBinp binp"
and good_rho: "goodEnv rho"
shows
"((Op delta inp binp) #[rho]) = Op delta (inp %[rho]) (binp %%[rho])"
proof-
let ?qrho = "pickE rho"
let ?sbs = "psubst rho" let ?qsbs = "qPsubst ?qrho"
let ?sbsI = "psubstInp rho" let ?qsbsI = "qPsubstInp ?qrho"
let ?sbsB = "psubstBinp rho" let ?qsbsB = "qPsubstBinp ?qrho"
let ?op = "Op delta" let ?qop = "qOp delta"
have good_qop_pickInp_inp: "qGood (?qop (pickInp inp) (pickBinp binp))"
using good_inp goodInp_imp_qGoodInp_pickInp
goodBinp_imp_qGoodBinp_pickBinp by auto
hence "qGood ((pick o asTerm) (?qop (pickInp inp) (pickBinp binp)))"
using good_imp_qGood_pick qGood_iff_good_asTerm by fastforce
moreover have good_qrho: "qGoodEnv ?qrho"
using good_rho goodEnv_imp_qGoodEnv_pickE by auto
ultimately have good: "qGood (?qsbs((pick o asTerm) (?qop (pickInp inp) (pickBinp binp))))"
using qPsubst_preserves_qGood by auto
(* *)
have "?sbs (?op inp binp) =
asTerm (?qsbs ((pick o asTerm) (?qop (pickInp inp) (pickBinp binp))))"
unfolding psubst_def Op_def by simp
also
{have "(pick o asTerm) (?qop (pickInp inp) (pickBinp binp)) #=
?qop (pickInp inp) (pickBinp binp)"
using good_qop_pickInp_inp pick_asTerm by fastforce
hence "?qsbs((pick o asTerm) (?qop (pickInp inp) (pickBinp binp))) #=
?qsbs(?qop (pickInp inp) (pickBinp binp))"
using good_qop_pickInp_inp good_qrho qPsubst_preserves_alpha1 by fastforce
moreover have "?qsbs (?qop (pickInp inp) (pickBinp binp)) =
?qop (?qsbsI (pickInp inp)) (?qsbsB (pickBinp binp))" by simp
moreover
{have "?qsbsI (pickInp inp) %= pickInp (?sbsI inp) \<and>
?qsbsB (pickBinp binp) %%= pickBinp (?sbsB binp)"
using good_rho good_inp pickInp_psubstInp_qPsubstInp[of inp rho]
pickBinp_psubstBinp_qPsubstBinp[of binp rho] alphaInp_sym alphaBinp_sym by auto
hence "?qop (?qsbsI (pickInp inp)) (?qsbsB (pickBinp binp)) #=
?qop (pickInp (?sbsI inp)) (pickBinp (?sbsB binp))" by simp
}
ultimately have "?qsbs((pick o asTerm) (?qop (pickInp inp) (pickBinp binp))) #=
?qop (pickInp (?sbsI inp)) (pickBinp (?sbsB binp))"
using good alpha_trans by force
hence "asTerm (?qsbs((pick o asTerm) (?qop (pickInp inp) (pickBinp binp)))) =
asTerm (?qop (pickInp (?sbsI inp)) (pickBinp (?sbsB binp)))"
using good by (auto simp add: asTerm_equal_iff_alpha)
}
also have "asTerm (?qop (pickInp (?sbsI inp)) (pickBinp (?sbsB binp))) =
?op (?sbsI inp) (?sbsB binp)" unfolding Op_def by simp
finally show ?thesis .
qed
lemma psubstAbs_simp[simp]:
assumes good_X: "good X" and good_rho: "goodEnv rho" and
x_fresh_rho: "freshEnv xs x rho"
shows "((Abs xs x X) $[rho]) = Abs xs x (X #[rho])"
proof-
let ?qrho = "pickE rho"
let ?sbs = "psubst rho" let ?qsbs = "qPsubst ?qrho"
let ?sbsA = "psubstAbs rho" let ?qsbsA = "qPsubstAbs ?qrho"
have good_qrho: "qGoodEnv ?qrho"
using good_rho goodEnv_imp_qGoodEnv_pickE by auto
have good_pick_X: "qGood (pick X)" using good_X good_imp_qGood_pick by auto
hence good_qsbs_pick_X: "qGood(?qsbs (pick X))"
using good_qrho qPsubst_preserves_qGood by auto
have good_qAbs_pick_X: "qGoodAbs (qAbs xs x (pick X))"
using good_X good_imp_qGood_pick by auto
hence "qGoodAbs ((pick o asAbs) (qAbs xs x (pick X)))"
using goodAbs_imp_qGoodAbs_pick qGoodAbs_iff_goodAbs_asAbs by fastforce
hence good: "qGoodAbs (?qsbsA ((pick o asAbs) (qAbs xs x (pick X))))"
using good_qrho qPsubstAbs_preserves_qGoodAbs by auto
have x_fresh_qrho: "qFreshEnv xs x ?qrho"
using x_fresh_rho unfolding freshEnv_def2 by auto
(* *)
have "?sbsA (Abs xs x X) = asAbs (?qsbsA ((pick o asAbs) (qAbs xs x (pick X))))"
unfolding psubstAbs_def Abs_def by simp
also
{have "(pick o asAbs) (qAbs xs x (pick X)) $= qAbs xs x (pick X)"
using good_qAbs_pick_X pick_asAbs by fastforce
hence "?qsbsA((pick o asAbs) (qAbs xs x (pick X))) $= ?qsbsA(qAbs xs x (pick X))"
using good_qAbs_pick_X good_qrho qPsubstAbs_preserves_alphaAbs1 by force
moreover have "?qsbsA(qAbs xs x (pick X)) $= qAbs xs x (?qsbs (pick X))"
using qFresh_qPsubst_commute_qAbs good_pick_X good_qrho x_fresh_qrho by auto
moreover
{have "?qsbs (pick X) #= pick (?sbs X)"
using good_rho good_X pick_psubst_qPsubst alpha_sym by fastforce
hence "qAbs xs x (?qsbs (pick X)) $= qAbs xs x (pick (?sbs X))"
using good_qsbs_pick_X qAbs_preserves_alpha by fastforce
}
ultimately
have "?qsbsA((pick o asAbs) (qAbs xs x (pick X))) $= qAbs xs x (pick (?sbs X))"
using good alphaAbs_trans by blast
hence "asAbs (?qsbsA((pick o asAbs) (qAbs xs x (pick X)))) =
asAbs (qAbs xs x (pick (?sbs X)))"
using good asAbs_equal_iff_alphaAbs by auto
}
also have "asAbs (qAbs xs x (pick (?sbs X))) = Abs xs x (?sbs X)"
unfolding Abs_def by simp
finally show ?thesis .
qed
lemmas good_psubstAll_simps =
psubst_Var_simp1 psubst_Var_simp2
psubst_Op_simp psubstAbs_simp
theorem getEnv_idEnv[simp]: "idEnv xs x = None"
unfolding idEnv_def by simp
lemma getEnv_updEnv[simp]:
"(rho [x \<leftarrow> X]_xs) ys y = (if ys = xs \<and> y = x then Some X else rho ys y)"
unfolding updEnv_def by auto
theorem getEnv_updEnv1:
"ys \<noteq> xs \<or> y \<noteq> x \<Longrightarrow> (rho [x \<leftarrow> X]_xs) ys y = rho ys y"
by auto
theorem getEnv_updEnv2:
"(rho [x \<leftarrow> X]_xs) xs x = Some X"
by auto
lemma subst_Var_simp1[simp]:
assumes "good Y"
and "ys \<noteq> xs \<or> y \<noteq> x"
shows "((Var xs x) #[Y / y]_ys) = Var xs x"
using assms unfolding subst_def by auto
lemma subst_Var_simp2[simp]:
assumes "good Y"
shows "((Var xs x) #[Y / x]_xs) = Y"
using assms unfolding subst_def by auto
lemma subst_Op_simp[simp]:
assumes "good Y"
and "goodInp inp" and "goodBinp binp"
shows
"((Op delta inp binp) #[Y / y]_ys) =
Op delta (inp %[Y / y]_ys) (binp %%[Y / y]_ys)"
using assms unfolding subst_def substInp_def substBinp_def by auto
lemma substAbs_simp[simp]:
assumes good: "good Y" and good_X: "good X" and
x_dif_y: "xs \<noteq> ys \<or> x \<noteq> y" and x_fresh: "fresh xs x Y"
shows "((Abs xs x X) $[Y / y]_ys) = Abs xs x (X #[Y / y]_ys)"
proof-
have "freshEnv xs x (idEnv [y \<leftarrow> Y]_ys)" unfolding freshEnv_def liftAll_def
using x_dif_y x_fresh by auto
thus ?thesis using assms unfolding subst_def substAbs_def by auto
qed
lemmas good_substAll_simps =
subst_Var_simp1 subst_Var_simp2
subst_Op_simp substAbs_simp
theorem vsubst_Var_simp[simp]:
"((Var xs x) #[y1 // y]_ys) = Var xs (x @xs[y1 / y]_ys)"
unfolding vsubst_def
apply(case_tac "ys = xs \<and> y = x") by simp_all
lemma vsubst_Op_simp[simp]:
assumes "goodInp inp" and "goodBinp binp"
shows
"((Op delta inp binp) #[y1 // y]_ys) =
Op delta (inp %[y1 // y]_ys) (binp %%[y1 // y]_ys)"
using assms unfolding vsubst_def vsubstInp_def vsubstBinp_def by auto
lemma vsubstAbs_simp[simp]:
assumes "good X" and
"xs \<noteq> ys \<or> x \<notin> {y,y1}"
shows "((Abs xs x X) $[y1 // y]_ys) = Abs xs x (X #[y1 // y]_ys)"
using assms unfolding vsubst_def vsubstAbs_def by auto
lemmas good_vsubstAll_simps =
vsubst_Op_simp vsubstAbs_simp
lemmas good_allOpers_simps =
good_swapAll_simps
good_freshAll_simps
good_skelAll_simps
good_psubstAll_simps
good_substAll_simps
good_vsubstAll_simps
-subsubsection {* The ability to pick fresh variables *}
+subsubsection \<open>The ability to pick fresh variables\<close>
lemma single_non_fresh_ordLess_var:
"good X \<Longrightarrow> |{x. \<not> fresh xs x X}| <o |UNIV :: 'var set|"
unfolding fresh_def
by(auto simp add: good_imp_qGood_pick single_non_qFresh_ordLess_var)
lemma single_non_freshAbs_ordLess_var:
"goodAbs A \<Longrightarrow> |{x. \<not> freshAbs xs x A}| <o |UNIV :: 'var set|"
unfolding freshAbs_def
by(auto simp add: goodAbs_imp_qGoodAbs_pick single_non_qFreshAbs_ordLess_var)
lemma obtain_fresh1:
fixes XS::"('index,'bindex,'varSort,'var,'opSym)term set" and
Rho::"('index,'bindex,'varSort,'var,'opSym)env set" and rho
assumes Vvar: "|V| <o |UNIV :: 'var set| \<or> finite V" and XSvar: "|XS| <o |UNIV :: 'var set| \<or> finite XS" and
good: "\<forall> X \<in> XS. good X" and
Rhovar: "|Rho| <o |UNIV :: 'var set| \<or> finite Rho" and RhoGood: "\<forall> rho \<in> Rho. goodEnv rho"
shows
"\<exists> z. z \<notin> V \<and>
(\<forall> X \<in> XS. fresh xs z X) \<and>
(\<forall> rho \<in> Rho. freshEnv xs z rho)"
proof-
let ?qXS = "pick ` XS" let ?qRho = "pickE ` Rho"
have "|?qXS| \<le>o |XS|" using card_of_image by auto
hence 1: "|?qXS| <o |UNIV :: 'var set| \<or> finite ?qXS"
using ordLeq_ordLess_trans card_of_ordLeq_finite XSvar by blast
have "|?qRho| \<le>o |Rho|" using card_of_image by auto
hence 2: "|?qRho| <o |UNIV :: 'var set| \<or> finite ?qRho"
using ordLeq_ordLess_trans card_of_ordLeq_finite Rhovar by blast
have 3: "\<forall> qX \<in> ?qXS. qGood qX" using good good_imp_qGood_pick by auto
have "\<forall> qrho \<in> ?qRho. qGoodEnv qrho" using RhoGood goodEnv_imp_qGoodEnv_pickE by auto
then obtain z where
"z \<notin> V \<and> (\<forall> qX \<in> ?qXS. qFresh xs z qX) \<and>
(\<forall> qrho \<in> ?qRho. qFreshEnv xs z qrho)"
using Vvar 1 2 3 obtain_qFreshEnv[of V ?qXS ?qRho] by fastforce
thus ?thesis unfolding fresh_def freshEnv_def2 by auto
qed
lemma obtain_fresh:
fixes V::"'var set" and
XS::"('index,'bindex,'varSort,'var,'opSym)term set" and
AS::"('index,'bindex,'varSort,'var,'opSym)abs set" and
Rho::"('index,'bindex,'varSort,'var,'opSym)env set"
assumes Vvar: "|V| <o |UNIV :: 'var set| \<or> finite V" and
XSvar: "|XS| <o |UNIV :: 'var set| \<or> finite XS" and
ASvar: "|AS| <o |UNIV :: 'var set| \<or> finite AS" and
Rhovar: "|Rho| <o |UNIV :: 'var set| \<or> finite Rho" and
good: "\<forall> X \<in> XS. good X" and
ASGood: "\<forall> A \<in> AS. goodAbs A" and
RhoGood: "\<forall> rho \<in> Rho. goodEnv rho"
shows
"\<exists> z. z \<notin> V \<and>
(\<forall> X \<in> XS. fresh xs z X) \<and>
(\<forall> A \<in> AS. freshAbs xs z A) \<and>
(\<forall> rho \<in> Rho. freshEnv xs z rho)"
proof-
have XS: "|XS| <o |UNIV :: 'var set|" and AS: "|AS| <o |UNIV :: 'var set|"
using XSvar ASvar finite_ordLess_var by auto
let ?phi = "% A Y. (good Y \<and> (EX ys y. A = Abs ys y Y))"
{fix A assume "A \<in> AS"
hence "goodAbs A" using ASGood by simp
hence "EX Y. ?phi A Y" using abs_nchotomy[of A] by auto
}
then obtain f where 1: "ALL A : AS. ?phi A (f A)"
using bchoice[of AS ?phi] by auto
let ?YS = "f ` AS"
have 2: "ALL Y : ?YS. good Y" using 1 by simp
have "|?YS| <=o |AS|" using card_of_image by auto
hence "|?YS| <o |UNIV :: 'var set|"
using AS ordLeq_ordLess_trans by blast
hence "|XS Un ?YS| <o |UNIV :: 'var set|"
using XS by (auto simp add: var_infinite_INNER card_of_Un_ordLess_infinite)
then obtain z where z: "z \<notin> V"
and XSYS: "\<forall> X \<in> XS Un ?YS. fresh xs z X"
and Rho: "\<forall> rho \<in> Rho. freshEnv xs z rho"
using Vvar Rhovar good 2 RhoGood
obtain_fresh1[of V "XS Un ?YS" Rho xs] by blast
moreover
{fix A
obtain Y where Y_def: "Y = f A" by blast
assume "A : AS"
hence "fresh xs z Y" unfolding Y_def using XSYS by simp
moreover obtain ys y where Y: "good Y" and A: "A = Abs ys y Y"
- unfolding Y_def using `A : AS` 1 by auto
+ unfolding Y_def using \<open>A : AS\<close> 1 by auto
ultimately have "freshAbs xs z A" unfolding A using z by auto
}
ultimately show ?thesis by auto
qed
-subsubsection {* Compositionality *}
+subsubsection \<open>Compositionality\<close>
lemma swap_ident[simp]:
assumes "good X"
shows "(X #[x \<and> x]_xs) = X"
using assms asTerm_pick qSwap_ident unfolding swap_def by auto
lemma swap_compose:
assumes good_X: "good X"
shows "((X #[x \<and> y]_zs) #[x' \<and> y']_zs') =
((X #[x' \<and> y']_zs') #[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]_zs)"
using assms qSwap_compose[of _ _ _ _ _ _ "pick X"] by(auto simp add: double_swap_qSwap)
lemma swap_commute:
"\<lbrakk>good X; zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {}\<rbrakk> \<Longrightarrow>
((X #[x \<and> y]_zs) #[x' \<and> y']_zs') = ((X #[x' \<and> y']_zs') #[x \<and> y]_zs)"
using swap_compose[of X zs' x' y' zs x y] by(auto simp add: sw_def)
lemma swap_involutive[simp]:
assumes good_X: "good X"
shows "((X #[x \<and> y]_zs) #[x \<and> y]_zs) = X"
using assms asTerm_pick[of X] by (auto simp add: double_swap_qSwap)
theorem swap_sym: "(X #[x \<and> y]_zs) = (X #[y \<and> x]_zs)"
unfolding swap_def by(auto simp add: qSwap_sym)
lemma swap_involutive2[simp]:
assumes "good X"
shows "((X #[x \<and> y]_zs) #[y \<and> x]_zs) = X"
using assms by(simp add: swap_sym)
lemma swap_preserves_fresh[simp]:
assumes "good X"
shows "fresh xs (x @xs[y1 \<and> y2]_ys) (X #[y1 \<and> y2]_ys) = fresh xs x X"
unfolding fresh_def[of _ _ X] using assms qSwap_preserves_qFresh[of _ _ _ _ _ "pick X"]
by(auto simp add: fresh_swap_qFresh_qSwap)
lemma swap_preserves_fresh_distinct:
assumes "good X" and
"xs \<noteq> ys \<or> x \<notin> {y1,y2}"
shows "fresh xs x (X #[y1 \<and> y2]_ys) = fresh xs x X"
unfolding fresh_def[of _ _ X] using assms
by(auto simp: fresh_swap_qFresh_qSwap qSwap_preserves_qFresh_distinct)
lemma fresh_swap_exchange1:
assumes "good X"
shows "fresh xs x2 (X #[x1 \<and> x2]_xs) = fresh xs x1 X"
unfolding fresh_def[of _ _ X]
using assms by(auto simp: fresh_swap_qFresh_qSwap qFresh_qSwap_exchange1)
lemma fresh_swap_exchange2:
assumes "good X" and "{x1,x2} \<subseteq> var xs"
shows "fresh xs x2 (X #[x2 \<and> x1]_xs) = fresh xs x1 X"
using assms by(simp add: fresh_swap_exchange1 swap_sym)
(* Note: the lemmas swap_preserves_fresh_distinct, fresh_swap_exchange1 and
fresh_swap_exchange2 do cover all possibilities of simplifying an
expression of the form "fresh ys y (X #[x2 \<and> x1]_xs)". *)
lemma fresh_swap_id[simp]:
assumes "good X" and "fresh xs x1 X" "fresh xs x2 X"
shows "(X #[x1 \<and> x2]_xs) = X"
by (metis (no_types, lifting) assms alpha_imp_asTerm_equal alpha_qFresh_qSwap_id asTerm_pick
fresh_def good_imp_qGood_pick local.swap_def qSwap_preserves_qGood1)
lemma freshAbs_swapAbs_id[simp]:
assumes "goodAbs A" "freshAbs xs x1 A" "freshAbs xs x2 A"
shows "(A $[x1 \<and> x2]_xs) = A"
using assms
by (meson alphaAbs_qFreshAbs_qSwapAbs_id alphaAll_trans freshAbs_def goodAbs_imp_qGoodAbs_pick
pick_alphaAbs_iff_equal pick_swapAbs_qSwapAbs swapAbs_preserves_good)
lemma fresh_swap_compose:
assumes "good X" "fresh xs y X" "fresh xs z X"
shows "((X #[y \<and> x]_xs) #[z \<and> y]_xs) = (X #[z \<and> x]_xs)"
using assms by (simp add: sw_def swap_compose)
lemma skel_swap:
assumes "good X"
shows "skel (X #[x1 \<and> x2]_xs) = skel X"
using assms by (metis alpha_qSkel pick_swap_qSwap qSkel_qSwap skel_def)
-subsubsection {* Compositionality for environments *}
+subsubsection \<open>Compositionality for environments\<close>
lemma swapEnv_ident[simp]:
assumes "goodEnv rho"
shows "(rho &[x \<and> x]_xs) = rho"
using assms unfolding swapEnv_defs lift_def
by (intro ext) (auto simp: option.case_eq_if)
lemma swapEnv_compose:
assumes good: "goodEnv rho"
shows "((rho &[x \<and> y]_zs) &[x' \<and> y']_zs') =
((rho &[x' \<and> y']_zs') &[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]_zs)"
proof(rule ext)+
let ?xsw = "x @zs[x' \<and> y']_zs'" let ?ysw = "y @zs[x' \<and> y']_zs'"
let ?xswsw = "?xsw @zs[x' \<and> y']_zs'" let ?yswsw = "?ysw @zs[x' \<and> y']_zs'"
let ?rhosw1 = "rho &[x \<and> y]_zs" let ?rhosw11 = "?rhosw1 &[x' \<and> y']_zs'"
let ?rhosw2 = "rho &[x' \<and> y']_zs'" let ?rhosw22 = "?rhosw2 &[?xsw \<and> ?ysw]_zs"
let ?Sw1 = "\<lambda>X. (X #[x \<and> y]_zs)" let ?Sw11 = "\<lambda>X. ((?Sw1 X) #[x' \<and> y']_zs')"
let ?Sw2 = "\<lambda>X. (X #[x' \<and> y']_zs')" let ?Sw22 = "\<lambda>X. ((?Sw2 X) #[?xsw \<and> ?ysw]_zs)"
fix us u
let ?usw1 = "u @us [x' \<and> y']_zs'" let ?usw11 = "?usw1 @us [x \<and> y]_zs"
let ?usw2 = "u @us [?xsw \<and> ?ysw]_zs" let ?usw22 = "?usw2 @us [x' \<and> y']_zs'"
have "(?xsw @zs[x' \<and> y']_zs') = x" and "(?ysw @zs[x' \<and> y']_zs') = y" by auto
have "?usw22 = (?usw1 @us[?xswsw \<and> ?yswsw]_zs)" using sw_compose .
hence *: "?usw22 = ?usw11" by simp
show "?rhosw11 us u = ?rhosw22 us u"
proof(cases "rho us ?usw11")
case None
hence "?rhosw11 us u = None" unfolding swapEnv_defs lift_def by simp
also have "\<dots> = ?rhosw22 us u"
using None unfolding * swapEnv_defs lift_def by simp
finally show ?thesis .
next
case (Some X)
hence "good X" using good unfolding goodEnv_def liftAll_def by simp
have "?rhosw11 us u = Some(?Sw11 X)" using Some unfolding swapEnv_defs lift_def by simp
also have "?Sw11 X = ?Sw22 X"
- using `good X` by(rule swap_compose)
+ using \<open>good X\<close> by(rule swap_compose)
also have "Some(?Sw22 X) = ?rhosw22 us u"
using Some unfolding * swapEnv_defs lift_def by simp
finally show ?thesis .
qed
qed
lemma swapEnv_commute:
"\<lbrakk>goodEnv rho; {x,y} \<subseteq> var zs; zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {}\<rbrakk> \<Longrightarrow>
((rho &[x \<and> y]_zs) &[x' \<and> y']_zs') = ((rho &[x' \<and> y']_zs') &[x \<and> y]_zs)"
using swapEnv_compose[of rho zs' x' y' zs x y] by(auto simp add: sw_def)
lemma swapEnv_involutive[simp]:
assumes "goodEnv rho"
shows "((rho &[x \<and> y]_zs) &[x \<and> y]_zs) = rho"
using assms unfolding swapEnv_defs lift_def
by (fastforce simp: option.case_eq_if)
theorem swapEnv_sym: "(rho &[x \<and> y]_zs) = (rho &[y \<and> x]_zs)"
proof(intro ext)
fix us u
have *: "(u @us[x \<and> y]_zs) = (u @us[y \<and> x]_zs)" using sw_sym by fastforce
show "(rho &[x \<and> y]_zs) us u = (rho &[y \<and> x]_zs) us u"
unfolding swapEnv_defs lift_def *
by(cases "rho us (u @us[y \<and> x]_zs)") (auto simp: swap_sym)
qed
lemma swapEnv_involutive2[simp]:
assumes good: "goodEnv rho"
shows "((rho &[x \<and> y]_zs) &[y \<and> x]_zs) = rho"
using assms by(simp add: swapEnv_sym)
lemma swapEnv_preserves_freshEnv[simp]:
assumes good: "goodEnv rho"
shows "freshEnv xs (x @xs[y1 \<and> y2]_ys) (rho &[y1 \<and> y2]_ys) = freshEnv xs x rho"
proof-
let ?xsw = "x @xs[y1 \<and> y2]_ys" let ?xswsw = "?xsw @xs[y1 \<and> y2]_ys"
let ?rhosw = "rho &[y1 \<and> y2]_ys"
let ?Left = "freshEnv xs ?xsw ?rhosw"
let ?Right = "freshEnv xs x rho"
have "(?rhosw xs ?xsw = None) = (rho xs x = None)"
unfolding freshEnv_def swapEnv_defs
by(simp add: lift_None sw_involutive)
moreover
have "(\<forall> zs z' Z'. ?rhosw zs z' = Some Z' \<longrightarrow> fresh xs ?xsw Z') =
(\<forall> zs z Z. rho zs z = Some Z \<longrightarrow> fresh xs x Z)"
proof(rule iff_allI, auto)
fix zs z Z assume *: "\<forall> z' Z'. ?rhosw zs z' = Some Z' \<longrightarrow> fresh xs ?xsw Z'"
and **: "rho zs z = Some Z" let ?z' = "z @zs[y1 \<and> y2]_ys" let ?Z' = "Z #[y1 \<and> y2]_ys"
have "?rhosw zs ?z' = Some ?Z'" using ** unfolding swapEnv_defs lift_def
by(simp add: sw_involutive)
hence "fresh xs ?xsw ?Z'" using * by simp
moreover have "good Z" using ** good unfolding goodEnv_def liftAll_def by simp
ultimately show "fresh xs x Z" using swap_preserves_fresh by auto
next
fix zs z' Z'
assume *: "\<forall>z Z. rho zs z = Some Z \<longrightarrow> fresh xs x Z" and **: "?rhosw zs z' = Some Z'"
let ?z = "z' @zs[y1 \<and> y2]_ys"
obtain Z where rho: "rho zs ?z = Some Z" and Z': "Z' = Z #[y1 \<and> y2]_ys"
using ** unfolding swapEnv_defs lift_def by(cases "rho zs ?z", auto)
hence "fresh xs x Z" using * by simp
moreover have "good Z" using rho good unfolding goodEnv_def liftAll_def by simp
ultimately show "fresh xs ?xsw Z'" unfolding Z' using swap_preserves_fresh by auto
qed
ultimately show ?thesis unfolding freshEnv_def swapEnv_defs
unfolding liftAll_def by simp
qed
lemma swapEnv_preserves_freshEnv_distinct:
assumes "goodEnv rho" and
"xs \<noteq> ys \<or> x \<notin> {y1,y2}"
shows "freshEnv xs x (rho &[y1 \<and> y2]_ys) = freshEnv xs x rho"
by (metis assms sw_simps3 swapEnv_preserves_freshEnv)
lemma freshEnv_swapEnv_exchange1:
assumes "goodEnv rho"
shows "freshEnv xs x2 (rho &[x1 \<and> x2]_xs) = freshEnv xs x1 rho"
by (metis assms sw_simps1 swapEnv_preserves_freshEnv)
lemma freshEnv_swapEnv_exchange2:
assumes "goodEnv rho"
shows "freshEnv xs x2 (rho &[x2 \<and> x1]_xs) = freshEnv xs x1 rho"
using assms by(simp add: freshEnv_swapEnv_exchange1 swapEnv_sym)
lemma freshEnv_swapEnv_id[simp]:
assumes good: "goodEnv rho" and
fresh: "freshEnv xs x1 rho" "freshEnv xs x2 rho"
shows "(rho &[x1 \<and> x2]_xs) = rho"
proof(intro ext)
fix us u
let ?usw = "u @us[x1 \<and> x2]_xs" let ?rhosw = "rho &[x1 \<and> x2]_xs"
let ?Sw = "\<lambda> X. (X #[x1 \<and> x2]_xs)"
show "?rhosw us u = rho us u"
proof(cases "rho us u")
case None
hence "rho us ?usw = None" using fresh unfolding freshEnv_def sw_def by auto
hence "?rhosw us u = None" unfolding swapEnv_defs lift_def by auto
with None show ?thesis by simp
next
case (Some X)
moreover have "?usw = u" using fresh Some unfolding freshEnv_def sw_def by auto
ultimately have "?rhosw us u = Some (?Sw X)" unfolding swapEnv_defs lift_def by auto
moreover
{have "good X" using Some good unfolding goodEnv_def liftAll_def by auto
moreover have "fresh xs x1 X" and "fresh xs x2 X"
using Some fresh unfolding freshEnv_def liftAll_def by auto
ultimately have "?Sw X = X" by simp
}
ultimately show ?thesis using Some by simp
qed
qed
lemma freshEnv_swapEnv_compose:
assumes good: "goodEnv rho" and
fresh: "freshEnv xs y rho" "freshEnv xs z rho"
shows "((rho &[y \<and> x]_xs) &[z \<and> y]_xs) = (rho &[z \<and> x]_xs)"
by (simp add: fresh good sw_def swapEnv_compose)
lemmas good_swapAll_freshAll_otherSimps =
swap_ident swap_involutive swap_involutive2 swap_preserves_fresh fresh_swap_id
freshAbs_swapAbs_id
swapEnv_ident swapEnv_involutive swapEnv_involutive2 swapEnv_preserves_freshEnv freshEnv_swapEnv_id
-subsubsection {* Properties of the relation of being swapped *}
+subsubsection \<open>Properties of the relation of being swapped\<close>
theorem swap_swapped: "(X, X #[x \<and> y]_zs) \<in> swapped"
by(auto simp add: swapped.Refl swapped.Swap)
lemma swapped_preserves_good:
assumes "good X" and "(X,Y) \<in> swapped"
shows "good Y"
using assms(2,1) by (induct rule: swapped.induct) auto
lemma swapped_skel:
assumes "good X" and "(X,Y) \<in> swapped"
shows "skel Y = skel X"
using assms(2,1)
by (induct rule: swapped.induct) (auto simp: swapped_preserves_good skel_swap)
lemma obtain_rep:
assumes GOOD: "good X" and FRESH: "fresh xs x' X"
shows "\<exists> X'. (X,X') \<in> swapped \<and> good X' \<and> Abs xs x X = Abs xs x' X'"
using Abs_swap_fresh FRESH GOOD swap_preserves_good swap_swapped by blast
-subsection {* Induction *}
-
-subsubsection {* Induction lifted from quasi-terms *}
+subsection \<open>Induction\<close>
+
+subsubsection \<open>Induction lifted from quasi-terms\<close>
lemma term_templateInduct[case_names rel Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and phi phiAbs rel
assumes
rel: "\<And> X Y. \<lbrakk>good X; (X,Y) \<in> rel\<rbrakk> \<Longrightarrow> good Y \<and> skel Y = skel X" and
var: "\<And> xs x. phi (Var xs x)" and
op: "\<And> delta inp binp. \<lbrakk>goodInp inp; goodBinp binp; liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
abs: "\<And> xs x X. \<lbrakk>good X; \<And> Y. (X,Y) \<in> rel \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (Abs xs x X)"
shows "(good X \<longrightarrow> phi X) \<and> (goodAbs A \<longrightarrow> phiAbs A)"
proof-
let ?qX = "pick X" let ?qA = "pick A"
let ?qphi = "phi o asTerm" let ?qphiAbs = "phiAbs o asAbs"
let ?qrel = "{(qY, qY')| qY qY'. (asTerm qY, asTerm qY') \<in> rel}"
(* *)
have "(good X \<longrightarrow> qGood ?qX) \<and> (goodAbs A \<longrightarrow> qGoodAbs ?qA)"
using good_imp_qGood_pick goodAbs_imp_qGoodAbs_pick by auto
moreover
have "(good X \<longrightarrow> (?qphi ?qX = phi X)) \<and> (goodAbs A \<longrightarrow> (?qphiAbs ?qA = phiAbs A))"
using asTerm_pick asAbs_pick by fastforce
moreover
have "(qGood ?qX \<longrightarrow> ?qphi ?qX) \<and> (qGoodAbs ?qA \<longrightarrow> ?qphiAbs ?qA)"
proof(induction rule: qGood_qTerm_templateInduct[of ?qrel])
case (Rel qX qY)
thus ?case using qGood_iff_good_asTerm pick_asTerm unfolding skel_def
using rel skel_asTerm_qSkel
by simp (smt qGood_iff_good_asTerm skel_asTerm_qSkel)
next
case (Var xs x)
then show ?case using var unfolding Var_def by simp
next
case (Op delta qinp qbinp)
hence good_qinp: "qGoodInp qinp \<and> qGoodBinp qbinp"
unfolding qGoodInp_def qGoodBinp_def liftAll_def by simp
let ?inp = "asInp qinp" let ?binp = "asBinp qbinp"
have good_inp: "goodInp ?inp \<and> goodBinp ?binp"
using good_qinp qGoodInp_iff_goodInp_asInp qGoodBinp_iff_goodBinp_asBinp by auto
have 1: "Op delta ?inp ?binp = asTerm (qOp delta qinp qbinp)"
using good_qinp Op_asInp_asTerm_qOp by fastforce
{fix i X
assume inp: "?inp i = Some X"
then obtain qX where qinp: "qinp i = Some qX" and X: "X = asTerm qX"
unfolding asInp_def lift_def by(cases "qinp i", auto)
have "qGood qX \<and> phi (asTerm qX)" using qinp Op.IH by (simp add: liftAll_def)
hence "good X \<and> phi X" unfolding X using qGood_iff_good_asTerm by auto
}
moreover
{fix i A
assume binp: "?binp i = Some A"
then obtain qA where qbinp: "qbinp i = Some qA" and A: "A = asAbs qA"
unfolding asBinp_def lift_def by(cases "qbinp i", auto)
have "qGoodAbs qA \<and> phiAbs (asAbs qA)" using qbinp Op.IH by (simp add: liftAll_def)
hence "goodAbs A \<and> phiAbs A" unfolding A using qGoodAbs_iff_goodAbs_asAbs by auto
}
ultimately show ?case
using op[of ?inp ?binp delta] good_inp unfolding 1 liftAll_def by simp
next
case (Abs xs x qX)
- have "good (asTerm qX)" using `qGood qX` qGood_iff_good_asTerm by auto
+ have "good (asTerm qX)" using \<open>qGood qX\<close> qGood_iff_good_asTerm by auto
moreover
{fix Y assume *: "(asTerm qX, Y) \<in> rel"
obtain qY where qY: "qY = pick Y" by blast
- have "good (asTerm qX)" using `qGood qX` qGood_iff_good_asTerm by auto
+ have "good (asTerm qX)" using \<open>qGood qX\<close> qGood_iff_good_asTerm by auto
hence "good Y" using * rel by auto
hence Y: "Y = asTerm qY" unfolding qY using asTerm_pick by auto
have "phi Y" using * Abs.IH unfolding Y by simp
}
ultimately have "phiAbs (Abs xs x (asTerm qX))" using abs by simp
- thus ?case using `qGood qX` Abs_asTerm_asAbs_qAbs by fastforce
+ thus ?case using \<open>qGood qX\<close> Abs_asTerm_asAbs_qAbs by fastforce
qed
(* *)
ultimately show ?thesis by blast
qed
lemma term_rawInduct[case_names Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (Var xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>goodInp inp; goodBinp binp; liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>good X; phi X\<rbrakk> \<Longrightarrow> phiAbs (Abs xs x X)"
shows "(good X \<longrightarrow> phi X) \<and> (goodAbs A \<longrightarrow> phiAbs A)"
by(rule term_templateInduct[of Id], auto simp add: assms)
lemma term_induct[case_names Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and phi phiAbs
assumes
Var: "\<And> xs x. phi (Var xs x)" and
Op: "\<And> delta inp binp. \<lbrakk>goodInp inp; goodBinp binp; liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
Abs: "\<And> xs x X. \<lbrakk>good X;
\<And> Y. (X,Y) \<in> swapped \<Longrightarrow> phi Y;
\<And> Y. \<lbrakk>good Y; skel Y = skel X\<rbrakk> \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (Abs xs x X)"
shows "(good X \<longrightarrow> phi X) \<and> (goodAbs A \<longrightarrow> phiAbs A)"
apply(induct rule: term_templateInduct[of "swapped \<union> {(X,Y). good Y \<and> skel Y = skel X}"])
by(auto simp: assms swapped_skel swapped_preserves_good)
-subsubsection {* Fresh induction *}
-
-text{* First a general situation, where parameters are of an unspecified type (that should be given by the user): *}
+subsubsection \<open>Fresh induction\<close>
+
+text\<open>First a general situation, where parameters are of an unspecified type (that should be given by the user):\<close>
lemma term_fresh_forall_induct[case_names PAR Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and A::"('index,'bindex,'varSort,'var,'opSym)abs"
and phi and phiAbs and varsOf :: "'param \<Rightarrow> 'varSort \<Rightarrow> 'var set"
assumes
PAR: "\<And> p xs. ( |varsOf xs p| <o |UNIV::'var set| )" and
var: "\<And> xs x p. phi (Var xs x) p" and
op: "\<And> delta inp binp p.
\<lbrakk>|{i. inp i \<noteq> None}| <o |UNIV::'var set|; |{i. binp i \<noteq> None}| <o |UNIV::'var set|;
liftAll (\<lambda> X. good X \<and> (\<forall> q. phi X p)) inp; liftAll (\<lambda> A. goodAbs A \<and> (\<forall> q. phiAbs A p)) binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp) p" and
abs: "\<And> xs x X p. \<lbrakk>good X; x \<notin> varsOf p xs; phi X p\<rbrakk> \<Longrightarrow> phiAbs (Abs xs x X) p"
shows "(good X \<longrightarrow> (\<forall> p. phi X p)) \<and> (goodAbs A \<longrightarrow> (\<forall> p. phiAbs A p))"
proof(induction rule: term_templateInduct[of swapped])
case (Abs xs x X)
show ?case proof safe
fix p
obtain x' where x'_freshP: "x' \<notin> varsOf p xs" and x'_fresh_X: "fresh xs x' X"
- using `good X` PAR obtain_fresh[of "varsOf p xs" "{X}" "{}" "{}" xs] by auto
+ using \<open>good X\<close> PAR obtain_fresh[of "varsOf p xs" "{X}" "{}" "{}" xs] by auto
then obtain X' where XX': "(X, X') \<in> swapped" and good_X': "good X'" and
Abs_eq: "Abs xs x X = Abs xs x' X'"
- using `good X` x'_freshP x'_fresh_X using obtain_rep[of X xs x' x] by auto
+ using \<open>good X\<close> x'_freshP x'_fresh_X using obtain_rep[of X xs x' x] by auto
thus "phiAbs (Abs xs x X) p"
unfolding Abs_eq using x'_freshP good_X' abs Abs.IH by simp
qed
qed(insert assms swapped_preserves_good swapped_skel,
unfold liftAll_def goodInp_def goodBinp_def, auto)
lemma term_templateInduct_fresh[case_names PAR Var Op Abs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
rel and phi and phiAbs and
vars :: "'varSort \<Rightarrow> 'var set" and
terms :: "('index,'bindex,'varSort,'var,'opSym)term set" and
abs :: "('index,'bindex,'varSort,'var,'opSym)abs set" and
envs :: "('index,'bindex,'varSort,'var,'opSym)env set"
assumes
PAR:
"\<And> xs.
( |vars xs| <o |UNIV :: 'var set| \<or> finite (vars xs)) \<and>
( |terms| <o |UNIV :: 'var set| \<or> finite terms) \<and> (\<forall> X \<in> terms. good X) \<and>
( |abs| <o |UNIV :: 'var set| \<or> finite abs) \<and> (\<forall> A \<in> abs. goodAbs A) \<and>
( |envs| <o |UNIV :: 'var set| \<or> finite envs) \<and> (\<forall> rho \<in> envs. goodEnv rho)" and
rel: "\<And> X Y. \<lbrakk>good X; (X,Y) \<in> rel\<rbrakk> \<Longrightarrow> good Y \<and> skel Y = skel X" and
Var: "\<And> xs x. phi (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>goodInp inp; goodBinp binp;
liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
abs:
"\<And> xs x X.
\<lbrakk>good X;
x \<notin> vars xs;
\<And> Y. Y \<in> terms \<Longrightarrow> fresh xs x Y;
\<And> A. A \<in> abs \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envs \<Longrightarrow> freshEnv xs x rho;
\<And> Y. (X,Y) \<in> rel \<Longrightarrow> phi Y\<rbrakk>
\<Longrightarrow> phiAbs (Abs xs x X)"
shows
"(good X \<longrightarrow> phi X) \<and>
(goodAbs A \<longrightarrow> phiAbs A)"
proof(induction rule: term_templateInduct[of "swapped O rel"])
- case (Abs xs x X) note good_X = `good X`
+ case (Abs xs x X) note good_X = \<open>good X\<close>
have "|{X} \<union> terms| <o |UNIV :: 'var set| \<or> finite ({X} \<union> terms)"
apply(cases "finite terms", auto simp add: PAR)
using PAR var_infinite_INNER card_of_Un_singl_ordLess_infinite by force
then obtain x' where x'_not: "x' \<notin> vars xs" and
x'_fresh_X: "fresh xs x' X" and
x'_freshP: "(\<forall> Y \<in> terms. fresh xs x' Y) \<and>
(\<forall> A \<in> abs. freshAbs xs x' A) \<and>
(\<forall> rho \<in> envs. freshEnv xs x' rho)"
using good_X PAR
using obtain_fresh[of "vars xs" "{X} \<union> terms" abs envs xs] by auto
then obtain X' where XX': "(X, X') \<in> swapped" and good_X': "good X'" and
Abs_eq: "Abs xs x X = Abs xs x' X'"
using good_X x'_not x'_fresh_X using obtain_rep[of X xs x' x] by auto
have "\<And>Y. (X', Y) \<in> rel \<Longrightarrow> phi Y" using XX' Abs.IH by auto
thus ?case
unfolding Abs_eq using x'_not x'_freshP good_X' abs by auto
qed(insert Op rel, unfold relcomp_unfold liftAll_def, simp_all add: Var,
metis rel swapped_preserves_good swapped_skel)
-text{* A version of the above not employing any relation for the bound-argument case: *}
+text\<open>A version of the above not employing any relation for the bound-argument case:\<close>
lemma term_rawInduct_fresh[case_names Par Var Op Obs]:
fixes X::"('index,'bindex,'varSort,'var,'opSym)term" and
A::"('index,'bindex,'varSort,'var,'opSym)abs" and
vars :: "'varSort \<Rightarrow> 'var set" and
terms :: "('index,'bindex,'varSort,'var,'opSym)term set" and
abs :: "('index,'bindex,'varSort,'var,'opSym)abs set" and
envs :: "('index,'bindex,'varSort,'var,'opSym)env set"
assumes
PAR:
"\<And> xs.
( |vars xs| <o |UNIV :: 'var set| \<or> finite (vars xs)) \<and>
( |terms| <o |UNIV :: 'var set| \<or> finite terms) \<and> (\<forall> X \<in> terms. good X) \<and>
( |abs| <o |UNIV :: 'var set| \<or> finite abs) \<and> (\<forall> A \<in> abs. goodAbs A) \<and>
( |envs| <o |UNIV :: 'var set| \<or> finite envs) \<and> (\<forall> rho \<in> envs. goodEnv rho)" and
Var: "\<And> xs x. phi (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>goodInp inp; goodBinp binp;
liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
Abs:
"\<And> xs x X.
\<lbrakk>good X;
x \<notin> vars xs;
\<And> Y. Y \<in> terms \<Longrightarrow> fresh xs x Y;
\<And> A. A \<in> abs \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envs \<Longrightarrow> freshEnv xs x rho;
phi X\<rbrakk>
\<Longrightarrow> phiAbs (Abs xs x X)"
shows
"(good X \<longrightarrow> phi X) \<and>
(goodAbs A \<longrightarrow> phiAbs A)"
apply(induct rule: term_templateInduct_fresh[of vars terms abs envs Id])
using assms by auto
(* Note that here, since we avoid variable-capture and hence will
not typically need to swap, term-inductRaw_fresh will suffice in proofs.
Therefore we do not prove a swapped-and-skel version of fresh induction, although such a version
could be easily inferred from ``term-templateInduct". *)
-text{* The typical raw induction with freshness is one dealing with
+text\<open>The typical raw induction with freshness is one dealing with
finitely many variables, terms, abstractions and environments as parameters --
we have all these condensed in the notion of a parameter (type
- constructor ``param"): *}
+ constructor ``param"):\<close>
lemma term_induct_fresh[case_names Par Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)term" and
A :: "('index,'bindex,'varSort,'var,'opSym)abs" and
P :: "('index,'bindex,'varSort,'var,'opSym)param"
assumes
goodP: "goodPar P" and
Var: "\<And> xs x. phi (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>goodInp inp; goodBinp binp;
liftAll phi inp; liftAll phiAbs binp\<rbrakk>
\<Longrightarrow> phi (Op delta inp binp)" and
Abs:
"\<And> xs x X.
\<lbrakk>good X;
x \<notin> varsOf P;
\<And> Y. Y \<in> termsOf P \<Longrightarrow> fresh xs x Y;
\<And> A. A \<in> absOf P \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envsOf P \<Longrightarrow> freshEnv xs x rho;
phi X\<rbrakk>
\<Longrightarrow> phiAbs (Abs xs x X)"
shows
"(good X \<longrightarrow> phi X) \<and>
(goodAbs A \<longrightarrow> phiAbs A)"
proof(induct rule: term_rawInduct_fresh
[of "\<lambda> xs. varsOf P" "termsOf P" "absOf P" "envsOf P"])
case (Par xs)
then show ?case unfolding goodPar_def
using goodP by(cases P) simp
qed(insert assms, auto)
end (* context FixVars *)
end
diff --git a/thys/Binding_Syntax_Theory/Variable.thy b/thys/Binding_Syntax_Theory/Variable.thy
--- a/thys/Binding_Syntax_Theory/Variable.thy
+++ b/thys/Binding_Syntax_Theory/Variable.thy
@@ -1,10 +1,10 @@
theory Variable
imports Main
begin
-subsection{* A common concrete type of variables *}
+subsection\<open>A common concrete type of variables\<close>
datatype var = Variable nat
end
\ No newline at end of file
diff --git a/thys/Binding_Syntax_Theory/Well_Sorted_Terms.thy b/thys/Binding_Syntax_Theory/Well_Sorted_Terms.thy
--- a/thys/Binding_Syntax_Theory/Well_Sorted_Terms.thy
+++ b/thys/Binding_Syntax_Theory/Well_Sorted_Terms.thy
@@ -1,2908 +1,2908 @@
-section {* Binding Signatures and well-sorted terms *}
+section \<open>Binding Signatures and well-sorted terms\<close>
theory Well_Sorted_Terms
imports Terms
begin
-text {* This section introduces binding signatures
+text \<open>This section introduces binding signatures
and well-sorted terms for them. All the properties we proved for good terms are then
lifted to well-sorted terms.
-*}
-
-subsection {* Binding signatures *}
-
-text{* A {\em (binding) signature} consists of:
+\<close>
+
+subsection \<open>Binding signatures\<close>
+
+text\<open>A {\em (binding) signature} consists of:
\\- an indication of which sorts of variables can be injected in
which sorts of terms;
\\- for any operation symbol, dwelling in a type ``opSym",
an indication of its result sort, its (nonbinding) arity, and its binding arity.
In addition, we have a predicate, ``wlsOpSym", that specifies which operations symbols
are well-sorted (or well-structured)
\footnote
{
We shall use ``wls" in many contexts as a prefix indicating well-sortedness or
well-structuredness.
}
-- only these operation symbols will be considered in
forming terms. In other words, the relevant collection of operation symbols is given not by the
whole type ``opSym", but by the predicate ``wlsOpSym". This bit of extra flexibility
will be useful when (pre)instantiating the signature to concrete syntaxes.
(Note that the ``wlsOpSym" condition will be required for well-sorted terms as part of the notion of
well-sorted (free and bound) input, ``wlsInp" and ``wlsBinp".)
- *}
+\<close>
record ('index,'bindex,'varSort,'sort,'opSym)signature =
varSortAsSort :: "'varSort \<Rightarrow> 'sort"
wlsOpSym :: "'opSym \<Rightarrow> bool"
sortOf :: "'opSym \<Rightarrow> 'sort"
arityOf :: "'opSym \<Rightarrow> ('index, 'sort)input"
barityOf :: "'opSym \<Rightarrow> ('bindex, 'varSort * 'sort)input"
-subsection {* The Binding Syntax locale *}
+subsection \<open>The Binding Syntax locale\<close>
(* From now on, all work on binding syntax shall be developed in this locale
(or an extension of it): *)
(*
Note also that currently locales do not allow datatype defs or records,
reason for which the type of models has been defined outside. *)
-text {* For our signatures, we shall make some assumptions:
+text \<open>For our signatures, we shall make some assumptions:
\\- For each sort of term, there is at most one sort of variables injectable
in terms of that sort (i.e., ``varSortAsSort" is injective");
\\- The domains of arities (sets of indexes) are smaller than the set of variables
of each sort;
\\- The type of sorts is smaller than the set of variables
of each sort.
These are satisfiable assumptions, and in particular they are trivially satisfied by any finitary syntax
with bindings.
-*}
+\<close>
definition varSortAsSort_inj where
"varSortAsSort_inj Delta ==
inj (varSortAsSort Delta)"
definition arityOf_lt_var where
"arityOf_lt_var (_ :: 'var) Delta ==
\<forall> delta.
wlsOpSym Delta delta \<longrightarrow> |{i. arityOf Delta delta i \<noteq> None}| <o |UNIV :: 'var set|"
definition barityOf_lt_var where
"barityOf_lt_var (_ :: 'var) Delta ==
\<forall> delta.
wlsOpSym Delta delta \<longrightarrow> |{i. barityOf Delta delta i \<noteq> None}| <o |UNIV :: 'var set|"
definition sort_lt_var where
"sort_lt_var (_ :: 'sort) (_ :: 'var) ==
|UNIV :: 'sort set| <o |UNIV :: 'var set|"
locale FixSyn =
fixes dummyV :: 'var
and Delta :: "('index,'bindex,'varSort,'sort,'opSym)signature"
assumes
(* The ``FixVars" assumptions, minus ``varSort-lt-var" (which follows form the rest): *)
FixSyn_var_infinite: "var_infinite (undefined :: 'var)"
and FixSyn_var_regular: "var_regular (undefined :: 'var)"
(* Plus the following: *)
and varSortAsSort_inj: "varSortAsSort_inj Delta"
and arityOf_lt_var: "arityOf_lt_var (undefined :: 'var) Delta"
and barityOf_lt_var: "barityOf_lt_var (undefined :: 'var) Delta"
and sort_lt_var: "sort_lt_var (undefined :: 'sort) (undefined :: 'var)"
context FixSyn
begin
lemmas FixSyn_assms =
FixSyn_var_infinite FixSyn_var_regular
varSortAsSort_inj arityOf_lt_var barityOf_lt_var
sort_lt_var
end
-subsection {* Definitions and basic properties of well-sortedness *}
-
-subsubsection {* Notations and definitions *}
+subsection \<open>Definitions and basic properties of well-sortedness\<close>
+
+subsubsection \<open>Notations and definitions\<close>
(* Sorted parameters (again, for use in proofs): *)
datatype ('index,'bindex,'varSort,'var,'opSym,'sort)paramS =
ParS "'varSort \<Rightarrow> 'var list"
"'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term list"
"('varSort * 'sort) \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs list"
"('index,'bindex,'varSort,'var,'opSym)env list"
fun varsOfS ::
"('index,'bindex,'varSort,'var,'opSym,'sort)paramS \<Rightarrow> 'varSort \<Rightarrow> 'var set"
where "varsOfS (ParS xLF _ _ _) xs = set (xLF xs)"
fun termsOfS ::
"('index,'bindex,'varSort,'var,'opSym,'sort)paramS \<Rightarrow>
'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term set"
where "termsOfS (ParS _ XLF _ _) s = set (XLF s)"
fun absOfS ::
"('index,'bindex,'varSort,'var,'opSym,'sort)paramS \<Rightarrow>
('varSort * 'sort) \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs set"
where "absOfS (ParS _ _ ALF _) (xs,s) = set (ALF (xs,s))"
fun envsOfS ::
"('index,'bindex,'varSort,'var,'opSym,'sort)paramS \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)env set"
where "envsOfS (ParS _ _ _ rhoL) = set rhoL"
-subsubsection {* Sublocale of ``FixVars" *}
+subsubsection \<open>Sublocale of ``FixVars"\<close>
lemma sort_lt_var_imp_varSort_lt_var:
assumes
**: "varSortAsSort_inj (Delta :: ('index,'bindex,'varSort,'sort,'opSym)signature)"
and ***: "sort_lt_var (undefined :: 'sort) (undefined :: 'var)"
shows "varSort_lt_var (undefined :: 'varSort) (undefined :: 'var)"
proof-
have "|UNIV::'varSort set| \<le>o |UNIV::'sort set|"
using card_of_ordLeq ** unfolding varSortAsSort_inj_def by auto
thus ?thesis
using ordLeq_ordLess_trans assms
unfolding sort_lt_var_def varSort_lt_var_def by blast
qed
sublocale FixSyn < FixVars
where dummyV = dummyV and dummyVS = "undefined::'varSort"
using FixSyn_assms
by unfold_locales (auto simp add: sort_lt_var_imp_varSort_lt_var)
-subsubsection {* Abbreviations *}
+subsubsection \<open>Abbreviations\<close>
(*********************************************)
context FixSyn (* scope all throughout the file *)
begin
abbreviation asSort where "asSort == varSortAsSort Delta"
abbreviation wlsOpS where "wlsOpS == wlsOpSym Delta"
abbreviation stOf where "stOf == sortOf Delta"
abbreviation arOf where "arOf == arityOf Delta"
abbreviation barOf where "barOf == barityOf Delta"
abbreviation empInp ::
"('index,('index,'bindex,'varSort,'var,'opSym)term)input"
where "empInp == \<lambda>i. None"
abbreviation empAr :: "('index,'sort)input"
where "empAr == \<lambda>i. None"
abbreviation empBinp :: "('bindex,('index,'bindex,'varSort,'var,'opSym)abs)input"
where "empBinp == \<lambda>i. None"
abbreviation empBar :: "('bindex,'varSort * 'sort)input"
where "empBar == \<lambda>i. None"
lemma freshInp_empInp[simp]:
"freshInp xs x empInp"
unfolding freshInp_def liftAll_def by simp
lemma swapInp_empInp[simp]:
"(empInp %[x1 \<and> x2]_xs) = empInp"
unfolding swapInp_def lift_def by simp
lemma psubstInp_empInp[simp]:
"(empInp %[rho]) = empInp"
unfolding psubstInp_def lift_def by simp
lemma substInp_empInp[simp]:
"(empInp %[Y / y]_ys) = empInp"
unfolding substInp_def by simp
lemma vsubstInp_empInp[simp]:
"(empInp %[y1 // y]_ys) = empInp"
unfolding vsubstInp_def by simp
lemma freshBinp_empBinp[simp]:
"freshBinp xs x empBinp"
unfolding freshBinp_def liftAll_def by simp
lemma swapBinp_empBinp[simp]:
"(empBinp %%[x1 \<and> x2]_xs) = empBinp"
unfolding swapBinp_def lift_def by simp
lemma psubstBinp_empBinp[simp]:
"(empBinp %%[rho]) = empBinp"
unfolding psubstBinp_def lift_def by simp
lemma substBinp_empBinp[simp]:
"(empBinp %%[Y / y]_ys) = empBinp"
unfolding substBinp_def by simp
lemma vsubstBinp_empBinp[simp]:
"(empBinp %%[y1 // y]_ys) = empBinp"
unfolding vsubstBinp_def by simp
lemmas empInp_simps =
freshInp_empInp swapInp_empInp psubstInp_empInp substInp_empInp vsubstInp_empInp
freshBinp_empBinp swapBinp_empBinp psubstBinp_empBinp substBinp_empBinp vsubstBinp_empBinp
-subsubsection {* Inner versions of the locale assumptions *}
+subsubsection \<open>Inner versions of the locale assumptions\<close>
lemma varSortAsSort_inj_INNER: "inj asSort"
using varSortAsSort_inj
unfolding varSortAsSort_inj_def by simp
lemma asSort_inj[simp]:
"(asSort xs = asSort ys) = (xs = ys)"
using varSortAsSort_inj_INNER unfolding inj_on_def by auto
lemma arityOf_lt_var_INNER:
assumes "wlsOpS delta"
shows "|{i. arityOf Delta delta i \<noteq> None}| <o |UNIV :: 'var set|"
using assms arityOf_lt_var unfolding arityOf_lt_var_def by simp
lemma barityOf_lt_var_INNER:
assumes "wlsOpS delta"
shows "|{i. barityOf Delta delta i \<noteq> None}| <o |UNIV :: 'var set|"
using assms barityOf_lt_var unfolding barityOf_lt_var_def by simp
lemma sort_lt_var_INNER:
"|UNIV :: 'sort set| <o |UNIV :: 'var set|"
using sort_lt_var unfolding sort_lt_var_def by simp
lemma sort_le_var:
"|UNIV :: 'sort set| \<le>o |UNIV :: 'var set|"
using sort_lt_var_INNER ordLess_imp_ordLeq by auto
lemma varSort_sort_lt_var:
"|UNIV :: ('varSort * 'sort) set| <o |UNIV :: 'var set|"
unfolding UNIV_Times_UNIV[symmetric]
using var_infinite_INNER varSort_lt_var_INNER sort_lt_var_INNER
by(rule card_of_Times_ordLess_infinite)
lemma varSort_sort_le_var:
"|UNIV :: ('varSort * 'sort) set| \<le>o |UNIV :: 'var set|"
using varSort_sort_lt_var ordLess_imp_ordLeq by auto
-subsubsection {* Definitions of well-sorted items *}
-
-text {* We shall only be interested in abstractions that pertain to some bound arities: *}
+subsubsection \<open>Definitions of well-sorted items\<close>
+
+text \<open>We shall only be interested in abstractions that pertain to some bound arities:\<close>
definition isInBar where
"isInBar xs_s ==
\<exists> delta i. wlsOpS delta \<and> barOf delta i = Some xs_s"
-text {* Well-sorted terms (according to the signature) are defined as expected (mutually inductively
-together with well-sorted abstractions and inputs): *}
+text \<open>Well-sorted terms (according to the signature) are defined as expected (mutually inductively
+together with well-sorted abstractions and inputs):\<close>
inductive
wls :: "'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term \<Rightarrow> bool"
and
wlsAbs :: "'varSort * 'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs \<Rightarrow> bool"
and
wlsInp :: "'opSym \<Rightarrow> ('index,('index,'bindex,'varSort,'var,'opSym)term)input \<Rightarrow> bool"
and
wlsBinp :: "'opSym \<Rightarrow> ('bindex,('index,'bindex,'varSort,'var,'opSym)abs)input \<Rightarrow> bool"
where
Var: "wls (asSort xs) (Var xs x)"
|
Op: "\<lbrakk>wlsInp delta inp; wlsBinp delta binp\<rbrakk> \<Longrightarrow> wls (stOf delta) (Op delta inp binp)"
|
Inp:
"\<lbrakk>wlsOpS delta;
\<And> i. (arOf delta i = None \<and> inp i = None) \<or>
(\<exists> s X. arOf delta i = Some s \<and> inp i = Some X \<and> wls s X)\<rbrakk>
\<Longrightarrow> wlsInp delta inp"
|
Binp:
"\<lbrakk>wlsOpS delta;
\<And> i. (barOf delta i = None \<and> binp i = None) \<or>
(\<exists> us s A. barOf delta i = Some (us,s) \<and> binp i = Some A \<and> wlsAbs (us,s) A)\<rbrakk>
\<Longrightarrow> wlsBinp delta binp"
|
Abs: "\<lbrakk>isInBar (xs,s); wls s X\<rbrakk> \<Longrightarrow> wlsAbs (xs,s) (Abs xs x X)"
lemmas Var_preserves_wls = wls_wlsAbs_wlsInp_wlsBinp.Var
lemmas Op_preserves_wls = wls_wlsAbs_wlsInp_wlsBinp.Op
lemmas Abs_preserves_wls = wls_wlsAbs_wlsInp_wlsBinp.Abs
lemma barOf_isInBar[simp]:
assumes "wlsOpS delta" and "barOf delta i = Some (us,s)"
shows "isInBar (us,s)"
unfolding isInBar_def using assms by blast
lemmas Cons_preserve_wls =
barOf_isInBar
Var_preserves_wls Op_preserves_wls
Abs_preserves_wls
declare Cons_preserve_wls [simp]
definition wlsEnv :: "('index,'bindex,'varSort,'var,'opSym)env \<Rightarrow> bool"
where
"wlsEnv rho ==
(\<forall> ys. liftAll (wls (asSort ys)) (rho ys)) \<and>
(\<forall> ys. |{y. rho ys y \<noteq> None}| <o |UNIV :: 'var set| )"
definition wlsPar :: "('index,'bindex,'varSort,'var,'opSym,'sort)paramS \<Rightarrow> bool"
where
"wlsPar P ==
(\<forall> s. \<forall> X \<in> termsOfS P s. wls s X) \<and>
(\<forall> xs s. \<forall> A \<in> absOfS P (xs,s). wlsAbs (xs,s) A) \<and>
(\<forall> rho \<in> envsOfS P. wlsEnv rho)"
lemma ParS_preserves_wls[simp]:
assumes "\<And> s X. X \<in> set (XLF s) \<Longrightarrow> wls s X"
and "\<And> xs s A. A \<in> set (ALF (xs,s)) \<Longrightarrow> wlsAbs (xs,s) A"
and "\<And> rho. rho \<in> set rhoF \<Longrightarrow> wlsEnv rho"
shows "wlsPar (ParS xLF XLF ALF rhoF)"
using assms unfolding wlsPar_def by auto
lemma termsOfS_preserves_wls[simp]:
assumes "wlsPar P" and "X : termsOfS P s"
shows "wls s X"
using assms unfolding wlsPar_def by auto
lemma absOfS_preserves_wls[simp]:
assumes "wlsPar P" and "A : absOfS P (us,s)"
shows "wlsAbs (us,s) A"
using assms unfolding wlsPar_def by auto
lemma envsOfS_preserves_wls[simp]:
assumes "wlsPar P" and "rho : envsOfS P "
shows "wlsEnv rho"
using assms unfolding wlsPar_def by blast
lemma not_isInBar_absOfS_empty[simp]:
assumes *: "\<not> isInBar (us,s)" and **: "wlsPar P"
shows "absOfS P (us,s) = {}"
proof-
{fix A assume "A : absOfS P (us,s)"
hence "wlsAbs (us,s) A" using ** by simp
hence False using * using wlsAbs.cases by auto
}
thus ?thesis by auto
qed
lemmas paramS_simps =
varsOfS.simps termsOfS.simps absOfS.simps envsOfS.simps
ParS_preserves_wls
termsOfS_preserves_wls absOfS_preserves_wls envsOfS_preserves_wls
not_isInBar_absOfS_empty
-subsubsection {* Well-sorted exists *}
+subsubsection \<open>Well-sorted exists\<close>
lemma wlsInp_iff:
"wlsInp delta inp =
(wlsOpS delta \<and> sameDom (arOf delta) inp \<and> liftAll2 wls (arOf delta) inp)"
by (simp add: wlsInp.simps wls_wlsAbs_wlsInp_wlsBinp.Inp sameDom_and_liftAll2_iff)
lemma wlsBinp_iff:
"wlsBinp delta binp =
(wlsOpS delta \<and> sameDom (barOf delta) binp \<and> liftAll2 wlsAbs (barOf delta) binp)"
by (simp add: wlsBinp.simps wls_wlsAbs_wlsInp_wlsBinp.Inp sameDom_and_liftAll2_iff)
lemma exists_asSort_wls:
"\<exists> X. wls (asSort xs) X"
by (intro exI[of _ "Var xs undefined"]) simp
lemma exists_wls_imp_exists_wlsAbs:
assumes *: "isInBar (us,s)" and **: "\<exists> X. wls s X"
shows "\<exists> A. wlsAbs (us,s) A"
proof-
obtain X where "wls s X" using ** by blast
hence "wlsAbs (us,s) (Abs us undefined X)" using * by simp
thus ?thesis by blast
qed
lemma exists_asSort_wlsAbs:
assumes "isInBar (us,asSort xs)"
shows "\<exists> A. wlsAbs (us,asSort xs) A"
proof-
obtain X where "wls (asSort xs) X" using exists_asSort_wls by auto
thus ?thesis using assms exists_wls_imp_exists_wlsAbs by auto
qed
-text {* Standard criterion for the non-emptiness of the sets of well-sorted terms for each sort,
+text \<open>Standard criterion for the non-emptiness of the sets of well-sorted terms for each sort,
by a well-founded relation and a function picking, for sorts not corresponding to varSorts,
an operation symbol as an ``inductive" witness for non-emptyness.
-``witOpS" stands for ``witness operation symbol". *}
+``witOpS" stands for ``witness operation symbol".\<close>
definition witOpS where
"witOpS s delta R ==
wlsOpS delta \<and> stOf delta = s \<and>
liftAll (\<lambda>s'. (s',s) : R) (arOf delta) \<and>
liftAll (\<lambda>(us,s'). (s',s) : R) (barOf delta)"
lemma wf_exists_wls:
assumes wf: "wf R" and *: "\<And>s. (\<exists> xs. s = asSort xs) \<or> witOpS s (f s) R"
shows "\<exists> X. wls s X"
proof(induction rule: wf_induct[of R])
case (2 s)
show ?case
proof(cases "\<exists> xs. s = asSort xs")
case True
thus ?thesis using exists_asSort_wls by auto
next
let ?delta = "f s"
case False
hence delta: "wlsOpS ?delta" and st: "stOf ?delta = s"
and ar: "liftAll (\<lambda>s'. (s',s) : R) (arOf ?delta)"
and bar: "liftAll (\<lambda>(us,s'). (s',s) : R) (barOf ?delta)"
using * unfolding witOpS_def by auto
(* *)
have 1: "\<forall> i s'. arOf ?delta i = Some s' \<longrightarrow> (\<exists> X. wls s' X)"
using ar 2 unfolding liftAll_def by simp
let ?chi = "\<lambda>i s' X. arOf ?delta i = Some s' \<longrightarrow> wls s' X"
define inp where
"inp \<equiv> (\<lambda>i. (if arOf ?delta i = None
then None
else Some (SOME X. \<forall> s'. ?chi i s' X)))"
have inp: "wlsInp ?delta inp"
unfolding wlsInp_iff sameDom_def liftAll2_def using delta
by (auto simp: inp_def 1 someI2_ex split: if_splits)
(* *)
have 1: "\<forall> i us s'. barOf ?delta i = Some (us,s') \<longrightarrow> (\<exists> A. wlsAbs (us,s') A)"
using bar 2 unfolding liftAll_def using delta exists_wls_imp_exists_wlsAbs by simp
let ?chi = "\<lambda>i us s' A. barOf ?delta i = Some (us,s') \<longrightarrow> wlsAbs (us,s') A"
define binp where
"binp \<equiv> (\<lambda>i. (if barOf ?delta i = None
then None
else Some (SOME A. \<forall> us s'. ?chi i us s' A)))"
have binp: "wlsBinp ?delta binp"
unfolding wlsBinp_iff sameDom_def liftAll2_def using delta
by (auto simp: binp_def 1 someI2_ex split: if_splits)
(* *)
have "wls s (Op ?delta inp binp)"
using inp binp st using Op_preserves_wls[of ?delta inp binp] by simp
thus ?thesis by blast
qed
qed(insert assms, auto)
lemma wf_exists_wlsAbs:
assumes "isInBar (us,s)"
and "wf R" and "\<And>s. (\<exists> xs. s = asSort xs) \<or> witOpS s (f s) R"
shows "\<exists> A. wlsAbs (us,s) A"
using assms by (auto intro: exists_wls_imp_exists_wlsAbs wf_exists_wls)
-subsubsection {* Well-sorted implies Good *}
+subsubsection \<open>Well-sorted implies Good\<close>
lemma wlsInp_empAr_empInp[simp]:
assumes "wlsOpS delta" and "arOf delta = empAr"
shows "wlsInp delta empInp"
using assms
unfolding wlsInp_iff sameDom_def liftAll2_def by auto
lemma wlsBinp_empBar_empBinp[simp]:
assumes "wlsOpS delta" and "barOf delta = empBar"
shows "wlsBinp delta empBinp"
using assms unfolding wlsBinp_iff sameDom_def liftAll2_def by auto
lemmas empInp_otherSimps =
wlsInp_empAr_empInp wlsBinp_empBar_empBinp
lemma wlsAll_implies_goodAll:
"(wls s X \<longrightarrow> good X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> goodAbs A) \<and>
(wlsInp delta inp \<longrightarrow> goodInp inp) \<and>
(wlsBinp delta binp \<longrightarrow> goodBinp binp)"
apply(induct rule: wls_wlsAbs_wlsInp_wlsBinp.induct)
subgoal by auto
subgoal by auto
subgoal unfolding goodInp_def liftAll_def
by simp (smt Collect_cong arityOf_lt_var_INNER option.distinct(1) option.sel)
subgoal unfolding goodBinp_def liftAll_def
by simp (smt Collect_cong barityOf_lt_var_INNER option.distinct(1) option.sel)
subgoal by auto .
corollary wls_imp_good[simp]: "wls s X \<Longrightarrow> good X"
by(simp add: wlsAll_implies_goodAll)
corollary wlsAbs_imp_goodAbs[simp]: "wlsAbs (xs,s) A \<Longrightarrow> goodAbs A"
by(simp add: wlsAll_implies_goodAll)
corollary wlsInp_imp_goodInp[simp]: "wlsInp delta inp \<Longrightarrow> goodInp inp"
by(simp add: wlsAll_implies_goodAll)
corollary wlsBinp_imp_goodBinp[simp]: "wlsBinp delta binp \<Longrightarrow> goodBinp binp"
by(simp add: wlsAll_implies_goodAll)
lemma wlsEnv_imp_goodEnv[simp]: "wlsEnv rho \<Longrightarrow> goodEnv rho"
unfolding wlsEnv_def goodEnv_def liftAll_def
by simp (insert wls_imp_good, blast)
lemmas wlsAll_imp_goodAll =
wls_imp_good wlsAbs_imp_goodAbs
wlsInp_imp_goodInp wlsBinp_imp_goodBinp
wlsEnv_imp_goodEnv
-subsubsection {* Swapping preserves well-sortedness *}
+subsubsection \<open>Swapping preserves well-sortedness\<close>
lemma swapAll_pres_wlsAll:
"(wls s X \<longrightarrow> wls s (X #[z1 \<and> z2]_zs)) \<and>
(wlsAbs (xs,s') A \<longrightarrow> wlsAbs (xs,s') (A $[z1 \<and> z2]_zs)) \<and>
(wlsInp delta inp \<longrightarrow> wlsInp delta (inp %[z1 \<and> z2]_zs)) \<and>
(wlsBinp delta binp \<longrightarrow> wlsBinp delta (binp %%[z1 \<and> z2]_zs))"
proof(induct rule: wls_wlsAbs_wlsInp_wlsBinp.induct)
case (Inp delta inp)
then show ?case
unfolding wlsInp_iff sameDom_def liftAll2_def lift_def swapInp_def
using option.sel by (fastforce simp add: split: option.splits)
next
case (Binp delta binp)
then show ?case
unfolding wlsBinp_iff sameDom_def liftAll2_def lift_def swapBinp_def
using option.sel by (fastforce simp add: split: option.splits)
qed(insert Cons_preserve_wls, simp_all)
lemma swap_preserves_wls[simp]:
"wls s X \<Longrightarrow> wls s (X #[z1 \<and> z2]_zs)"
by(simp add: swapAll_pres_wlsAll)
lemma swap_preserves_wls2[simp]:
assumes "good X"
shows "wls s (X #[z1 \<and> z2]_zs) = wls s X"
using assms swap_preserves_wls[of s "X #[z1 \<and> z2]_zs" zs z1 z2] by auto
lemma swap_preserves_wls3:
assumes "good X" and "good Y"
and "(X #[x1 \<and> x2]_xs) = (Y #[y1 \<and> y2]_ys)"
shows "wls s X = wls s Y"
by (metis assms swap_preserves_wls2)
lemma swapAbs_preserves_wls[simp]:
"wlsAbs (xs,x) A \<Longrightarrow> wlsAbs (xs,x) (A $[z1 \<and> z2]_zs)"
by(simp add: swapAll_pres_wlsAll)
lemma swapInp_preserves_wls[simp]:
"wlsInp delta inp \<Longrightarrow> wlsInp delta (inp %[z1 \<and> z2]_zs)"
by(simp add: swapAll_pres_wlsAll)
lemma swapBinp_preserves_wls[simp]:
"wlsBinp delta binp \<Longrightarrow> wlsBinp delta (binp %%[z1 \<and> z2]_zs)"
by(simp add: swapAll_pres_wlsAll)
lemma swapEnvDom_preserves_wls:
assumes "wlsEnv rho"
shows "wlsEnv (swapEnvDom xs x y rho)"
proof-
{fix xsa ys let ?Left = "{ya. swapEnvDom xs x y rho ys ya \<noteq> None}"
have "|{y} \<union> {ya. rho ys ya \<noteq> None}| <o |UNIV :: 'var set|"
using assms var_infinite_INNER card_of_Un_singl_ordLess_infinite
unfolding wlsEnv_def by fastforce
hence "|{x,y} \<union> {ya. rho ys ya \<noteq> None}| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_singl_ordLess_infinite by fastforce
moreover
{have "?Left \<subseteq> {x,y} \<union> {ya. rho ys ya \<noteq> None}"
unfolding swapEnvDom_def sw_def[abs_def] by auto
hence "|?Left| \<le>o |{x,y} \<union> {ya. rho ys ya \<noteq> None}|"
using card_of_mono1 by auto
}
ultimately have "|?Left| <o |UNIV :: 'var set|"
using ordLeq_ordLess_trans by blast
}
thus ?thesis using assms unfolding wlsEnv_def liftAll_def
by (auto simp add: swapEnvDom_def)
qed
lemma swapEnvIm_preserves_wls:
assumes "wlsEnv rho"
shows "wlsEnv (swapEnvIm xs x y rho)"
using assms unfolding wlsEnv_def swapEnvIm_def liftAll_def lift_def
by (auto split: option.splits)
lemma swapEnv_preserves_wls[simp]:
assumes "wlsEnv rho"
shows "wlsEnv (rho &[z1 \<and> z2]_zs)"
unfolding swapEnv_def comp_def
using assms by(auto simp: swapEnvDom_preserves_wls swapEnvIm_preserves_wls)
lemmas swapAll_preserve_wls =
swap_preserves_wls swapAbs_preserves_wls
swapInp_preserves_wls swapBinp_preserves_wls
swapEnv_preserves_wls
lemma swapped_preserves_wls:
assumes "wls s X" and "(X,Y) \<in> swapped"
shows "wls s Y"
proof-
have "(X,Y) \<in> swapped \<Longrightarrow> wls s X \<longrightarrow> wls s Y"
by (induct rule: swapped.induct) auto
thus ?thesis using assms by simp
qed
-subsubsection {* Inversion rules for well-sortedness *}
+subsubsection \<open>Inversion rules for well-sortedness\<close>
lemma wlsAll_inversion:
"(wls s X \<longrightarrow>
(\<forall> xs x. X = Var xs x \<longrightarrow> s = asSort xs) \<and>
(\<forall> delta inp binp. goodInp inp \<and> goodBinp binp \<and> X = Op delta inp binp \<longrightarrow>
stOf delta = s \<and> wlsInp delta inp \<and> wlsBinp delta binp))
\<and>
(wlsAbs xs_s A \<longrightarrow>
isInBar xs_s \<and>
(\<forall> x X. good X \<and> A = Abs (fst xs_s) x X \<longrightarrow>
wls (snd xs_s) X))
\<and>
(wlsInp delta inp \<longrightarrow> True)
\<and>
(wlsBinp delta binp \<longrightarrow> True)"
proof(induct rule: wls_wlsAbs_wlsInp_wlsBinp.induct)
case (Abs xs s X x)
then show ?case using swap_preserves_wls3 wls_imp_good
by (metis FixVars.Abs_ainj_ex FixVars_axioms snd_conv)
qed (auto simp: Abs_ainj_ex)
lemma conjLeft: "\<lbrakk>phi1 \<and> phi2; phi1 \<Longrightarrow> chi\<rbrakk> \<Longrightarrow> chi"
by blast
lemma conjRight: "\<lbrakk>phi1 \<and> phi2; phi2 \<Longrightarrow> chi\<rbrakk> \<Longrightarrow> chi"
by blast
(* For the next three lemmas, I could not use the simp, auto etc.
powerhorses freely -- for some reason, they loop\<dots> *)
lemma wls_inversion[rule_format]:
"wls s X \<longrightarrow>
(\<forall> xs x. X = Var xs x \<longrightarrow> s = asSort xs) \<and>
(\<forall> delta inp binp. goodInp inp \<and> goodBinp binp \<and> X = Op delta inp binp \<longrightarrow>
stOf delta = s \<and> wlsInp delta inp \<and> wlsBinp delta binp)"
using wlsAll_inversion
[of s X undefined undefined undefined undefined undefined]
by (rule conjLeft)
lemma wlsAbs_inversion[rule_format]:
"wlsAbs (xs,s) A \<longrightarrow>
isInBar (xs,s) \<and>
(\<forall> x X. good X \<and> A = Abs xs x X \<longrightarrow> wls s X)"
using wlsAll_inversion
[of undefined undefined "(xs,s)" A undefined undefined undefined]
by auto
lemma wls_Var_simp[simp]:
"wls s (Var xs x) = (s = asSort xs)"
using wls_inversion by auto
lemma wls_Op_simp[simp]:
assumes "goodInp inp" and "goodBinp binp"
shows
"wls s (Op delta inp binp) =
(stOf delta = s \<and> wlsInp delta inp \<and> wlsBinp delta binp)"
using Op assms wls_inversion by blast
lemma wls_Abs_simp[simp]:
assumes "good X"
shows "wlsAbs (xs,s) (Abs xs x X) = (isInBar (xs,s) \<and> wls s X)"
using Abs assms wlsAbs_inversion by blast
lemma wlsAll_inversion2:
"(wls s X \<longrightarrow> True)
\<and>
(wlsAbs xs_s A \<longrightarrow>
isInBar xs_s \<and>
(\<exists> x X. wls (snd xs_s) X \<and> A = Abs (fst xs_s) x X))
\<and>
(wlsInp delta inp \<longrightarrow> True)
\<and>
(wlsBinp delta binp \<longrightarrow> True)"
by (induct rule: wls_wlsAbs_wlsInp_wlsBinp.induct)
(auto simp add: Abs_ainj_ex simp del: not_None_eq)
lemma wlsAbs_inversion2[rule_format]:
"wlsAbs (xs,s) A \<longrightarrow>
isInBar (xs,s) \<and> (\<exists> x X. wls s X \<and> A = Abs xs x X)"
using wlsAll_inversion2 by auto
corollary wlsAbs_Abs_varSort:
assumes X: "good X" and wlsAbs: "wlsAbs (xs,s) (Abs xs' x X)"
shows "xs = xs'"
by (metis Abs_ainj_all X wlsAbs wlsAbs_inversion2 wls_imp_good)
lemma wlsAbs:
"wlsAbs (xs,s) A \<longleftrightarrow>
isInBar (xs,s) \<and> (\<exists> x X. wls s X \<and> A = Abs xs x X)"
using Abs wlsAbs_inversion2 by blast
lemma wlsAbs_Abs[simp]:
assumes X: "good X"
shows "wlsAbs (xs',s) (Abs xs x X) = (isInBar (xs',s) \<and> xs = xs' \<and> wls s X)"
using assms wlsAbs_Abs_varSort by fastforce
lemmas Cons_wls_simps =
wls_Var_simp wls_Op_simp wls_Abs_simp wlsAbs_Abs
-subsection {* Induction principles for well-sorted terms *}
-
-subsubsection {* Regular induction *}
+subsection \<open>Induction principles for well-sorted terms\<close>
+
+subsubsection \<open>Regular induction\<close>
(* See also the subsection on substitution of variables for variables
from the section on compositional properties
for an induction principle involving a choice between swap, subst and skeleton
for the abstraction case. *)
theorem wls_templateInduct[case_names rel Var Op Abs]:
assumes
rel: "\<And> s X Y. \<lbrakk>wls s X; (X,Y) \<in> rel s\<rbrakk> \<Longrightarrow> wls s Y \<and> skel Y = skel X" and
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 phi (arOf delta) inp; liftAll2 phiAbs (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs:
"\<And> s xs x X.
\<lbrakk>isInBar (xs,s); wls s X; \<And> Y. (X,Y) \<in> rel s \<Longrightarrow> phi s Y\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
proof-
have "(good X \<longrightarrow> (\<forall> s. wls s X \<longrightarrow> phi s X)) \<and>
(goodAbs A \<longrightarrow> (\<forall> xs s. wlsAbs (xs,s) A \<longrightarrow> phiAbs (xs,s) A))"
apply(induct rule: term_templateInduct[of "{(X,Y). \<exists> s. wls s X \<and> (X,Y) \<in> rel s}"])
subgoal using rel wls_imp_good by blast
subgoal using Var by auto
subgoal by (auto intro!: Op simp: wlsInp_iff wlsBinp_iff liftAll_def liftAll2_def)
subgoal using Abs rel by simp blast .
thus ?thesis by auto
qed
theorem wls_rawInduct[case_names Var Op Abs]:
assumes
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 phi (arOf delta) inp; liftAll2 phiAbs (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs: "\<And> s xs x X. \<lbrakk>isInBar (xs,s); wls s X; phi s X\<rbrakk> \<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
by (induct rule: wls_templateInduct[of "\<lambda>s. Id"]) (simp_all add: assms)
-subsubsection {* Fresh induction *}
-
-text {* First for an unspecified notion of parameter: *}
+subsubsection \<open>Fresh induction\<close>
+
+text \<open>First for an unspecified notion of parameter:\<close>
theorem wls_templateInduct_fresh[case_names Par Rel Var Op Abs]:
fixes s X xs s' A phi phiAbs rel
and vars :: "'varSort \<Rightarrow> 'var set"
and terms :: "'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term set"
and abs :: "('varSort * 'sort) \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs set"
and envs :: "('index,'bindex,'varSort,'var,'opSym)env set"
assumes
PAR:
"\<And> xs us s.
( |vars xs| <o |UNIV :: 'var set| \<or> finite (vars xs)) \<and>
( |terms s| <o |UNIV :: 'var set| \<or> finite (terms s)) \<and>
( |abs (us,s)| <o |UNIV :: 'var set| \<or> finite (abs (us,s))) \<and>
(\<forall> X \<in> terms s. wls s X) \<and>
(\<forall> A \<in> abs (us,s). wlsAbs (us,s) A) \<and>
( |envs| <o |UNIV :: 'var set| \<or> finite (envs)) \<and>
(\<forall> rho \<in> envs. wlsEnv rho)" and
rel: "\<And> s X Y. \<lbrakk>wls s X; (X,Y) \<in> rel s\<rbrakk> \<Longrightarrow> wls s Y \<and> skel Y = skel X" and
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 (\<lambda>s X. phi s X) (arOf delta) inp;
liftAll2 (\<lambda>(us,s) A. phiAbs (us,s) A) (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs:
"\<And> s xs x X.
\<lbrakk>isInBar (xs,s); wls s X;
x \<notin> vars xs;
\<And> s' Y. Y \<in> terms s' \<Longrightarrow> fresh xs x Y;
\<And> xs' s' A. A \<in> abs (xs',s') \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envs \<Longrightarrow> freshEnv xs x rho;
\<And> Y. (X,Y) \<in> rel s \<Longrightarrow> phi s Y\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
proof-
let ?terms = "\<Union> s. terms s"
let ?abs = "\<Union> xs s. abs (xs,s)"
have "\<forall> s. |terms s| <o |UNIV :: 'var set|"
using PAR var_infinite_INNER finite_ordLess_infinite2 by blast
hence 1:"|\<Union>s. terms s| <o |UNIV :: 'var set|"
using sort_lt_var_INNER var_regular_INNER regular_UNION by blast
have "\<forall> us s. |abs (us,s)| <o |UNIV :: 'var set|"
using PAR var_infinite_INNER finite_ordLess_infinite2 by blast
hence "\<forall> us. |\<Union>s. abs (us,s)| <o |UNIV :: 'var set|"
by(auto simp add: sort_lt_var_INNER var_regular_INNER regular_UNION)
hence 2: "|\<Union> us s. abs (us,s)| <o |UNIV :: 'var set|"
using varSort_lt_var_INNER var_regular_INNER by(auto simp add: regular_UNION)
(* *)
have "(good X \<longrightarrow> (\<forall> s. wls s X \<longrightarrow> phi s X)) \<and>
(goodAbs A \<longrightarrow> (\<forall> xs s. wlsAbs (xs,s) A \<longrightarrow> phiAbs (xs,s) A))"
apply(induct rule: term_templateInduct_fresh
[of vars ?terms ?abs envs
"{(X,Y). \<exists> s. wls s X \<and> (X,Y) \<in> rel s}"])
subgoal for xs
using PAR 1 2 apply simp_all using wls_imp_good wlsAbs_imp_goodAbs by blast+
subgoal using assms by simp (meson wls_imp_good)
subgoal using assms by simp
subgoal using assms by simp
(smt liftAll2_def liftAll_def option.distinct(1)
option.sel wlsBinp.cases wlsInp_iff)
subgoal using assms by simp metis .
thus ?thesis by auto
qed
-text{* A version of the above not employing any relation for
- the abstraction case: *}
+text\<open>A version of the above not employing any relation for
+ the abstraction case:\<close>
theorem wls_rawInduct_fresh[case_names Par Var Op Abs]:
fixes s X xs s' A phi phiAbs
and vars :: "'varSort \<Rightarrow> 'var set"
and terms :: "'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term set"
and abs :: "('varSort * 'sort) \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs set"
and envs :: "('index,'bindex,'varSort,'var,'opSym)env set"
assumes
PAR:
"\<And> xs us s.
( |vars xs| <o |UNIV :: 'var set| \<or> finite (vars xs)) \<and>
( |terms s| <o |UNIV :: 'var set| \<or> finite (terms s)) \<and>
(\<forall> X \<in> terms s. wls s X) \<and>
( |abs (us,s)| <o |UNIV :: 'var set| \<or> finite (abs (us,s))) \<and>
(\<forall> A \<in> abs (us,s). wlsAbs (us,s) A) \<and>
( |envs| <o |UNIV :: 'var set| \<or> finite (envs)) \<and>
(\<forall> rho \<in> envs. wlsEnv rho)" and
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 (\<lambda>s X. phi s X) (arOf delta) inp;
liftAll2 (\<lambda>(us,s) A. phiAbs (us,s) A) (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs:
"\<And> s xs x X.
\<lbrakk>isInBar (xs,s); wls s X;
x \<notin> vars xs;
\<And> s' Y. Y \<in> terms s' \<Longrightarrow> fresh xs x Y;
\<And> us s' A. A \<in> abs (us,s') \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envs \<Longrightarrow> freshEnv xs x rho;
phi s X\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
apply(induct rule: wls_templateInduct_fresh[of vars terms abs envs "\<lambda>s. Id"])
using assms by auto
-text{* Then for our notion of sorted parameter: *}
+text\<open>Then for our notion of sorted parameter:\<close>
theorem wls_induct_fresh[case_names Par Var Op Abs]:
fixes X :: "('index,'bindex,'varSort,'var,'opSym)term" and s and
A :: "('index,'bindex,'varSort,'var,'opSym)abs" and xs s' and
P :: "('index,'bindex,'varSort,'var,'opSym,'sort)paramS" and phi phiAbs
assumes
P: "wlsPar P" and
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 (\<lambda>s X. phi s X) (arOf delta) inp;
liftAll2 (\<lambda>(us,s) A. phiAbs (us,s) A) (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs:
"\<And> s xs x X.
\<lbrakk>isInBar (xs,s); wls s X;
x \<notin> varsOfS P xs;
\<And> s' Y. Y \<in> termsOfS P s' \<Longrightarrow> fresh xs x Y;
\<And> us s' A. A \<in> absOfS P (us,s') \<Longrightarrow> freshAbs xs x A;
\<And> rho. rho \<in> envsOfS P \<Longrightarrow> freshEnv xs x rho;
phi s X\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
proof(induct rule: wls_rawInduct_fresh
[of "varsOfS P" "termsOfS P" "absOfS P" "envsOfS P" _ _ s X xs s' A])
case (Par xs us s)
then show ?case using assms by(cases P) simp
qed(insert assms, simp_all)
-subsubsection {* The syntactic constructs are almost free (on well-sorted terms) *}
+subsubsection \<open>The syntactic constructs are almost free (on well-sorted terms)\<close>
(* Recall theorem Var_inj. *)
theorem wls_Op_inj[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
and "wlsInp delta' inp'" and "wlsBinp delta' binp'"
shows
"(Op delta inp binp = Op delta' inp' binp') =
(delta = delta' \<and> inp = inp' \<and> binp = binp')"
using assms by simp
lemma wls_Abs_ainj_all:
assumes "wls s X" and "wls s' X'"
shows
"(Abs xs x X = Abs xs' x' X') =
(xs = xs' \<and>
(\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X') \<longrightarrow>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
using assms by(simp add: Abs_ainj_all)
theorem wls_Abs_swap_all:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X') \<longrightarrow>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs))"
using assms by(simp add: wls_Abs_ainj_all)
lemma wls_Abs_ainj_ex:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs' x' X') =
(xs = xs' \<and>
(\<exists> y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X' \<and>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)))"
using assms by(simp add: Abs_ainj_ex)
theorem wls_Abs_swap_ex:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<exists> y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X' \<and>
(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs))"
using assms by(simp add: wls_Abs_ainj_ex)
theorem wls_Abs_inj[simp]:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x X') =
(X = X')"
using assms by (auto simp: wls_Abs_swap_all)
theorem wls_Abs_swap_cong[fundef_cong]:
assumes "wls s X" and "wls s X'"
and "fresh xs y X" and "fresh xs y X'" and "(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)"
shows "Abs xs x X = Abs xs x' X'"
using assms by (intro Abs_cong) auto
theorem wls_Abs_swap_fresh[simp]:
assumes "wls s X" and "fresh xs x' X"
shows "Abs xs x' (X #[x' \<and> x]_xs) = Abs xs x X"
using assms by(simp add: Abs_swap_fresh)
theorem wls_Var_diff_Op[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows "Var xs x \<noteq> Op delta inp binp"
using assms by auto
theorem wls_Op_diff_Var[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows "Op delta inp binp \<noteq> Var xs x"
using assms by auto
theorem wls_nchotomy:
assumes "wls s X"
shows
"(\<exists> xs x. asSort xs = s \<and> X = Var xs x) \<or>
(\<exists> delta inp binp. stOf delta = s \<and> wlsInp delta inp \<and> wlsBinp delta binp
\<and> X = Op delta inp binp)"
using assms wls.simps by force
lemmas wls_cases = wls_wlsAbs_wlsInp_wlsBinp.inducts(1)
lemmas wlsAbs_nchotomy = wlsAbs_inversion2
theorem wlsAbs_cases:
assumes "wlsAbs (xs,s) A"
and "\<And> x X. \<lbrakk>isInBar (xs,s); wls s X\<rbrakk> \<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows "phiAbs (xs,s) A"
using assms wlsAbs_nchotomy by blast
lemma wls_disjoint:
assumes "wls s X" and "wls s' X"
shows "s = s'"
using assms term_nchotomy wls_imp_good by fastforce
lemma wlsAbs_disjoint:
assumes "wlsAbs (xs,s) A" and "wlsAbs (xs',s') A"
shows "xs = xs' \<and> s = s'"
using assms abs_nchotomy wlsAbs_imp_goodAbs wls_disjoint by fastforce
lemmas wls_freeCons =
Var_inj wls_Op_inj wls_Var_diff_Op wls_Op_diff_Var wls_Abs_swap_fresh
-subsection {* The non-construct operators preserve well-sortedness *}
+subsection \<open>The non-construct operators preserve well-sortedness\<close>
lemma idEnv_preserves_wls[simp]:
"wlsEnv idEnv"
proof-
have "goodEnv idEnv" by simp
thus ?thesis unfolding wlsEnv_def goodEnv_def liftAll_def idEnv_def by auto
qed
lemma updEnv_preserves_wls[simp]:
assumes "wlsEnv rho" and "wls (asSort xs) X"
shows "wlsEnv (rho [x \<leftarrow> X]_xs)"
proof-
{fix ys
let ?L = "{y. rho ys y \<noteq> None}"
let ?R = "{y. (rho [x \<leftarrow> X]_xs) ys y \<noteq> None}"
have "?R \<le> ?L Un {x}" by auto
hence "|?R| \<le>o |?L Un {x}|" by simp
moreover
{have "|?L| <o |UNIV :: 'var set|"
using assms unfolding wlsEnv_def by simp
moreover have "|{x}| <o |UNIV :: 'var set|"
using var_infinite_INNER finite_ordLess_infinite by auto
ultimately have "|?L Un {x}| <o |UNIV :: 'var set|"
using var_infinite_INNER card_of_Un_ordLess_infinite by blast
}
ultimately have "|?R| <o |UNIV :: 'var set|"
using ordLeq_ordLess_trans by blast
} note 0 = this
have 1: "goodEnv (rho [x \<leftarrow> X]_xs)" using assms by simp
show ?thesis unfolding wlsEnv_def goodEnv_def
using 0 1 assms unfolding wlsEnv_def liftAll_def by auto
qed
lemma getEnv_preserves_wls[simp]:
assumes "wlsEnv rho" and "rho xs x = Some X"
shows "wls (asSort xs) X"
using assms unfolding wlsEnv_def liftAll_def by simp
lemmas envOps_preserve_wls =
idEnv_preserves_wls updEnv_preserves_wls
getEnv_preserves_wls
lemma psubstAll_preserves_wlsAll:
assumes P: "wlsPar P"
shows
"(wls s X \<longrightarrow> (\<forall> rho \<in> envsOfS P. wls s (X #[rho]))) \<and>
(wlsAbs (xs,s') A \<longrightarrow> (\<forall> rho \<in> envsOfS P. wlsAbs (xs,s') (A $[rho])))"
proof(induct rule: wls_induct_fresh[of P])
case (Var xs x)
show ?case
using assms apply safe subgoal for rho
apply(cases "rho xs x") apply simp_all
using getEnv_preserves_wls wlsPar_def by blast+ .
next
case (Op delta inp binp)
then show ?case using assms
by (auto simp:
wlsInp_iff psubstInp_def wlsBinp_iff psubstBinp_def liftAll2_def lift_def
sameDom_def intro!: Op_preserves_wls split: option.splits)
qed(insert assms, auto)
lemma psubst_preserves_wls[simp]:
"\<lbrakk>wls s X; wlsEnv rho\<rbrakk> \<Longrightarrow> wls s (X #[rho])"
using psubstAll_preserves_wlsAll[of "ParS (\<lambda>_. []) (\<lambda>_. []) (\<lambda>_. []) [rho]"]
unfolding wlsPar_def by auto
lemma psubstAbs_preserves_wls[simp]:
"\<lbrakk>wlsAbs (xs,s) A; wlsEnv rho\<rbrakk> \<Longrightarrow> wlsAbs (xs,s) (A $[rho])"
using psubstAll_preserves_wlsAll[of "ParS (\<lambda>_. []) (\<lambda>_. []) (\<lambda>_. []) [rho]"]
unfolding wlsPar_def by auto
lemma psubstInp_preserves_wls[simp]:
assumes "wlsInp delta inp" and "wlsEnv rho"
shows "wlsInp delta (inp %[rho])"
using assms by (auto simp: wlsInp_iff psubstInp_def liftAll2_def lift_def
sameDom_def intro!: Op_preserves_wls split: option.splits)
lemma psubstBinp_preserves_wls[simp]:
assumes "wlsBinp delta binp" and "wlsEnv rho"
shows "wlsBinp delta (binp %%[rho])"
using assms by (auto simp: wlsBinp_iff psubstBinp_def liftAll2_def lift_def
sameDom_def intro!: Op_preserves_wls split: option.splits)
lemma psubstEnv_preserves_wls[simp]:
assumes "wlsEnv rho" and "wlsEnv rho'"
shows "wlsEnv (rho &[rho'])"
proof-
{fix ys y Y
assume "(rho &[rho']) ys y = Some Y"
hence "wls (asSort ys) Y"
using assms unfolding psubstEnv_def wlsEnv_def liftAll_def
by (cases "rho ys y") (auto simp add: assms)
}
moreover have "goodEnv (rho &[rho'])" using assms by simp
ultimately show ?thesis
unfolding goodEnv_def wlsEnv_def psubstEnv_def wlsEnv_def liftAll_def
by (auto simp add: assms)
qed
lemmas psubstAll_preserve_wls =
psubst_preserves_wls psubstAbs_preserves_wls
psubstInp_preserves_wls psubstBinp_preserves_wls
psubstEnv_preserves_wls
lemma subst_preserves_wls[simp]:
assumes "wls s X" and "wls (asSort ys) Y"
shows "wls s (X #[Y / y]_ys)"
using assms unfolding subst_def by simp
lemma substAbs_preserves_wls[simp]:
assumes "wlsAbs (xs,s) A" and "wls (asSort ys) Y"
shows "wlsAbs (xs,s) (A $[Y / y]_ys)"
using assms unfolding substAbs_def by simp
lemma substInp_preserves_wls[simp]:
assumes "wlsInp delta inp" and "wls (asSort ys) Y"
shows "wlsInp delta (inp %[Y / y]_ys)"
using assms unfolding substInp_def by simp
lemma substBinp_preserves_wls[simp]:
assumes "wlsBinp delta binp" and "wls (asSort ys) Y"
shows "wlsBinp delta (binp %%[Y / y]_ys)"
using assms unfolding substBinp_def by simp
lemma substEnv_preserves_wls[simp]:
assumes "wlsEnv rho" and "wls (asSort ys) Y"
shows "wlsEnv (rho &[Y / y]_ys)"
using assms unfolding substEnv_def by simp
lemmas substAll_preserve_wls =
subst_preserves_wls substAbs_preserves_wls
substInp_preserves_wls substBinp_preserves_wls
substEnv_preserves_wls
lemma vsubst_preserves_wls[simp]:
assumes "wls s Y"
shows "wls s (Y #[x1 // x]_xs)"
using assms unfolding vsubst_def by simp
lemma vsubstAbs_preserves_wls[simp]:
assumes "wlsAbs (us,s) A"
shows "wlsAbs (us,s) (A $[x1 // x]_xs)"
using assms unfolding vsubstAbs_def by simp
lemma vsubstInp_preserves_wls[simp]:
assumes "wlsInp delta inp"
shows "wlsInp delta (inp %[x1 // x]_xs)"
using assms unfolding vsubstInp_def by simp
lemma vsubstBinp_preserves_wls[simp]:
assumes "wlsBinp delta binp"
shows "wlsBinp delta (binp %%[x1 // x]_xs)"
using assms unfolding vsubstBinp_def by simp
lemma vsubstEnv_preserves_wls[simp]:
assumes "wlsEnv rho"
shows "wlsEnv (rho &[x1 // x]_xs)"
using assms unfolding vsubstEnv_def by simp
lemmas vsubstAll_preserve_wls = vsubst_preserves_wls vsubstAbs_preserves_wls
vsubstInp_preserves_wls vsubstBinp_preserves_wls vsubstEnv_preserves_wls
lemmas all_preserve_wls = Cons_preserve_wls swapAll_preserve_wls psubstAll_preserve_wls envOps_preserve_wls
substAll_preserve_wls vsubstAll_preserve_wls
-subsection {* Simplification rules for swapping, substitution, freshness and skeleton *}
+subsection \<open>Simplification rules for swapping, substitution, freshness and skeleton\<close>
(* Recall theorem swap_Var_simp. *)
theorem wls_swap_Op_simp[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows
"((Op delta inp binp) #[x1 \<and> x2]_xs) =
Op delta (inp %[x1 \<and> x2]_xs) (binp %%[x1 \<and> x2]_xs)"
using assms by simp
theorem wls_swapAbs_simp[simp]:
assumes "wls s X"
shows "((Abs xs x X) $[y1 \<and> y2]_ys) = Abs xs (x @xs[y1 \<and> y2]_ys) (X #[y1 \<and> y2]_ys)"
using assms by simp
lemmas wls_swapAll_simps =
swap_Var_simp wls_swap_Op_simp wls_swapAbs_simp
(* Recall theorem fresh_Var_simp. *)
theorem wls_fresh_Op_simp[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows
"fresh xs x (Op delta inp binp) =
(freshInp xs x inp \<and> freshBinp xs x binp)"
using assms by simp
theorem wls_freshAbs_simp[simp]:
assumes "wls s X"
shows "freshAbs ys y (Abs xs x X) = (ys = xs \<and> y = x \<or> fresh ys y X)"
using assms by simp
lemmas wls_freshAll_simps =
fresh_Var_simp wls_fresh_Op_simp wls_freshAbs_simp
(* Recall theorem skel_Var_simp *)
theorem wls_skel_Op_simp[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows
"skel (Op delta inp binp) = Branch (skelInp inp) (skelBinp binp)"
using assms by simp
(* The next is not a simplification rule, but belongs here: *)
lemma wls_skelInp_def2:
assumes "wlsInp delta inp"
shows "skelInp inp = lift skel inp"
using assms by(simp add: skelInp_def2)
lemma wls_skelBinp_def2:
assumes "wlsBinp delta binp"
shows "skelBinp binp = lift skelAbs binp"
using assms by(simp add: skelBinp_def2)
theorem wls_skelAbs_simp[simp]:
assumes "wls s X"
shows "skelAbs (Abs xs x X) = Branch (\<lambda>i. Some (skel X)) Map.empty"
using assms by simp
lemmas wls_skelAll_simps =
skel_Var_simp wls_skel_Op_simp wls_skelAbs_simp
theorem wls_psubst_Var_simp1[simp]:
assumes "wlsEnv rho" and "rho xs x = None"
shows "((Var xs x) #[rho]) = Var xs x"
using assms by simp
theorem wls_psubst_Var_simp2[simp]:
assumes "wlsEnv rho" and "rho xs x = Some X"
shows "((Var xs x) #[rho]) = X"
using assms by simp
theorem wls_psubst_Op_simp[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp" and "wlsEnv rho"
shows
"((Op delta inp binp) #[rho]) = Op delta (inp %[rho]) (binp %%[rho])"
using assms by simp
theorem wls_psubstAbs_simp[simp]:
assumes "wls s X" and "wlsEnv rho" and "freshEnv xs x rho"
shows "((Abs xs x X) $[rho]) = Abs xs x (X #[rho])"
using assms by simp
lemmas wls_psubstAll_simps =
wls_psubst_Var_simp1 wls_psubst_Var_simp2 wls_psubst_Op_simp wls_psubstAbs_simp
(* Recall lemmas getEnv_idEnv, getEnv_updEnv1 and getEnv_updEnv2. *)
lemmas wls_envOps_simps =
getEnv_idEnv getEnv_updEnv1 getEnv_updEnv2
theorem wls_subst_Var_simp1[simp]:
assumes "wls (asSort ys) Y"
and "ys \<noteq> xs \<or> y \<noteq> x "
shows "((Var xs x) #[Y / y]_ys) = Var xs x"
using assms unfolding subst_def by auto
theorem wls_subst_Var_simp2[simp]:
assumes "wls (asSort xs) Y"
shows "((Var xs x) #[Y / x]_xs) = Y"
using assms unfolding subst_def by auto
theorem wls_subst_Op_simp[simp]:
assumes "wls (asSort ys) Y"
and "wlsInp delta inp" and "wlsBinp delta binp"
shows
"((Op delta inp binp) #[Y / y]_ys) =
Op delta (inp %[Y / y]_ys) (binp %%[Y / y]_ys)"
using assms unfolding subst_def substInp_def
substAbs_def substBinp_def by auto
theorem wls_substAbs_simp[simp]:
assumes "wls (asSort ys) Y"
and "wls s X" and "xs \<noteq> ys \<or> x \<noteq> y" and "fresh xs x Y"
shows "((Abs xs x X) $[Y / y]_ys) = Abs xs x (X #[Y / y]_ys)"
proof-
have "freshEnv xs x (idEnv [y \<leftarrow> Y]_ys)" unfolding freshEnv_def liftAll_def
using assms by simp
thus ?thesis using assms unfolding subst_def substAbs_def by auto
qed
lemmas wls_substAll_simps =
wls_subst_Var_simp1 wls_subst_Var_simp2 wls_subst_Op_simp wls_substAbs_simp
(* Recall theorem vsubst_Var_simp. *)
theorem wls_vsubst_Op_simp[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows
"((Op delta inp binp) #[y1 // y]_ys) =
Op delta (inp %[y1 // y]_ys) (binp %%[y1 // y]_ys)"
using assms unfolding vsubst_def vsubstInp_def
vsubstAbs_def vsubstBinp_def by simp
theorem wls_vsubstAbs_simp[simp]:
assumes "wls s X" and
"xs \<noteq> ys \<or> x \<notin> {y,y1}"
shows "((Abs xs x X) $[y1 // y]_ys) = Abs xs x (X #[y1 // y]_ys)"
using assms unfolding vsubst_def vsubstAbs_def by simp
lemmas wls_vsubstAll_simps =
vsubst_Var_simp wls_vsubst_Op_simp wls_vsubstAbs_simp
(* Recall theorem swap_swapped. *)
theorem wls_swapped_skel:
assumes "wls s X" and "(X,Y) \<in> swapped"
shows "skel Y = skel X"
apply(rule swapped_skel) using assms by auto
theorem wls_obtain_rep:
assumes "wls s X" and FRESH: "fresh xs x' X"
shows "\<exists> X'. skel X' = skel X \<and> (X,X') \<in> swapped \<and> wls s X' \<and> Abs xs x X = Abs xs x' X'"
proof-
have 0: "skel (X #[x' \<and> x]_xs) = skel X" using assms by(simp add: skel_swap)
have 1: "wls s (X #[x' \<and> x]_xs)" using assms swap_preserves_wls by auto
have 2: "(X, X #[x' \<and> x]_xs) \<in> swapped" using Var swap_swapped by auto
show ?thesis using assms 0 1 2 by fastforce
qed
lemmas wls_allOpers_simps =
wls_swapAll_simps
wls_freshAll_simps
wls_skelAll_simps
wls_envOps_simps
wls_psubstAll_simps
wls_substAll_simps
wls_vsubstAll_simps
-subsection {* The ability to pick fresh variables *}
+subsection \<open>The ability to pick fresh variables\<close>
theorem wls_single_non_fresh_ordLess_var:
"wls s X \<Longrightarrow> |{x. \<not> fresh xs x X}| <o |UNIV :: 'var set|"
by(simp add: single_non_fresh_ordLess_var)
theorem wls_single_non_freshAbs_ordLess_var:
"wlsAbs (us,s) A \<Longrightarrow> |{x. \<not> freshAbs xs x A}| <o |UNIV :: 'var set|"
by(simp add: single_non_freshAbs_ordLess_var)
theorem wls_obtain_fresh:
fixes V::"'varSort \<Rightarrow> 'var set" and
XS::"'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)term set" and
AS::"'varSort \<Rightarrow> 'sort \<Rightarrow> ('index,'bindex,'varSort,'var,'opSym)abs set" and
Rho::"('index,'bindex,'varSort,'var,'opSym)env set" and zs
assumes VVar: "\<forall> xs. |V xs| <o |UNIV :: 'var set| \<or> finite (V xs)"
and XSVar: "\<forall> s. |XS s| <o |UNIV :: 'var set| \<or> finite (XS s)"
and ASVar: "\<forall> xs s. |AS xs s| <o |UNIV :: 'var set| \<or> finite (AS xs s)"
and XSwls: "\<forall> s. \<forall> X \<in> XS s. wls s X"
and ASwls: "\<forall> xs s. \<forall> A \<in> AS xs s. wlsAbs (xs,s) A"
and RhoVar: "|Rho| <o |UNIV :: 'var set| \<or> finite Rho"
and Rhowls: "\<forall> rho \<in> Rho. wlsEnv rho"
shows
"\<exists> z. (\<forall> xs. z \<notin> V xs) \<and>
(\<forall> s. \<forall> X \<in> XS s. fresh zs z X) \<and>
(\<forall> xs s. \<forall> A \<in> AS xs s. freshAbs zs z A) \<and>
(\<forall> rho \<in> Rho. freshEnv zs z rho)"
proof-
let ?VG = "\<Union> xs. V xs" let ?XSG = "\<Union> s. XS s" let ?ASG = "\<Union> xs s. AS xs s"
have "\<forall> xs. |V xs| <o |UNIV :: 'var set|" using VVar finite_ordLess_var by auto
hence 1: "|?VG| <o |UNIV :: 'var set|"
using var_regular_INNER varSort_lt_var_INNER regular_UNION by blast
have "\<forall> s. |XS s| <o |UNIV :: 'var set|" using XSVar finite_ordLess_var by auto
hence 2: "|?XSG| <o |UNIV :: 'var set|"
using var_regular_INNER sort_lt_var_INNER regular_UNION by blast
have "\<forall> xs s. |AS xs s| <o |UNIV :: 'var set|" using ASVar finite_ordLess_var by auto
hence "\<forall> xs. |\<Union> s. AS xs s| <o |UNIV :: 'var set|"
using var_regular_INNER sort_lt_var_INNER regular_UNION by blast
hence 3: "|?ASG| <o |UNIV :: 'var set|"
using var_regular_INNER varSort_lt_var_INNER by (auto simp add: regular_UNION)
have "\<exists> z. z \<notin> ?VG \<and>
(\<forall> X \<in> ?XSG. fresh zs z X) \<and>
(\<forall> A \<in> ?ASG. freshAbs zs z A) \<and>
(\<forall> rho \<in> Rho. freshEnv zs z rho)"
using assms 1 2 3 by (intro obtain_fresh) fastforce+
thus ?thesis by auto
qed
theorem wls_obtain_fresh_paramS:
assumes "wlsPar P"
shows
"\<exists> z.
(\<forall> xs. z \<notin> varsOfS P xs) \<and>
(\<forall> s. \<forall> X \<in> termsOfS P s. fresh zs z X) \<and>
(\<forall> us s. \<forall> A \<in> absOfS P (us,s). freshAbs zs z A) \<and>
(\<forall> rho \<in> envsOfS P. freshEnv zs z rho)"
using assms by(cases P) (auto intro: wls_obtain_fresh)
lemma wlsAbs_freshAbs_nchotomy:
assumes A: "wlsAbs (xs,s) A" and fresh: "freshAbs xs x A"
shows "\<exists> X. wls s X \<and> A = Abs xs x X"
proof-
{assume "wlsAbs (xs,s) A"
hence "freshAbs xs x A \<longrightarrow> (\<exists> X. wls s X \<and> A = Abs xs x X)"
using fresh wls_obtain_rep[of s _ xs x] by (fastforce elim!: wlsAbs_cases)
}
thus ?thesis using assms by auto
qed
theorem wlsAbs_fresh_nchotomy:
assumes A: "wlsAbs (xs,s) A" and P: "wlsPar P"
shows "\<exists> x X. A = Abs xs x X \<and>
wls s X \<and>
(\<forall> ys. x \<notin> varsOfS P ys) \<and>
(\<forall> s'. \<forall> Y \<in> termsOfS P s'. fresh xs x Y) \<and>
(\<forall> us s'. \<forall> B \<in> absOfS P (us,s'). freshAbs xs x B) \<and>
(\<forall> rho \<in> envsOfS P. freshEnv xs x rho)"
proof-
let ?chi =
"\<lambda> x. (\<forall> xs. x \<notin> varsOfS P xs) \<and>
(\<forall> s'. \<forall> Y \<in> termsOfS P s'. fresh xs x Y) \<and>
(\<forall> us s'.\<forall> B \<in> (if us = xs \<and> s' = s then {A} else {}) \<union> absOfS P (us,s'). freshAbs xs x B) \<and>
(\<forall> rho \<in> envsOfS P. freshEnv xs x rho)"
have "\<exists> x. ?chi x"
using A P by (intro wls_obtain_fresh) (cases P, auto)+
then obtain x where 1: "?chi x" by blast
hence "freshAbs xs x A" by fastforce
then obtain X where X: "wls s X" and 2: "A = Abs xs x X"
using A 1 wlsAbs_freshAbs_nchotomy[of xs s A x] by auto
thus ?thesis using 1 by blast
qed
theorem wlsAbs_fresh_cases:
assumes "wlsAbs (xs,s) A" and "wlsPar P"
and "\<And> x X.
\<lbrakk>wls s X;
\<And> ys. x \<notin> varsOfS P ys;
\<And> s' Y. Y \<in> termsOfS P s' \<Longrightarrow> fresh xs x Y;
\<And> us s' B. B \<in> absOfS P (us,s') \<Longrightarrow> freshAbs xs x B;
\<And> rho. rho \<in> envsOfS P \<Longrightarrow> freshEnv xs x rho\<rbrakk>
\<Longrightarrow> phi (xs,s) (Abs xs x X) P"
shows "phi (xs,s) A P"
by (metis assms wlsAbs_fresh_nchotomy)
-subsection {* Compositionality properties of freshness and swapping *}
-
-subsubsection {* W.r.t. terms *}
+subsection \<open>Compositionality properties of freshness and swapping\<close>
+
+subsubsection \<open>W.r.t. terms\<close>
theorem wls_swap_ident[simp]:
assumes "wls s X"
shows "(X #[x \<and> x]_xs) = X"
using assms by simp
theorem wls_swap_compose:
assumes "wls s X"
shows "((X #[x \<and> y]_zs) #[x' \<and> y']_zs') =
((X #[x' \<and> y']_zs') #[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]_zs)"
using assms by (intro swap_compose) auto
theorem wls_swap_commute:
"\<lbrakk>wls s X; zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {}\<rbrakk> \<Longrightarrow>
((X #[x \<and> y]_zs) #[x' \<and> y']_zs') = ((X #[x' \<and> y']_zs') #[x \<and> y]_zs)"
by (intro swap_commute) auto
theorem wls_swap_involutive[simp]:
assumes "wls s X"
shows "((X #[x \<and> y]_zs) #[x \<and> y]_zs) = X"
using assms by simp
theorem wls_swap_inj[simp]:
assumes "wls s X" and "wls s X'"
shows
"((X #[x \<and> y]_zs) = (X' #[x \<and> y]_zs)) =
(X = X')"
using assms by (metis wls_swap_involutive)
(* Recall theorem swap_sym. *)
theorem wls_swap_involutive2[simp]:
assumes "wls s X"
shows "((X #[x \<and> y]_zs) #[y \<and> x]_zs) = X"
using assms by (simp add: swap_sym)
theorem wls_swap_preserves_fresh[simp]:
assumes "wls s X"
shows "fresh xs (x @xs[y1 \<and> y2]_ys) (X #[y1 \<and> y2]_ys) = fresh xs x X"
using assms by simp
theorem wls_swap_preserves_fresh_distinct:
assumes "wls s X" and
"xs \<noteq> ys \<or> x \<notin> {y1,y2}"
shows "fresh xs x (X #[y1 \<and> y2]_ys) = fresh xs x X"
using assms by(intro swap_preserves_fresh_distinct) auto
theorem wls_fresh_swap_exchange1:
assumes "wls s X"
shows "fresh xs x2 (X #[x1 \<and> x2]_xs) = fresh xs x1 X"
using assms by (intro fresh_swap_exchange1) auto
theorem wls_fresh_swap_exchange2:
assumes "wls s X"
shows "fresh xs x2 (X #[x2 \<and> x1]_xs) = fresh xs x1 X"
using assms by (intro fresh_swap_exchange2) fastforce+
theorem wls_fresh_swap_id[simp]:
assumes "wls s X" and "fresh xs x1 X" and "fresh xs x2 X"
shows "(X #[x1 \<and> x2]_xs) = X"
using assms by simp
theorem wls_fresh_swap_compose:
assumes "wls s X" and "fresh xs y X" and "fresh xs z X"
shows "((X #[y \<and> x]_xs) #[z \<and> y]_xs) = (X #[z \<and> x]_xs)"
using assms by (intro fresh_swap_compose) auto
theorem wls_skel_swap:
assumes "wls s X"
shows "skel (X #[x1 \<and> x2]_xs) = skel X"
using assms by (intro skel_swap) auto
-subsubsection {* W.r.t. environments *}
+subsubsection \<open>W.r.t. environments\<close>
theorem wls_swapEnv_ident[simp]:
assumes "wlsEnv rho"
shows "(rho &[x \<and> x]_xs) = rho"
using assms by simp
theorem wls_swapEnv_compose:
assumes "wlsEnv rho"
shows "((rho &[x \<and> y]_zs) &[x' \<and> y']_zs') =
((rho &[x' \<and> y']_zs') &[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]_zs)"
using assms by (intro swapEnv_compose) auto
theorem wls_swapEnv_commute:
"\<lbrakk>wlsEnv rho; zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {}\<rbrakk> \<Longrightarrow>
((rho &[x \<and> y]_zs) &[x' \<and> y']_zs') = ((rho &[x' \<and> y']_zs') &[x \<and> y]_zs)"
by (intro swapEnv_commute) fastforce+
theorem wls_swapEnv_involutive[simp]:
assumes "wlsEnv rho"
shows "((rho &[x \<and> y]_zs) &[x \<and> y]_zs) = rho"
using assms by simp
theorem wls_swapEnv_inj[simp]:
assumes "wlsEnv rho" and "wlsEnv rho'"
shows
"((rho &[x \<and> y]_zs) = (rho' &[x \<and> y]_zs)) =
(rho = rho')"
by (metis assms wls_swapEnv_involutive)
(* Recall theorem swapEnv_sym. *)
theorem wls_swapEnv_involutive2[simp]:
assumes "wlsEnv rho"
shows "((rho &[x \<and> y]_zs) &[y \<and> x]_zs) = rho"
using assms by(simp add: swapEnv_sym)
theorem wls_swapEnv_preserves_freshEnv[simp]:
assumes "wlsEnv rho"
shows "freshEnv xs (x @xs[y1 \<and> y2]_ys) (rho &[y1 \<and> y2]_ys) = freshEnv xs x rho"
using assms by simp
theorem wls_swapEnv_preserves_freshEnv_distinct:
assumes "wlsEnv rho"
"xs \<noteq> ys \<or> x \<notin> {y1,y2}"
shows "freshEnv xs x (rho &[y1 \<and> y2]_ys) = freshEnv xs x rho"
using assms by (intro swapEnv_preserves_freshEnv_distinct) auto
theorem wls_freshEnv_swapEnv_exchange1:
assumes "wlsEnv rho"
shows "freshEnv xs x2 (rho &[x1 \<and> x2]_xs) = freshEnv xs x1 rho"
using assms by (intro freshEnv_swapEnv_exchange1) auto
theorem wls_freshEnv_swapEnv_exchange2:
assumes "wlsEnv rho"
shows "freshEnv xs x2 (rho &[x2 \<and> x1]_xs) = freshEnv xs x1 rho"
using assms by (intro freshEnv_swapEnv_exchange2) auto
theorem wls_freshEnv_swapEnv_id[simp]:
assumes "wlsEnv rho" and "freshEnv xs x1 rho" and "freshEnv xs x2 rho"
shows "(rho &[x1 \<and> x2]_xs) = rho"
using assms by simp
theorem wls_freshEnv_swapEnv_compose:
assumes "wlsEnv rho" and "freshEnv xs y rho" and "freshEnv xs z rho"
shows "((rho &[y \<and> x]_xs) &[z \<and> y]_xs) = (rho &[z \<and> x]_xs)"
using assms by (intro freshEnv_swapEnv_compose) auto
-subsubsection {* W.r.t. abstractions *}
+subsubsection \<open>W.r.t. abstractions\<close>
theorem wls_swapAbs_ident[simp]:
"wlsAbs (us,s) A \<Longrightarrow> (A $[x \<and> x]_xs) = A"
by (elim wlsAbs_cases) auto
theorem wls_swapAbs_compose:
"wlsAbs (us,s) A \<Longrightarrow>
((A $[x \<and> y]_zs) $[x' \<and> y']_zs') =
((A $[x' \<and> y']_zs') $[(x @zs[x' \<and> y']_zs') \<and> (y @zs[x' \<and> y']_zs')]_zs)"
by (erule wlsAbs_cases) (simp, metis sw_compose wls_swap_compose)
theorem wls_swapAbs_commute:
assumes "zs \<noteq> zs' \<or> {x,y} \<inter> {x',y'} = {}"
shows
"wlsAbs (us,s) A \<Longrightarrow>
((A $[x \<and> y]_zs) $[x' \<and> y']_zs') = ((A $[x' \<and> y']_zs') $[x \<and> y]_zs)"
using assms by (elim wlsAbs_cases) (simp add: sw_commute wls_swap_commute)
theorem wls_swapAbs_involutive[simp]:
"wlsAbs (us,s) A \<Longrightarrow> ((A $[x \<and> y]_zs) $[x \<and> y]_zs) = A"
by (erule wlsAbs_cases) simp_all
theorem wls_swapAbs_sym:
"wlsAbs (us,s) A \<Longrightarrow> (A $[x \<and> y]_zs) = (A $[y \<and> x]_zs)"
by (erule wlsAbs_cases) (auto simp add: swap_sym sw_sym)
theorem wls_swapAbs_inj[simp]:
assumes "wlsAbs (us,s) A" and "wlsAbs (us,s) A'"
shows
"((A $[x \<and> y]_zs) = (A' $[x \<and> y]_zs)) =
(A = A')"
by (metis assms wls_swapAbs_involutive)
theorem wls_swapAbs_involutive2[simp]:
"wlsAbs (us,s) A \<Longrightarrow> ((A $[x \<and> y]_zs) $[y \<and> x]_zs) = A"
using wls_swapAbs_sym[of us s A zs x y] by auto
theorem wls_swapAbs_preserves_freshAbs[simp]:
"wlsAbs (us,s) A
\<Longrightarrow> freshAbs xs (x @xs[y1 \<and> y2]_ys) (A $[y1 \<and> y2]_ys) = freshAbs xs x A"
by (erule wlsAbs_cases)
(simp_all add: sw_def wls_fresh_swap_exchange1 wls_fresh_swap_exchange2
wls_swap_preserves_fresh_distinct)
theorem wls_swapAbs_preserves_freshAbs_distinct:
"\<lbrakk>wlsAbs (us,s) A; xs \<noteq> ys \<or> x \<notin> {y1,y2}\<rbrakk>
\<Longrightarrow> freshAbs xs x (A $[y1 \<and> y2]_ys) = freshAbs xs x A"
apply(erule wlsAbs_cases) apply simp_all
unfolding sw_def by (auto simp: wls_swap_preserves_fresh_distinct)
theorem wls_freshAbs_swapAbs_exchange1:
"wlsAbs (us,s) A
\<Longrightarrow> freshAbs xs x2 (A $[x1 \<and> x2]_xs) = freshAbs xs x1 A"
apply(erule wlsAbs_cases) apply simp_all
unfolding sw_def by (auto simp add: wls_fresh_swap_exchange1)
theorem wls_freshAbs_swapAbs_exchange2:
"wlsAbs (us,s) A
\<Longrightarrow> freshAbs xs x2 (A $[x2 \<and> x1]_xs) = freshAbs xs x1 A"
apply(erule wlsAbs_cases) apply simp_all
unfolding sw_def by (auto simp add: wls_fresh_swap_exchange2)
theorem wls_freshAbs_swapAbs_id[simp]:
assumes "wlsAbs (us,s) A"
and "freshAbs xs x1 A" and "freshAbs xs x2 A"
shows "(A $[x1 \<and> x2]_xs) = A"
using assms by simp
lemma wls_freshAbs_swapAbs_compose_aux:
"\<lbrakk>wlsAbs (us,s) A; wlsPar P\<rbrakk> \<Longrightarrow>
\<forall> x y z. {x,y,z} \<subseteq> varsOfS P xs \<and> freshAbs xs y A \<and> freshAbs xs z A \<longrightarrow>
((A $[y \<and> x]_xs) $[z \<and> y]_xs) = (A $[z \<and> x]_xs)"
apply(erule wlsAbs_fresh_cases)
by simp_all (metis fresh_swap_compose sw_def wls_imp_good)
theorem wls_freshAbs_swapAbs_compose:
assumes "wlsAbs (us,s) A"
and "freshAbs xs y A" and "freshAbs xs z A"
shows "((A $[y \<and> x]_xs) $[z \<and> y]_xs) = (A $[z \<and> x]_xs)"
proof-
let ?P =
"ParS (\<lambda>xs'. if xs' = xs then [x,y,z] else []) (\<lambda>s.[]) (\<lambda>_. []) [] ::
('index, 'bindex, 'varSort, 'var, 'opSym, 'sort) paramS"
show ?thesis
using assms wls_freshAbs_swapAbs_compose_aux[of us s A ?P xs]
unfolding wlsPar_def by simp
qed
theorem wls_skelAbs_swapAbs:
"wlsAbs (us,s) A
\<Longrightarrow> skelAbs (A $[x1 \<and> x2]_xs) = skelAbs A"
by (erule wlsAbs_cases) (auto simp: wls_skel_swap)
lemmas wls_swapAll_freshAll_otherSimps =
wls_swap_ident wls_swap_involutive wls_swap_inj wls_swap_involutive2 wls_swap_preserves_fresh wls_fresh_swap_id
wls_swapAbs_ident wls_swapAbs_involutive wls_swapAbs_inj wls_swapAbs_involutive2 wls_swapAbs_preserves_freshAbs
wls_freshAbs_swapAbs_id
wls_swapEnv_ident wls_swapEnv_involutive wls_swapEnv_inj wls_swapEnv_involutive2 wls_swapEnv_preserves_freshEnv
wls_freshEnv_swapEnv_id
-subsection {* Compositionality properties for the other operators *}
-
-subsubsection {* Environment identity, update and ``get" versus other operators *}
+subsection \<open>Compositionality properties for the other operators\<close>
+
+subsubsection \<open>Environment identity, update and ``get" versus other operators\<close>
(* Recall lemmas getEnv_idEnv, getEnv_updEnv_idEnv, getEnv_updEnv1, getEnv_updEnv2,
subst_psubst_idEnv, vsubst_psubst_idEnv, substEnv_psubstEnv_idEnv, vsubstEnv_psubstEnv_idEnv,
freshEnv_idEnv,
swapEnv_idEnv, psubstEnv_idEnv, substEnv_idEnv, vsubstEnv_idEnv. *)
theorem wls_psubst_idEnv[simp]:
"wls s X \<Longrightarrow> (X #[idEnv]) = X"
by simp
theorem wls_psubstEnv_idEnv_id[simp]:
"wlsEnv rho \<Longrightarrow> (rho &[idEnv]) = rho"
by simp
(* Recall lemmas updEnv_overwrite, updEnv_commute,
freshEnv_updEnv_E1, freshEnv_updEnv_E2, freshEnv_updEnv_E3, freshEnv_updEnv_E4,
freshEnv_updEnv_I,
swapEnv_updEnv *)
theorem wls_swapEnv_updEnv_fresh:
assumes "zs \<noteq> ys \<or> y \<notin> {z1,z2}" and "wls (asSort ys) Y"
and "fresh zs z1 Y" and "fresh zs z2 Y"
shows "((rho [y \<leftarrow> Y]_ys) &[z1 \<and> z2]_zs) = ((rho &[z1 \<and> z2]_zs) [y \<leftarrow> Y]_ys)"
using assms by (simp add: swapEnv_updEnv_fresh)
(* Recall lemmas psubstEnv_updEnv, psubstEnv_updEnv_idEnv,
substEnv_updEnv, vsubstEnv_updEnv, getEnv_ext *)
(* Recall lemmas freshEnv_getEnv1 freshEnv_getEnv2 freshEnv_getEnv swapEnv_getEnv1 swapEnv_getEnv2
getEnv_psubstEnv_None getEnv_psubstEnv_Some
getEnv_substEnv1 getEnv_substEnv2 getEnv_substEnv3 getEnv_substEnv4
getEnv_vsubstEnv1 getEnv_vsubstEnv2 getEnv_vsubstEnv3 getEnv_vsubstEnv4. *)
-subsubsection {* Substitution versus other operators *}
+subsubsection \<open>Substitution versus other operators\<close>
(* Recall definition freshImEnvAt_def. *)
theorem wls_fresh_psubst:
assumes "wls s X" and "wlsEnv rho"
shows
"fresh zs z (X #[rho]) =
(\<forall> ys y. fresh ys y X \<or> freshImEnvAt zs z rho ys y)"
using assms by(simp add: fresh_psubst)
theorem wls_fresh_psubst_E1:
assumes "wls s X" and "wlsEnv rho"
and "rho ys y = None" and "fresh zs z (X #[rho])"
shows "fresh ys y X \<or> (ys \<noteq> zs \<or> y \<noteq> z)"
using assms fresh_psubst_E1[of X rho ys y zs z] by simp
theorem wls_fresh_psubst_E2:
assumes "wls s X" and "wlsEnv rho"
and "rho ys y = Some Y" and "fresh zs z (X #[rho])"
shows "fresh ys y X \<or> fresh zs z Y"
using assms fresh_psubst_E2[of X rho ys y Y zs z] by simp
theorem wls_fresh_psubst_I1:
assumes "wls s X" and "wlsEnv rho"
and "fresh zs z X" and "freshEnv zs z rho"
shows "fresh zs z (X #[rho])"
using assms by(simp add: fresh_psubst_I1)
theorem wls_psubstEnv_preserves_freshEnv:
assumes "wlsEnv rho" and "wlsEnv rho'"
and fresh: "freshEnv zs z rho" "freshEnv zs z rho'"
shows "freshEnv zs z (rho &[rho'])"
using assms by(simp add: psubstEnv_preserves_freshEnv)
theorem wls_fresh_psubst_I:
assumes "wls s X" and "wlsEnv rho"
and "rho zs z = None \<Longrightarrow> fresh zs z X" and
"\<And> ys y Y. rho ys y = Some Y \<Longrightarrow> fresh ys y X \<or> fresh zs z Y"
shows "fresh zs z (X #[rho])"
using assms by(simp add: fresh_psubst_I)
theorem wls_fresh_subst:
assumes "wls s X" and "wls (asSort ys) Y"
shows "fresh zs z (X #[Y / y]_ys) =
(((zs = ys \<and> z = y) \<or> fresh zs z X) \<and> (fresh ys y X \<or> fresh zs z Y))"
using assms by(simp add: fresh_subst)
theorem wls_fresh_vsubst:
assumes "wls s X"
shows "fresh zs z (X #[y1 // y]_ys) =
(((zs = ys \<and> z = y) \<or> fresh zs z X) \<and> (fresh ys y X \<or> (zs \<noteq> ys \<or> z \<noteq> y1)))"
using assms by(simp add: fresh_vsubst)
theorem wls_subst_preserves_fresh:
assumes "wls s X" and "wls (asSort ys) Y"
and "fresh zs z X" and "fresh zs z Y"
shows "fresh zs z (X #[Y / y]_ys)"
using assms by(simp add: subst_preserves_fresh)
theorem wls_substEnv_preserves_freshEnv:
assumes "wlsEnv rho" and "wls (asSort ys) Y"
and "freshEnv zs z rho" and "fresh zs z Y" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "freshEnv zs z (rho &[Y / y]_ys)"
using assms by(simp add: substEnv_preserves_freshEnv)
theorem wls_vsubst_preserves_fresh:
assumes "wls s X"
and "fresh zs z X" and "zs \<noteq> ys \<or> z \<noteq> y1"
shows "fresh zs z (X #[y1 // y]_ys)"
using assms by(simp add: vsubst_preserves_fresh)
theorem wls_vsubstEnv_preserves_freshEnv:
assumes "wlsEnv rho"
and "freshEnv zs z rho" and "zs \<noteq> ys \<or> z \<notin> {y,y1}"
shows "freshEnv zs z (rho &[y1 // y]_ys)"
using assms by(simp add: vsubstEnv_preserves_freshEnv)
theorem wls_fresh_fresh_subst[simp]:
assumes "wls (asSort ys) Y" and "wls s X"
and "fresh ys y Y"
shows "fresh ys y (X #[Y / y]_ys)"
using assms by(simp add: fresh_fresh_subst)
theorem wls_diff_fresh_vsubst[simp]:
assumes "wls s X"
and "y \<noteq> y1"
shows "fresh ys y (X #[y1 // y]_ys)"
using assms by(simp add: diff_fresh_vsubst)
theorem wls_fresh_subst_E1:
assumes "wls s X" and "wls (asSort ys) Y"
and "fresh zs z (X #[Y / y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "fresh zs z X"
using assms fresh_subst_E1[of X Y zs z ys y] by simp
theorem wls_fresh_vsubst_E1:
assumes "wls s X"
and "fresh zs z (X #[y1 // y]_ys)" and "zs \<noteq> ys \<or> z \<noteq> y"
shows "fresh zs z X"
using assms fresh_vsubst_E1[of X zs z ys y1 y] by simp
theorem wls_fresh_subst_E2:
assumes "wls s X" and "wls (asSort ys) Y"
and "fresh zs z (X #[Y / y]_ys)"
shows "fresh ys y X \<or> fresh zs z Y"
using assms fresh_subst_E2[of X Y zs z ys y] by simp
theorem wls_fresh_vsubst_E2:
assumes "wls s X"
and "fresh zs z (X #[y1 // y]_ys)"
shows "fresh ys y X \<or> zs \<noteq> ys \<or> z \<noteq> y1"
using assms fresh_vsubst_E2[of X zs z ys y1 y] by simp
theorem wls_psubst_cong[fundef_cong]:
assumes "wls s X" and "wlsEnv rho" and "wlsEnv rho'"
and "\<And> ys y. fresh ys y X \<or> rho ys y = rho' ys y"
shows "(X #[rho]) = (X #[rho'])"
using assms by (simp add: psubst_cong)
theorem wls_fresh_psubst_updEnv:
assumes "wls (asSort ys) Y" and "wls s X" and "wlsEnv rho"
and "fresh ys y X"
shows "(X #[rho [y \<leftarrow> Y]_ys]) = (X #[rho])"
using assms by(simp add: fresh_psubst_updEnv)
theorem wls_freshEnv_psubst_ident[simp]:
assumes "wls s X" and "wlsEnv rho"
and "\<And> zs z. freshEnv zs z rho \<or> fresh zs z X"
shows "(X #[rho]) = X"
using assms by simp
theorem wls_fresh_subst_ident[simp]:
assumes "wls (asSort ys) Y" and "wls s X" and "fresh ys y X"
shows "(X #[Y / y]_ys) = X"
using assms by(simp add: fresh_subst_ident)
theorem wls_substEnv_updEnv_fresh:
assumes "wls (asSort xs) X" and "wls (asSort ys) Y" and "fresh ys y X"
shows "((rho [x \<leftarrow> X]_xs) &[Y / y]_ys) = ((rho &[Y / y]_ys) [x \<leftarrow> X]_xs)"
using assms by(simp add: substEnv_updEnv_fresh)
theorem wls_fresh_substEnv_updEnv[simp]:
assumes "wlsEnv rho" and "wls (asSort ys) Y"
and "freshEnv ys y rho"
shows "(rho &[Y / y]_ys) = (rho [y \<leftarrow> Y]_ys)"
using assms by simp
theorem wls_fresh_vsubst_ident[simp]:
assumes "wls s X" and "fresh ys y X"
shows "(X #[y1 // y]_ys) = X"
using assms by(simp add: fresh_vsubst_ident)
theorem wls_vsubstEnv_updEnv_fresh:
assumes "wls s X" and "fresh ys y X"
shows "((rho [x \<leftarrow> X]_xs) &[y1 // y]_ys) = ((rho &[y1 // y]_ys) [x \<leftarrow> X]_xs)"
using assms by(simp add: vsubstEnv_updEnv_fresh)
theorem wls_fresh_vsubstEnv_updEnv[simp]:
assumes "wlsEnv rho"
and "freshEnv ys y rho"
shows "(rho &[y1 // y]_ys) = (rho [y \<leftarrow> Var ys y1]_ys)"
using assms by simp
theorem wls_swap_psubst:
assumes "wls s X" and "wlsEnv rho"
shows "((X #[rho]) #[z1 \<and> z2]_zs) = ((X #[z1 \<and> z2]_zs) #[rho &[z1 \<and> z2]_zs])"
using assms by(simp add: swap_psubst)
theorem wls_swap_subst:
assumes "wls s X" and "wls (asSort ys) Y"
shows "((X #[Y / y]_ys) #[z1 \<and> z2]_zs) = ((X #[z1 \<and> z2]_zs) #[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swap_subst)
theorem wls_swap_vsubst:
assumes "wls s X"
shows "((X #[y1 // y]_ys) #[z1 \<and> z2]_zs) = ((X #[z1 \<and> z2]_zs) #[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swap_vsubst)
theorem wls_swapEnv_psubstEnv:
assumes "wlsEnv rho" and "wlsEnv rho'"
shows "((rho &[rho']) &[z1 \<and> z2]_zs) = ((rho &[z1 \<and> z2]_zs) &[rho' &[z1 \<and> z2]_zs])"
using assms by(simp add: swapEnv_psubstEnv)
theorem wls_swapEnv_substEnv:
assumes "wls (asSort ys) Y" and "wlsEnv rho"
shows "((rho &[Y / y]_ys) &[z1 \<and> z2]_zs) =
((rho &[z1 \<and> z2]_zs) &[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swapEnv_substEnv)
theorem wls_swapEnv_vsubstEnv:
assumes "wlsEnv rho"
shows "((rho &[y1 // y]_ys) &[z1 \<and> z2]_zs) =
((rho &[z1 \<and> z2]_zs) &[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swapEnv_vsubstEnv)
theorem wls_psubst_compose:
assumes "wls s X" and "wlsEnv rho" and "wlsEnv rho'"
shows "((X #[rho]) #[rho']) = (X #[(rho &[rho'])])"
using assms by(simp add: psubst_compose)
theorem wls_psubstEnv_compose:
assumes "wlsEnv rho" and "wlsEnv rho'" and "wlsEnv rho''"
shows "((rho &[rho']) &[rho'']) = (rho &[(rho' &[rho''])])"
using assms by(simp add: psubstEnv_compose)
theorem wls_psubst_subst_compose:
assumes "wls s X" and "wls (asSort ys) Y" and "wlsEnv rho"
shows "((X #[Y / y]_ys) #[rho]) = (X #[(rho [y \<leftarrow> (Y #[rho])]_ys)])"
using assms by(simp add: psubst_subst_compose)
theorem wls_psubst_subst_compose_freshEnv:
assumes "wlsEnv rho" and "wls s X" and "wls (asSort ys) Y"
and "freshEnv ys y rho"
shows "((X #[Y / y]_ys) #[rho]) = ((X #[rho]) #[(Y #[rho]) / y]_ys)"
using assms by (simp add: psubst_subst_compose_freshEnv)
theorem wls_psubstEnv_substEnv_compose_freshEnv:
assumes "wlsEnv rho" and "wlsEnv rho'" and "wls (asSort ys) Y"
assumes "freshEnv ys y rho'"
shows "((rho &[Y / y]_ys) &[rho']) = ((rho &[rho']) &[(Y #[rho']) / y]_ys)"
using assms by (simp add: psubstEnv_substEnv_compose_freshEnv)
theorem wls_psubstEnv_substEnv_compose:
assumes "wlsEnv rho" and "wls (asSort ys) Y" and "wlsEnv rho'"
shows "((rho &[Y / y]_ys) &[rho']) = (rho &[(rho' [y \<leftarrow> (Y #[rho'])]_ys)])"
using assms by(simp add: psubstEnv_substEnv_compose)
theorem wls_psubst_vsubst_compose:
assumes "wls s X" and "wlsEnv rho"
shows "((X #[y1 // y]_ys) #[rho]) = (X #[(rho [y \<leftarrow> ((Var ys y1) #[rho])]_ys)])"
using assms by(simp add: psubst_vsubst_compose)
theorem wls_psubstEnv_vsubstEnv_compose:
assumes "wlsEnv rho" and "wlsEnv rho'"
shows "((rho &[y1 // y]_ys) &[rho']) = (rho &[(rho' [y \<leftarrow> ((Var ys y1) #[rho'])]_ys)])"
using assms by(simp add: psubstEnv_vsubstEnv_compose)
theorem wls_subst_psubst_compose:
assumes "wls s X" and "wls (asSort ys) Y" and "wlsEnv rho"
shows "((X #[rho]) #[Y / y]_ys) = (X #[(rho &[Y / y]_ys)])"
using assms by(simp add: subst_psubst_compose)
theorem wls_substEnv_psubstEnv_compose:
assumes "wlsEnv rho" and "wls (asSort ys) Y" and "wlsEnv rho'"
shows "((rho &[rho']) &[Y / y]_ys) = (rho &[(rho' &[Y / y]_ys)])"
using assms by(simp add: substEnv_psubstEnv_compose)
theorem wls_vsubst_psubst_compose:
assumes "wls s X" and "wlsEnv rho"
shows "((X #[rho]) #[y1 // y]_ys) = (X #[(rho &[y1 // y]_ys)])"
using assms by(simp add: vsubst_psubst_compose)
theorem wls_vsubstEnv_psubstEnv_compose:
assumes "wlsEnv rho" and "wlsEnv rho'"
shows "((rho &[rho']) &[y1 // y]_ys) = (rho &[(rho' &[y1 // y]_ys)])"
using assms by(simp add: vsubstEnv_psubstEnv_compose)
theorem wls_subst_compose1:
assumes "wls s X" and "wls (asSort ys) Y1" and "wls (asSort ys) Y2"
shows "((X #[Y1 / y]_ys) #[Y2 / y]_ys) = (X #[(Y1 #[Y2 / y]_ys) / y]_ys)"
using assms by(simp add: subst_compose1)
theorem wls_substEnv_compose1:
assumes "wlsEnv rho" and "wls (asSort ys) Y1" and "wls (asSort ys) Y2"
shows "((rho &[Y1 / y]_ys) &[Y2 / y]_ys) = (rho &[(Y1 #[Y2 / y]_ys) / y]_ys)"
using assms by(simp add: substEnv_compose1)
theorem wls_subst_vsubst_compose1:
assumes "wls s X" and "wls (asSort ys) Y" and "y \<noteq> y1"
shows "((X #[y1 // y]_ys) #[Y / y]_ys) = (X #[y1 // y]_ys)"
using assms by(simp add: subst_vsubst_compose1)
theorem wls_substEnv_vsubstEnv_compose1:
assumes "wlsEnv rho" and "wls (asSort ys) Y" and "y \<noteq> y1"
shows "((rho &[y1 // y]_ys) &[Y / y]_ys) = (rho &[y1 // y]_ys)"
using assms by(simp add: substEnv_vsubstEnv_compose1)
theorem wls_vsubst_subst_compose1:
assumes "wls s X" and "wls (asSort ys) Y"
shows "((X #[Y / y]_ys) #[y1 // y]_ys) = (X #[(Y #[y1 // y]_ys) / y]_ys)"
using assms by(simp add: vsubst_subst_compose1)
theorem wls_vsubstEnv_substEnv_compose1:
assumes "wlsEnv rho" and "wls (asSort ys) Y"
shows "((rho &[Y / y]_ys) &[y1 // y]_ys) = (rho &[(Y #[y1 // y]_ys) / y]_ys)"
using assms by(simp add: vsubstEnv_substEnv_compose1)
theorem wls_vsubst_compose1:
assumes "wls s X"
shows "((X #[y1 // y]_ys) #[y2 // y]_ys) = (X #[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms by(simp add: vsubst_compose1)
theorem wls_vsubstEnv_compose1:
assumes "wlsEnv rho"
shows "((rho &[y1 // y]_ys) &[y2 // y]_ys) = (rho &[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms by(simp add: vsubstEnv_compose1)
theorem wls_subst_compose2:
assumes "wls s X" and "wls (asSort ys) Y" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((X #[Y / y]_ys) #[Z / z]_zs) = ((X #[Z / z]_zs) #[(Y #[Z / z]_zs) / y]_ys)"
using assms by(simp add: subst_compose2)
theorem wls_substEnv_compose2:
assumes "wlsEnv rho" and "wls (asSort ys) Y" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((rho &[Y / y]_ys) &[Z / z]_zs) = ((rho &[Z / z]_zs) &[(Y #[Z / z]_zs) / y]_ys)"
using assms by(simp add: substEnv_compose2)
theorem wls_subst_vsubst_compose2:
assumes "wls s X" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((X #[y1 // y]_ys) #[Z / z]_zs) = ((X #[Z / z]_zs) #[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms by(simp add: subst_vsubst_compose2)
theorem wls_substEnv_vsubstEnv_compose2:
assumes "wlsEnv rho" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((rho &[y1 // y]_ys) &[Z / z]_zs) = ((rho &[Z / z]_zs) &[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms by(simp add: substEnv_vsubstEnv_compose2)
theorem wls_vsubst_subst_compose2:
assumes "wls s X" and "wls (asSort ys) Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((X #[Y / y]_ys) #[z1 // z]_zs) = ((X #[z1 // z]_zs) #[(Y #[z1 // z]_zs) / y]_ys)"
using assms by(simp add: vsubst_subst_compose2)
theorem wls_vsubstEnv_substEnv_compose2:
assumes "wlsEnv rho" and "wls (asSort ys) Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((rho &[Y / y]_ys) &[z1 // z]_zs) = ((rho &[z1 // z]_zs) &[(Y #[z1 // z]_zs) / y]_ys)"
using assms by(simp add: vsubstEnv_substEnv_compose2)
theorem wls_vsubst_compose2:
assumes "wls s X"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((X #[y1 // y]_ys) #[z1 // z]_zs) = ((X #[z1 // z]_zs) #[(y1 @ys[z1 / z]_zs) // y]_ys)"
using assms by(simp add: vsubst_compose2)
theorem wls_vsubstEnv_compose2:
assumes "wlsEnv rho"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((rho &[y1 // y]_ys) &[z1 // z]_zs) =
((rho &[z1 // z]_zs) &[(y1 @ys[z1 / z]_zs) // y]_ys)"
using assms by(simp add: vsubstEnv_compose2)
-subsubsection {* Properties specific to variable-for-variable substitution *}
+subsubsection \<open>Properties specific to variable-for-variable substitution\<close>
theorem wls_vsubst_ident[simp]:
assumes "wls s X"
shows "(X #[z // z]_zs) = X"
using assms by(simp add: vsubst_ident)
theorem wls_subst_ident[simp]:
assumes "wls s X"
shows "(X #[(Var zs z) / z]_zs) = X"
using assms by simp
theorem wls_vsubst_eq_swap:
assumes "wls s X" and "y1 = y2 \<or> fresh ys y1 X"
shows "(X #[y1 // y2]_ys) = (X #[y1 \<and> y2]_ys)"
using assms by(simp add: vsubst_eq_swap)
theorem wls_skel_vsubst:
assumes "wls s X"
shows "skel (X #[y1 // y2]_ys) = skel X"
using assms by(simp add: skel_vsubst)
theorem wls_subst_vsubst_trans:
assumes "wls s X" and "wls (asSort ys) Y" and "fresh ys y1 X"
shows "((X #[y1 // y]_ys) #[Y / y1]_ys) = (X #[Y / y]_ys)"
using assms by (simp add: subst_vsubst_trans)
theorem wls_vsubst_trans:
assumes "wls s X" and "fresh ys y1 X"
shows "((X #[y1 // y]_ys) #[y2 // y1]_ys) = (X #[y2 // y]_ys)"
using assms by (simp add: vsubst_trans)
theorem wls_vsubst_commute:
assumes "wls s X"
and "xs \<noteq> xs' \<or> {x,y} \<inter> {x',y'} = {}" and "fresh xs x X" and "fresh xs' x' X"
shows "((X #[x // y]_xs) #[x' // y']_xs') = ((X #[x' // y']_xs') #[x // y]_xs)"
using assms by(simp add: vsubst_commute)
(* The next lemmas do not have ``good" counterparts: *)
theorem wls_induct[case_names Var Op Abs]:
assumes
Var: "\<And> xs x. phi (asSort xs) (Var xs x)" and
Op:
"\<And> delta inp binp.
\<lbrakk>wlsInp delta inp; wlsBinp delta binp;
liftAll2 phi (arOf delta) inp; liftAll2 phiAbs (barOf delta) binp\<rbrakk>
\<Longrightarrow> phi (stOf delta) (Op delta inp binp)" and
Abs:
"\<And> s xs x X.
\<lbrakk>isInBar (xs,s); wls s X;
\<And> Y. (X,Y) \<in> swapped \<Longrightarrow> phi s Y;
\<And> ys y1 y2. phi s (X #[y1 // y2]_ys);
\<And> Y. \<lbrakk>wls s Y; skel Y = skel X\<rbrakk> \<Longrightarrow> phi s Y\<rbrakk>
\<Longrightarrow> phiAbs (xs,s) (Abs xs x X)"
shows
"(wls s X \<longrightarrow> phi s X) \<and>
(wlsAbs (xs,s') A \<longrightarrow> phiAbs (xs,s') A)"
apply(induction rule: wls_templateInduct
[of "\<lambda>s. swapped \<union> {(X, X #[y1 // y2]_ys)| X ys y1 y2. True}
\<union> {(X,Y). wls s Y \<and> skel Y = skel X}"])
by (auto simp add: assms swapped_preserves_wls swapped_skel wls_skel_vsubst
intro!: Abs)
theorem wls_Abs_vsubst_all_aux:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X') \<longrightarrow>
(X #[y // x]_xs) = (X' #[y // x']_xs))"
using assms wls_Abs_swap_all by (simp add: wls_vsubst_eq_swap)
theorem wls_Abs_vsubst_ex:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<exists> y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X' \<and>
(X #[y // x]_xs) = (X' #[y // x']_xs))"
proof-
let ?phi = "\<lambda> f y. y \<notin> {x,x'} \<and> fresh xs y X \<and> fresh xs y X'
\<and> (f xs y x X) = (f xs y x' X')"
{assume "Abs xs x X = Abs xs x' X'"
then obtain y where "?phi swap y" using assms wls_Abs_swap_ex by auto
hence "?phi (\<lambda> xs y x X. (X #[y // x]_xs)) y"
using assms by(simp add: wls_vsubst_eq_swap)
hence "\<exists> y. ?phi (\<lambda> xs y x X. (X #[y // x]_xs)) y" by auto
}
moreover
{fix y assume "?phi (\<lambda> xs y x X. (X #[y // x]_xs)) y"
hence "?phi swap y" using assms by(auto simp add: wls_vsubst_eq_swap)
hence "Abs xs x X = Abs xs x' X'" using assms wls_Abs_swap_ex by auto
}
ultimately show ?thesis by auto
qed
theorem wls_Abs_vsubst_all:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<forall> y. (X #[y // x]_xs) = (X' #[y // x']_xs))"
proof(rule iffI, clarify)
assume "\<forall> y. (X #[y // x]_xs) = (X' #[y // x']_xs)"
thus "Abs xs x X = Abs xs x' X'"
using assms by(auto simp add: wls_Abs_vsubst_all_aux)
next
fix y
assume "Abs xs x X = Abs xs x' X'"
then obtain z where z_fresh: "fresh xs z X \<and> fresh xs z X'"
and "(X #[z // x]_xs) = (X' #[z // x']_xs)"
using assms by(auto simp add: wls_Abs_vsubst_ex)
hence "((X #[z // x]_xs) #[y // z]_xs) = ((X' #[z // x']_xs) #[y // z]_xs)" by simp
thus "(X #[y // x]_xs) = (X' #[y // x']_xs)"
using assms z_fresh wls_vsubst_trans by auto
qed
theorem wls_Abs_subst_all:
assumes "wls s X" and "wls s X'"
shows
"(Abs xs x X = Abs xs x' X') =
(\<forall> Y. wls (asSort xs) Y \<longrightarrow> (X #[Y / x]_xs) = (X' #[Y / x']_xs))"
proof(rule iffI, clarify)
assume "\<forall> Y. wls (asSort xs) Y \<longrightarrow> (X #[Y / x]_xs) = (X' #[Y / x']_xs)"
hence "\<forall> y. (X #[y // x]_xs) = (X' #[y // x']_xs)"
unfolding vsubst_def by simp
thus "Abs xs x X = Abs xs x' X'"
using assms wls_Abs_vsubst_all by auto
next
fix Y assume Y: "wls (asSort xs) Y"
assume "Abs xs x X = Abs xs x' X'"
then obtain z where z_fresh: "fresh xs z X \<and> fresh xs z X'"
and "(X #[z // x]_xs) = (X' #[z // x']_xs)"
using assms by(auto simp add: wls_Abs_vsubst_ex)
hence "((X #[z // x]_xs) #[Y / z]_xs) = ((X' #[z // x']_xs) #[Y / z]_xs)" by simp
thus "(X #[Y / x]_xs) = (X' #[Y / x']_xs)"
using assms z_fresh Y wls_subst_vsubst_trans by auto
qed
lemma Abs_inj_fresh[simp]:
assumes X: "wls s X" and X': "wls s X'"
and fresh_X: "fresh ys x X" and fresh_X': "fresh ys x' X'"
and eq: "Abs ys x X = Abs ys x' X'"
shows "X = X'"
proof-
obtain z where "(X #[z // x]_ys) = (X' #[z // x']_ys)"
using X X' eq by(auto simp add: wls_Abs_vsubst_ex)
thus ?thesis using X X' fresh_X fresh_X' by simp
qed
theorem wls_Abs_vsubst_cong:
assumes "wls s X" and "wls s X'"
and "fresh xs y X" and "fresh xs y X'" and "(X #[y // x]_xs) = (X' #[y // x']_xs)"
shows "Abs xs x X = Abs xs x' X'"
using assms by (intro wls_Abs_swap_cong) (auto simp: wls_vsubst_eq_swap)
theorem wls_Abs_vsubst_fresh[simp]:
assumes "wls s X" and "fresh xs x' X"
shows "Abs xs x' (X #[x' // x]_xs) = Abs xs x X"
using assms by(simp add: wls_vsubst_eq_swap)
theorem wls_Abs_subst_Var_fresh[simp]:
assumes "wls s X" and "fresh xs x' X"
shows "Abs xs x' (subst xs (Var xs x') x X) = Abs xs x X"
using assms wls_Abs_vsubst_fresh unfolding vsubst_def by simp
theorem wls_Abs_vsubst_congSTR:
assumes "wls s X" and "wls s X'"
and "y = x \<or> fresh xs y X" "y = x' \<or> fresh xs y X'"
and "(X #[y // x]_xs) = (X' #[y // x']_xs)"
shows "Abs xs x X = Abs xs x' X'"
by (metis assms wls_Abs_vsubst_fresh wls_vsubst_ident)
-subsubsection {* Abstraction versions of the properties *}
+subsubsection \<open>Abstraction versions of the properties\<close>
(* Environment identity and update versus other operators *)
theorem wls_psubstAbs_idEnv[simp]:
"wlsAbs (us,s) A \<Longrightarrow> (A $[idEnv]) = A"
by simp
(* Substitution versus other operators *)
theorem wls_freshAbs_psubstAbs:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
shows
"freshAbs zs z (A $[rho]) =
(\<forall> ys y. freshAbs ys y A \<or> freshImEnvAt zs z rho ys y)"
using assms by(simp add: freshAbs_psubstAbs)
theorem wls_freshAbs_psubstAbs_E1:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
and "rho ys y = None" and "freshAbs zs z (A $[rho])"
shows "freshAbs ys y A \<or> (ys \<noteq> zs \<or> y \<noteq> z)"
using assms freshAbs_psubstAbs_E1[of A rho ys y zs z] by simp
theorem wls_freshAbs_psubstAbs_E2:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
and "rho ys y = Some Y" and "freshAbs zs z (A $[rho])"
shows "freshAbs ys y A \<or> fresh zs z Y"
using assms freshAbs_psubstAbs_E2[of A rho ys y Y zs z] by simp
theorem wls_freshAbs_psubstAbs_I1:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
and "freshAbs zs z A" and "freshEnv zs z rho"
shows "freshAbs zs z (A $[rho])"
using assms by(simp add: freshAbs_psubstAbs_I1)
theorem wls_freshAbs_psubstAbs_I:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
and "rho zs z = None \<Longrightarrow> freshAbs zs z A" and
"\<And> ys y Y. rho ys y = Some Y \<Longrightarrow> freshAbs ys y A \<or> fresh zs z Y"
shows "freshAbs zs z (A $[rho])"
using assms by(simp add: freshAbs_psubstAbs_I)
theorem wls_freshAbs_substAbs:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
shows "freshAbs zs z (A $[Y / y]_ys) =
(((zs = ys \<and> z = y) \<or> freshAbs zs z A) \<and> (freshAbs ys y A \<or> fresh zs z Y))"
using assms by(simp add: freshAbs_substAbs)
theorem wls_freshAbs_vsubstAbs:
assumes "wlsAbs (us,s) A"
shows "freshAbs zs z (A $[y1 // y]_ys) =
(((zs = ys \<and> z = y) \<or> freshAbs zs z A) \<and>
(freshAbs ys y A \<or> (zs \<noteq> ys \<or> z \<noteq> y1)))"
using assms by(simp add: freshAbs_vsubstAbs)
theorem wls_substAbs_preserves_freshAbs:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
and "freshAbs zs z A" and "fresh zs z Y"
shows "freshAbs zs z (A $[Y / y]_ys)"
using assms by(simp add: substAbs_preserves_freshAbs)
theorem wls_vsubstAbs_preserves_freshAbs:
assumes "wlsAbs (us,s) A"
and "freshAbs zs z A" and "zs \<noteq> ys \<or> z \<noteq> y1"
shows "freshAbs zs z (A $[y1 // y]_ys)"
using assms by(simp add: vsubstAbs_preserves_freshAbs)
theorem wls_fresh_freshAbs_substAbs[simp]:
assumes "wls (asSort ys) Y" and "wlsAbs (us,s) A"
and "fresh ys y Y"
shows "freshAbs ys y (A $[Y / y]_ys)"
using assms by simp
theorem wls_diff_freshAbs_vsubstAbs[simp]:
assumes "wlsAbs (us,s) A"
and "y \<noteq> y1"
shows "freshAbs ys y (A $[y1 // y]_ys)"
using assms by simp
theorem wls_freshAbs_substAbs_E1:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
and "freshAbs zs z (A $[Y / y]_ys)" and "z \<noteq> y \<or> zs \<noteq> ys"
shows "freshAbs zs z A"
using assms freshAbs_substAbs_E1[of A Y zs z ys y] by auto
theorem wls_freshAbs_vsubstAbs_E1:
assumes "wlsAbs (us,s) A"
and "freshAbs zs z (A $[y1 // y]_ys)" and "z \<noteq> y \<or> zs \<noteq> ys"
shows "freshAbs zs z A"
using assms freshAbs_vsubstAbs_E1[of A zs z ys y1 y] by auto
theorem wls_freshAbs_substAbs_E2:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
and "freshAbs zs z (A $[Y / y]_ys)"
shows "freshAbs ys y A \<or> fresh zs z Y"
using assms freshAbs_substAbs_E2[of A Y zs z ys] by simp
theorem wls_freshAbs_vsubstAbs_E2:
assumes "wlsAbs (us,s) A"
and "freshAbs zs z (A $[y1 // y]_ys)"
shows "freshAbs ys y A \<or> zs \<noteq> ys \<or> z \<noteq> y1"
using assms freshAbs_vsubstAbs_E2[of A zs z ys y1 y] by simp
theorem wls_psubstAbs_cong[fundef_cong]:
assumes "wlsAbs (us,s) A" and "wlsEnv rho" and "wlsEnv rho'"
and "\<And> ys y. freshAbs ys y A \<or> rho ys y = rho' ys y"
shows "(A $[rho]) = (A $[rho'])"
using assms by(simp add: psubstAbs_cong)
theorem wls_freshAbs_psubstAbs_updEnv:
assumes "wls (asSort xs) X" and "wlsAbs (us,s) A" and "wlsEnv rho"
and "freshAbs xs x A"
shows "(A $[rho [x \<leftarrow> X]_xs]) = (A $[rho])"
using assms by(simp add: freshAbs_psubstAbs_updEnv)
lemma wls_freshEnv_psubstAbs_ident[simp]:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
and "\<And> zs z. freshEnv zs z rho \<or> freshAbs zs z A"
shows "(A $[rho]) = A"
using assms by simp
theorem wls_freshAbs_substAbs_ident[simp]:
assumes "wls (asSort xs) X" and "wlsAbs (us,s) A" and "freshAbs xs x A"
shows "(A $[X / x]_xs) = A"
using assms by simp
theorem wls_substAbs_Abs[simp]:
assumes "wls s X" and "wls (asSort xs) Y"
shows "((Abs xs x X) $[Y / x]_xs) = Abs xs x X"
using assms by simp
theorem wls_freshAbs_vsubstAbs_ident[simp]:
assumes "wlsAbs (us,s) A" and "freshAbs xs x A"
shows "(A $[x1 // x]_xs) = A"
using assms by(simp add: freshAbs_vsubstAbs_ident)
theorem wls_swapAbs_psubstAbs:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
shows "((A $[rho]) $[z1 \<and> z2]_zs) = ((A $[z1 \<and> z2]_zs) $[rho &[z1 \<and> z2]_zs])"
using assms by(simp add: swapAbs_psubstAbs)
theorem wls_swapAbs_substAbs:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
shows "((A $[Y / y]_ys) $[z1 \<and> z2]_zs) =
((A $[z1 \<and> z2]_zs) $[(Y #[z1 \<and> z2]_zs) / (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swapAbs_substAbs)
theorem wls_swapAbs_vsubstAbs:
assumes "wlsAbs (us,s) A"
shows "((A $[y1 // y]_ys) $[z1 \<and> z2]_zs) =
((A $[z1 \<and> z2]_zs) $[(y1 @ys[z1 \<and> z2]_zs) // (y @ys[z1 \<and> z2]_zs)]_ys)"
using assms by(simp add: swapAbs_vsubstAbs)
theorem wls_psubstAbs_compose:
assumes "wlsAbs (us,s) A" and "wlsEnv rho" and "wlsEnv rho'"
shows "((A $[rho]) $[rho']) = (A $[(rho &[rho'])])"
using assms by(simp add: psubstAbs_compose)
theorem wls_psubstAbs_substAbs_compose:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y" and "wlsEnv rho"
shows "((A $[Y / y]_ys) $[rho]) = (A $[(rho [y \<leftarrow> (Y #[rho])]_ys)])"
using assms by(simp add: psubstAbs_substAbs_compose)
theorem wls_psubstAbs_substAbs_compose_freshEnv:
assumes "wlsEnv rho" and "wlsAbs (us,s) A" and "wls (asSort ys) Y"
assumes "freshEnv ys y rho"
shows "((A $[Y / y]_ys) $[rho]) = ((A $[rho]) $[(Y #[rho]) / y]_ys)"
using assms by (simp add: psubstAbs_substAbs_compose_freshEnv)
theorem wls_psubstAbs_vsubstAbs_compose:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
shows "((A $[y1 // y]_ys) $[rho]) = (A $[(rho [y \<leftarrow> ((Var ys y1) #[rho])]_ys)])"
using assms by(simp add: psubstAbs_vsubstAbs_compose)
theorem wls_substAbs_psubstAbs_compose:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y" and "wlsEnv rho"
shows "((A $[rho]) $[Y / y]_ys) = (A $[(rho &[Y / y]_ys)])"
using assms by(simp add: substAbs_psubstAbs_compose)
theorem wls_vsubstAbs_psubstAbs_compose:
assumes "wlsAbs (us,s) A" and "wlsEnv rho"
shows "((A $[rho]) $[y1 // y]_ys) = (A $[(rho &[y1 // y]_ys)])"
using assms by(simp add: vsubstAbs_psubstAbs_compose)
theorem wls_substAbs_compose1:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y1" and "wls (asSort ys) Y2"
shows "((A $[Y1 / y]_ys) $[Y2 / y]_ys) = (A $[(Y1 #[Y2 / y]_ys) / y]_ys)"
using assms by(simp add: substAbs_compose1)
theorem wls_substAbs_vsubstAbs_compose1:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y" and "y \<noteq> y1"
shows "((A $[y1 // y]_ys) $[Y / y]_ys) = (A $[y1 // y]_ys)"
using assms by(simp add: substAbs_vsubstAbs_compose1)
theorem wls_vsubstAbs_substAbs_compose1:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
shows "((A $[Y / y]_ys) $[y1 // y]_ys) = (A $[(Y #[y1 // y]_ys) / y]_ys)"
using assms by(simp add: vsubstAbs_substAbs_compose1)
theorem wls_vsubstAbs_compose1:
assumes "wlsAbs (us,s) A"
shows "((A $[y1 // y]_ys) $[y2 // y]_ys) = (A $[(y1 @ys[y2 / y]_ys) // y]_ys)"
using assms by(simp add: vsubstAbs_compose1)
theorem wls_substAbs_compose2:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((A $[Y / y]_ys) $[Z / z]_zs) = ((A $[Z / z]_zs) $[(Y #[Z / z]_zs) / y]_ys)"
using assms by(simp add: substAbs_compose2)
theorem wls_substAbs_vsubstAbs_compose2:
assumes "wlsAbs (us,s) A" and "wls (asSort zs) Z"
and "ys \<noteq> zs \<or> y \<noteq> z" and fresh: "fresh ys y Z"
shows "((A $[y1 // y]_ys) $[Z / z]_zs) = ((A $[Z / z]_zs) $[((Var ys y1) #[Z / z]_zs) / y]_ys)"
using assms by(simp add: substAbs_vsubstAbs_compose2)
theorem wls_vsubstAbs_substAbs_compose2:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((A $[Y / y]_ys) $[z1 // z]_zs) = ((A $[z1 // z]_zs) $[(Y #[z1 // z]_zs) / y]_ys)"
using assms by(simp add: vsubstAbs_substAbs_compose2)
theorem wls_vsubstAbs_compose2:
assumes "wlsAbs (us,s) A"
and "ys \<noteq> zs \<or> y \<notin> {z,z1}"
shows "((A $[y1 // y]_ys) $[z1 // z]_zs) = ((A $[z1 // z]_zs) $[(y1 @ys[z1 / z]_zs) // y]_ys)"
using assms by(simp add: vsubstAbs_compose2)
(* Properties specific to variable-for-variable substitution *)
theorem wls_vsubstAbs_ident[simp]:
assumes "wlsAbs (us,s) A"
shows "(A $[z // z]_zs) = A"
using assms by(simp add: vsubstAbs_ident)
theorem wls_substAbs_ident[simp]:
assumes "wlsAbs (us,s) A"
shows "(A $[(Var zs z) / z]_zs) = A"
using assms by simp
theorem wls_vsubstAbs_eq_swapAbs:
assumes "wlsAbs (us,s) A" and "y1 = y2 \<or> freshAbs ys y1 A"
shows "(A $[y1 // y2]_ys) = (A $[y1 \<and> y2]_ys)"
using assms vsubstAll_swapAll[of "Par [y1, y2] [] [] []" _ _ A]
unfolding goodPar_def by auto
theorem wls_skelAbs_vsubstAbs:
assumes "wlsAbs (us,s) A"
shows "skelAbs (A $[y1 // y2]_ys) = skelAbs A"
using assms by(simp add: skelAbs_vsubstAbs)
theorem wls_substAbs_vsubstAbs_trans:
assumes "wlsAbs (us,s) A" and "wls (asSort ys) Y" and "freshAbs ys y1 A"
shows "((A $[y1 // y]_ys) $[Y / y1]_ys) = (A $[Y / y]_ys)"
using assms by(simp add: substAbs_vsubstAbs_trans)
theorem wls_vsubstAbs_trans:
assumes "wlsAbs (us,s) A" and "freshAbs ys y1 A"
shows "((A $[y1 // y]_ys) $[y2 // y1]_ys) = (A $[y2 // y]_ys)"
using assms by(simp add: vsubstAbs_trans)
theorem wls_vsubstAbs_commute:
assumes "wlsAbs (us,s) A"
and "xs \<noteq> xs' \<or> {x,y} \<inter> {x',y'} = {}" and "freshAbs xs x A" and "freshAbs xs' x' A"
shows "((A $[x // y]_xs) $[x' // y']_xs') = ((A $[x' // y']_xs') $[x // y]_xs)"
proof-
have "freshAbs xs' x' (A $[x // y]_xs)"
using assms by(auto simp: vsubstAbs_preserves_freshAbs)
moreover have "freshAbs xs x (A $[x' // y']_xs')"
using assms by(auto simp: vsubstAbs_preserves_freshAbs)
ultimately show ?thesis using assms
by (auto simp: vsubstAbs_eq_swapAbs intro!: wls_swapAbs_commute)
qed
lemmas wls_psubstAll_freshAll_otherSimps =
wls_psubst_idEnv wls_psubstEnv_idEnv_id wls_psubstAbs_idEnv
wls_freshEnv_psubst_ident wls_freshEnv_psubstAbs_ident
lemmas wls_substAll_freshAll_otherSimps =
wls_fresh_fresh_subst wls_fresh_subst_ident wls_fresh_substEnv_updEnv wls_subst_ident
wls_fresh_freshAbs_substAbs wls_freshAbs_substAbs_ident wls_substAbs_ident
wls_Abs_subst_Var_fresh
lemmas wls_vsubstAll_freshAll_otherSimps =
wls_diff_fresh_vsubst wls_fresh_vsubst_ident wls_fresh_vsubstEnv_updEnv wls_vsubst_ident
wls_diff_freshAbs_vsubstAbs wls_freshAbs_vsubstAbs_ident wls_vsubstAbs_ident
wls_Abs_vsubst_fresh
lemmas wls_allOpers_otherSimps =
wls_swapAll_freshAll_otherSimps
wls_psubstAll_freshAll_otherSimps
wls_substAll_freshAll_otherSimps
wls_vsubstAll_freshAll_otherSimps
-subsection {* Operators for down-casting and case-analyzing well-sorted items *}
-
-text{* The features developed here may occasionally turn out more convenient than obtaining
+subsection \<open>Operators for down-casting and case-analyzing well-sorted items\<close>
+
+text\<open>The features developed here may occasionally turn out more convenient than obtaining
the desired effect by hand, via the corresponding nchotomies.
E.g., when we want to perform the case-analysis uniformly, as part of a
function definition, the operators defined in the subsection save some tedious
-definitions and proofs pertaining to Hilbert choice. *}
-
-subsubsection {* For terms *}
+definitions and proofs pertaining to Hilbert choice.\<close>
+
+subsubsection \<open>For terms\<close>
(* Definitions: *)
definition isVar where
"isVar s (X :: ('index,'bindex,'varSort,'var,'opSym)term) ==
\<exists> xs x. s = asSort xs \<and> X = Var xs x"
definition castVar where
"castVar s (X :: ('index,'bindex,'varSort,'var,'opSym)term) ==
SOME xs_x. s = asSort (fst xs_x) \<and> X = Var (fst xs_x) (snd xs_x)"
definition isOp where
"isOp s X \<equiv>
\<exists> delta inp binp.
wlsInp delta inp \<and> wlsBinp delta binp \<and> s = stOf delta \<and> X = Op delta inp binp"
definition castOp where
"castOp s X \<equiv>
SOME delta_inp_binp.
wlsInp (fst3 delta_inp_binp) (snd3 delta_inp_binp) \<and>
wlsBinp (fst3 delta_inp_binp) (trd3 delta_inp_binp) \<and>
s = stOf (fst3 delta_inp_binp) \<and>
X = Op (fst3 delta_inp_binp) (snd3 delta_inp_binp) (trd3 delta_inp_binp)"
definition sortTermCase where
"sortTermCase fVar fOp s X \<equiv>
if isVar s X then fVar (fst (castVar s X)) (snd (castVar s X))
else if isOp s X then fOp (fst3 (castOp s X)) (snd3 (castOp s X)) (trd3 (castOp s X))
else undefined"
(* Properties of isVar and castVar: *)
lemma isVar_asSort_Var[simp]:
"isVar (asSort xs) (Var xs x)"
unfolding isVar_def by auto
lemma not_isVar_Op[simp]:
"\<not> isVar s (Op delta inp binp)"
unfolding isVar_def by auto
lemma isVar_imp_wls:
"isVar s X \<Longrightarrow> wls s X"
unfolding isVar_def by auto
lemmas isVar_simps =
isVar_asSort_Var not_isVar_Op
lemma castVar_asSort_Var[simp]:
"castVar (asSort xs) (Var xs x) = (xs,x)"
unfolding castVar_def by (rule some_equality) auto
lemma isVar_castVar:
assumes "isVar s X"
shows "asSort (fst (castVar s X)) = s \<and>
Var (fst (castVar s X)) (snd (castVar s X)) = X"
using assms isVar_def by auto
lemma asSort_castVar[simp]:
"isVar s X \<Longrightarrow> asSort (fst (castVar s X)) = s"
using isVar_castVar by auto
lemma Var_castVar[simp]:
"isVar s X \<Longrightarrow> Var (fst (castVar s X)) (snd (castVar s X)) = X"
using isVar_castVar by auto
lemma castVar_inj[simp]:
assumes *: "isVar s X" and **: "isVar s' X'"
shows "(castVar s X = castVar s' X') = (s = s' \<and> X = X')"
using assms Var_castVar asSort_castVar by fastforce
lemmas castVar_simps =
castVar_asSort_Var
asSort_castVar Var_castVar castVar_inj
(* Properties of isOp and castOp: *)
lemma isOp_stOf_Op[simp]:
"\<lbrakk>wlsInp delta inp; wlsBinp delta binp\<rbrakk>
\<Longrightarrow> isOp (stOf delta) (Op delta inp binp)"
unfolding isOp_def by auto
lemma not_isOp_Var[simp]:
"\<not> isOp s (Var xs X)"
unfolding isOp_def by auto
lemma isOp_imp_wls:
"isOp s X \<Longrightarrow> wls s X"
unfolding isOp_def by auto
lemmas isOp_simps =
isOp_stOf_Op not_isOp_Var
lemma castOp_stOf_Op[simp]:
assumes "wlsInp delta inp" and "wlsBinp delta binp"
shows "castOp (stOf delta) (Op delta inp binp) = (delta,inp,binp)"
using assms unfolding castOp_def by (intro some_equality) auto
lemma isOp_castOp:
assumes "isOp s X"
shows "wlsInp (fst3 (castOp s X)) (snd3 (castOp s X)) \<and>
wlsBinp (fst3 (castOp s X)) (trd3 (castOp s X)) \<and>
stOf (fst3 (castOp s X)) = s \<and>
Op (fst3 (castOp s X)) (snd3 (castOp s X)) (trd3 (castOp s X)) = X"
proof-
let ?phi = "\<lambda> DIB. wlsInp (fst3 DIB) (snd3 DIB) \<and>
wlsBinp (fst3 DIB) (trd3 DIB) \<and>
s = stOf (fst3 DIB) \<and>
X = Op (fst3 DIB) (snd3 DIB) (trd3 DIB)"
obtain delta inp binp where "?phi (delta,inp,binp)"
using assms unfolding isOp_def by auto
hence "?phi (castOp s X)" using someI[of ?phi] by simp
thus ?thesis by simp
qed
lemma wlsInp_castOp[simp]:
"isOp s X \<Longrightarrow> wlsInp (fst3 (castOp s X)) (snd3 (castOp s X))"
using isOp_castOp by auto
lemma wlsBinp_castOp[simp]:
"isOp s X \<Longrightarrow> wlsBinp (fst3 (castOp s X)) (trd3 (castOp s X))"
using isOp_castOp by auto
lemma stOf_castOp[simp]:
"isOp s X \<Longrightarrow> stOf (fst3 (castOp s X)) = s"
using isOp_castOp by auto
lemma Op_castOp[simp]:
"isOp s X \<Longrightarrow>
Op (fst3 (castOp s X)) (snd3 (castOp s X)) (trd3 (castOp s X)) = X"
using isOp_castOp by auto
lemma castOp_inj[simp]:
assumes "isOp s X" and "isOp s' X'"
shows "(castOp s X = castOp s' X') = (s = s' \<and> X = X')"
using assms Op_castOp stOf_castOp by fastforce
lemmas castOp_simps =
castOp_stOf_Op wlsInp_castOp wlsBinp_castOp
stOf_castOp Op_castOp castOp_inj
(* isVar and castVar versus isOp and castOp: *)
lemma not_isVar_isOp:
"\<not> (isVar s X \<and> isOp s X)"
unfolding isVar_def isOp_def by auto
lemma isVar_or_isOp:
"wls s X \<Longrightarrow> isVar s X \<or> isOp s X"
by(erule wls_cases) auto
(* Properties of the case-analysis operator: *)
lemma sortTermCase_asSort_Var_simp[simp]:
"sortTermCase fVar fOp (asSort xs) (Var xs x) = fVar xs x"
unfolding sortTermCase_def by auto
lemma sortTermCase_stOf_Op_simp[simp]:
"\<lbrakk>wlsInp delta inp; wlsBinp delta binp\<rbrakk> \<Longrightarrow>
sortTermCase fVar fOp (stOf delta) (Op delta inp binp) = fOp delta inp binp"
unfolding sortTermCase_def by auto
lemma sortTermCase_cong[fundef_cong]:
assumes "\<And> xs x. fVar xs x = gVar xs x"
and "\<And> delta inp binp. \<lbrakk>wlsInp delta inp; wlsInp delta inp\<rbrakk>
\<Longrightarrow> fOp delta inp binp = gOp delta inp binp"
shows "wls s X \<Longrightarrow>
sortTermCase fVar fOp s X = sortTermCase gVar gOp s X"
apply(erule wls_cases) using assms by auto
lemmas sortTermCase_simps =
sortTermCase_asSort_Var_simp
sortTermCase_stOf_Op_simp
lemmas term_cast_simps =
isOp_simps castOp_simps sortTermCase_simps
-subsubsection {* For abstractions *}
-
-text {* Here, the situation will be different than that of terms, since:
+subsubsection \<open>For abstractions\<close>
+
+text \<open>Here, the situation will be different than that of terms, since:
\\- an abstraction can only be built using ``Abs", hence we need no ``is" operators;
\\- the constructor ``Abs" for abstractions is not injective, so need a more subtle condition
on the case-analysis operator.
Yet another difference is that when casting an abstraction ``A" such that ``wlsAbs (xs,s) A",
we need to cast only the value ``A", and not the sorting part``xs s", since the latter
already contains the desired information. Consequently, below, in the arguments for the case-analysis
-operator, the sorts ``xs s" come before the function ``f", and the latter doesnot take sorts into account. *}
+operator, the sorts ``xs s" come before the function ``f", and the latter doesnot take sorts into account.\<close>
(* Definitions: *)
definition castAbs where
"castAbs xs s A \<equiv> SOME x_X. wls s (snd x_X) \<and> A = Abs xs (fst x_X) (snd x_X)"
definition absCase where
"absCase xs s f A \<equiv> if wlsAbs (xs,s) A then f (fst (castAbs xs s A)) (snd (castAbs xs s A)) else undefined"
definition compatAbsSwap where
"compatAbsSwap xs s f \<equiv>
\<forall> x X x' X'. (\<forall> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X')
\<longrightarrow> (X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs))
\<longrightarrow> f x X = f x' X'"
definition compatAbsSubst where
"compatAbsSubst xs s f \<equiv>
\<forall> x X x' X'. (\<forall> Y. wls (asSort xs) Y \<longrightarrow> (X #[Y / x]_xs) = (X' #[Y / x']_xs))
\<longrightarrow> f x X = f x' X'"
definition compatAbsVsubst where
"compatAbsVsubst xs s f \<equiv>
\<forall> x X x' X'. (\<forall> y. (X #[y // x]_xs) = (X' #[y // x']_xs))
\<longrightarrow> f x X = f x' X'"
(* Properties of castAbs: *)
lemma wlsAbs_castAbs:
assumes "wlsAbs (xs,s) A"
shows "wls s (snd (castAbs xs s A)) \<and>
Abs xs (fst (castAbs xs s A)) (snd (castAbs xs s A)) = A"
proof-
let ?phi = "\<lambda> x_X. wls s (snd x_X) \<and>
A = Abs xs (fst x_X) (snd x_X)"
obtain x X where "?phi (x,X)" using assms wlsAbs_nchotomy[of xs s A] by auto
hence "?phi (castAbs xs s A)" unfolding castAbs_def using someI[of ?phi] by auto
thus ?thesis by simp
qed
lemma wls_castAbs[simp]:
"wlsAbs (xs,s) A \<Longrightarrow> wls s (snd (castAbs xs s A))"
using wlsAbs_castAbs by auto
lemma Abs_castAbs[simp]:
"wlsAbs (xs,s) A \<Longrightarrow> Abs xs (fst (castAbs xs s A)) (snd (castAbs xs s A)) = A"
using wlsAbs_castAbs by auto
lemma castAbs_Abs_swap:
assumes "isInBar (xs,s)" and X: "wls s X"
and yxX: "y = x \<or> fresh xs y X" and yx'X': "y = x' \<or> fresh xs y X'"
and *: "castAbs xs s (Abs xs x X) = (x',X')"
shows "(X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)"
proof-
have "wlsAbs (xs,s) (Abs xs x X)" using assms by simp
moreover
have "x' = fst (castAbs xs s (Abs xs x X))" and
"X' = snd (castAbs xs s (Abs xs x X))" using * by auto
ultimately
have "wls s X'" and "Abs xs x X = Abs xs x' X'" by auto
thus ?thesis using yxX yx'X' X by(auto simp add: wls_Abs_swap_all)
qed
lemma castAbs_Abs_subst:
assumes isInBar: "isInBar (xs,s)"
and X: "wls s X" and Y: "wls (asSort xs) Y"
and *: "castAbs xs s (Abs xs x X) = (x',X')"
shows "(X #[Y / x]_xs) = (X' #[Y / x']_xs)"
proof-
have "wlsAbs (xs,s) (Abs xs x X)" using isInBar X by simp
moreover
have "x' = fst (castAbs xs s (Abs xs x X))" and
"X' = snd (castAbs xs s (Abs xs x X))" using * by auto
ultimately
have "wls s X'" and "Abs xs x X = Abs xs x' X'" by auto
thus ?thesis using Y X by(auto simp add: wls_Abs_subst_all)
qed
lemma castAbs_Abs_vsubst:
assumes "isInBar (xs,s)" and "wls s X"
and "castAbs xs s (Abs xs x X) = (x',X')"
shows "(X #[y // x]_xs) = (X' #[y // x']_xs)"
using assms unfolding vsubst_def
by (intro castAbs_Abs_subst) auto
lemma castAbs_inj[simp]:
assumes *: "wlsAbs (xs,s) A" and **: "wlsAbs (xs,s) A'"
shows "(castAbs xs s A = castAbs xs s A') = (A = A')"
using assms Abs_castAbs by fastforce
lemmas castAbs_simps =
wls_castAbs Abs_castAbs castAbs_inj
(* Properties of the case-analysis operator: *)
lemma absCase_Abs_swap[simp]:
assumes isInBar: "isInBar (xs,s)" and X: "wls s X"
and f_compat: "compatAbsSwap xs s f"
shows "absCase xs s f (Abs xs x X) = f x X"
proof-
obtain x' X' where 1: "castAbs xs s (Abs xs x X) = (x',X')"
by (cases "castAbs xs s (Abs xs x X)", auto)
hence 2: "absCase xs s f (Abs xs x X) = f x' X'"
unfolding absCase_def using isInBar X by auto
have "\<And> y. (y = x \<or> fresh xs y X) \<and> (y = x' \<or> fresh xs y X')
\<Longrightarrow> (X #[y \<and> x]_xs) = (X' #[y \<and> x']_xs)"
using isInBar X 1 by(simp add: castAbs_Abs_swap)
hence "f x X = f x' X'" using f_compat
unfolding compatAbsSwap_def by fastforce
thus ?thesis using 2 by simp
qed
lemma absCase_Abs_subst[simp]:
assumes isInBar: "isInBar (xs,s)" and X: "wls s X"
and f_compat: "compatAbsSubst xs s f"
shows "absCase xs s f (Abs xs x X) = f x X"
proof-
obtain x' X' where 1: "castAbs xs s (Abs xs x X) = (x',X')"
by (cases "castAbs xs s (Abs xs x X)") auto
hence 2: "absCase xs s f (Abs xs x X) = f x' X'"
unfolding absCase_def using isInBar X by auto
have "\<And> Y. wls (asSort xs) Y \<Longrightarrow> (X #[Y / x]_xs) = (X' #[Y / x']_xs)"
using isInBar X 1 by(simp add: castAbs_Abs_subst)
hence "f x X = f x' X'" using f_compat unfolding compatAbsSubst_def by blast
thus ?thesis using 2 by simp
qed
lemma compatAbsVsubst_imp_compatAbsSubst[simp]:
"compatAbsVsubst xs s f \<Longrightarrow> compatAbsSubst xs s f"
unfolding compatAbsSubst_def compatAbsVsubst_def
vsubst_def by auto
lemma absCase_Abs_vsubst[simp]:
assumes "isInBar (xs,s)" and "wls s X"
and "compatAbsVsubst xs s f"
shows "absCase xs s f (Abs xs x X) = f x X"
using assms by(simp add: absCase_Abs_subst)
lemma absCase_cong[fundef_cong]:
assumes "compatAbsSwap xs s f \<or> compatAbsSubst xs s f \<or> compatAbsVsubst xs s f"
and "compatAbsSwap xs s f' \<or> compatAbsSubst xs s f' \<or> compatAbsVsubst xs s f'"
and "\<And> x X. wls s X \<Longrightarrow> f x X = f' x X"
shows "wlsAbs (xs,s) A \<Longrightarrow>
absCase xs s f A = absCase xs s f' A"
apply(erule wlsAbs_cases) using assms by auto
lemmas absCase_simps = absCase_Abs_swap absCase_Abs_subst
compatAbsVsubst_imp_compatAbsSubst absCase_Abs_vsubst
lemmas abs_cast_simps = castAbs_simps absCase_simps
lemmas cast_simps = term_cast_simps abs_cast_simps
lemmas wls_item_simps =
wlsAll_imp_goodAll paramS_simps Cons_wls_simps all_preserve_wls
wls_freeCons wls_allOpers_simps wls_allOpers_otherSimps Abs_inj_fresh cast_simps
(* Since the transition from good terms to well-sorted terms is complete, we
no longer need the ``good" layer: *)
lemmas wls_copy_of_good_item_simps = good_freeCons good_allOpers_simps good_allOpers_otherSimps
param_simps all_preserve_good
declare wls_copy_of_good_item_simps [simp del]
declare qItem_simps [simp del] declare qItem_versus_item_simps [simp del]
end (* context FixSyn *)
end
diff --git a/thys/Category3/CartesianCategory.thy b/thys/Category3/CartesianCategory.thy
--- a/thys/Category3/CartesianCategory.thy
+++ b/thys/Category3/CartesianCategory.thy
@@ -1,2285 +1,2285 @@
(* Title: CartesianCategory
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2020
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter "Cartesian Category"
text\<open>
In this chapter, we explore the notion of a ``cartesian category'', which we define
to be a category having binary products and a terminal object.
We show that every cartesian category extends to an ``elementary cartesian category'',
whose definition assumes that specific choices have been made for projections and
terminal object.
Conversely, the underlying category of an elementary cartesian category is a
cartesian category.
We also show that cartesian categories are the same thing as categories with
finite products.
\<close>
theory CartesianCategory
imports Limit SetCat
begin
section "Category with Binary Products"
subsection "Binary Product Diagrams"
text \<open>
The ``shape'' of a binary product diagram is a category having two distinct identity arrows
and no non-identity arrows.
\<close>
locale binary_product_shape
begin
sublocale concrete_category \<open>UNIV :: bool set\<close> \<open>\<lambda>a b. if a = b then {()} else {}\<close>
\<open>\<lambda>_. ()\<close> \<open>\<lambda>_ _ _ _ _. ()\<close>
apply (unfold_locales, auto)
apply (meson empty_iff)
by (meson empty_iff)
abbreviation comp
where "comp \<equiv> COMP"
abbreviation FF
where "FF \<equiv> MkIde False"
abbreviation TT
where "TT \<equiv> MkIde True"
lemma arr_char:
shows "arr f \<longleftrightarrow> f = FF \<or> f = TT"
using arr_char by (cases f, simp_all)
lemma ide_char:
shows "ide f \<longleftrightarrow> f = FF \<or> f = TT"
using ide_char ide_MkIde by (cases f, auto)
lemma is_discrete:
shows "ide f \<longleftrightarrow> arr f"
using arr_char ide_char by simp
lemma dom_simp [simp]:
assumes "arr f"
shows "dom f = f"
using assms is_discrete by simp
lemma cod_simp [simp]:
assumes "arr f"
shows "cod f = f"
using assms is_discrete by simp
lemma seq_char:
shows "seq f g \<longleftrightarrow> arr f \<and> f = g"
by auto
lemma comp_simp [simp]:
assumes "seq f g"
shows "comp f g = f"
using assms seq_char by fastforce
end
locale binary_product_diagram =
J: binary_product_shape +
C: category C
for C :: "'c comp" (infixr "\<cdot>" 55)
and a0 :: 'c
and a1 :: 'c +
assumes is_discrete: "C.ide a0 \<and> C.ide a1"
begin
notation J.comp (infixr "\<cdot>\<^sub>J" 55)
fun map
where "map J.FF = a0"
| "map J.TT = a1"
| "map _ = C.null"
sublocale diagram J.comp C map
proof
show "\<And>f. \<not> J.arr f \<Longrightarrow> map f = C.null"
using J.arr_char map.elims by auto
fix f
assume f: "J.arr f"
show "C.arr (map f)"
using f J.arr_char is_discrete C.ideD(1) map.simps(1-2) by metis
show "C.dom (map f) = map (J.dom f)"
using f J.arr_char J.dom_char is_discrete by force
show "C.cod (map f) = map (J.cod f)"
using f J.arr_char J.cod_char is_discrete by force
next
fix f g
assume fg: "J.seq g f"
show "map (g \<cdot>\<^sub>J f) = map g \<cdot> map f"
using fg J.arr_char J.seq_char J.null_char J.not_arr_null is_discrete
by (metis (no_types, lifting) C.comp_ide_self J.comp_simp map.simps(1-2))
qed
end
subsection "Category with Binary Products"
text \<open>
A \emph{binary product} in a category @{term C} is a limit of a binary product diagram
in @{term C}.
\<close>
context binary_product_diagram
begin
definition mkCone
where "mkCone p0 p1 \<equiv> \<lambda>j. if j = J.FF then p0 else if j = J.TT then p1 else C.null"
abbreviation is_rendered_commutative_by
where "is_rendered_commutative_by p0 p1 \<equiv>
C.seq a0 p0 \<and> C.seq a1 p1 \<and> C.dom p0 = C.dom p1"
abbreviation has_as_binary_product
where "has_as_binary_product p0 p1 \<equiv> limit_cone (C.dom p0) (mkCone p0 p1)"
lemma cone_mkCone:
assumes "is_rendered_commutative_by p0 p1"
shows "cone (C.dom p0) (mkCone p0 p1)"
proof -
interpret E: constant_functor J.comp C \<open>C.dom p0\<close>
using assms by unfold_locales auto
show "cone (C.dom p0) (mkCone p0 p1)"
using assms mkCone_def J.arr_char E.map_simp is_discrete C.comp_ide_arr C.comp_arr_dom
by unfold_locales auto
qed
lemma is_rendered_commutative_by_cone:
assumes "cone a \<chi>"
shows "is_rendered_commutative_by (\<chi> J.FF) (\<chi> J.TT)"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
show ?thesis
using is_discrete by simp
qed
lemma mkCone_cone:
assumes "cone a \<chi>"
shows "mkCone (\<chi> J.FF) (\<chi> J.TT) = \<chi>"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
interpret mkCone_\<chi>: cone J.comp C map \<open>C.dom (\<chi> J.FF)\<close> \<open>mkCone (\<chi> J.FF) (\<chi> J.TT)\<close>
using assms is_rendered_commutative_by_cone cone_mkCone by blast
show ?thesis
using mkCone_def \<chi>.is_extensional J.ide_char mkCone_def
NaturalTransformation.eqI [of J.comp C]
\<chi>.natural_transformation_axioms mkCone_\<chi>.natural_transformation_axioms
by fastforce
qed
end
locale binary_product_cone =
J: binary_product_shape +
C: category C +
D: binary_product_diagram C f0 f1 +
limit_cone J.comp C D.map \<open>C.dom p0\<close> \<open>D.mkCone p0 p1\<close>
for C :: "'c comp" (infixr "\<cdot>" 55)
and f0 :: 'c
and f1 :: 'c
and p0 :: 'c
and p1 :: 'c
begin
lemma renders_commutative:
shows "D.is_rendered_commutative_by p0 p1"
using cone_axioms D.is_rendered_commutative_by_cone D.mkCone_def \<Phi>.Ya.Cop_S.arr.simps(1)
by (metis (no_types, lifting)) (* TODO: pretty opaque *)
lemma is_universal':
assumes "D.is_rendered_commutative_by p0' p1'"
shows "\<exists>!h. \<guillemotleft>h : C.dom p0' \<rightarrow> C.dom p0\<guillemotright> \<and> p0 \<cdot> h = p0' \<and> p1 \<cdot> h = p1'"
proof -
have "D.cone (C.dom p0') (D.mkCone p0' p1')"
using assms D.cone_mkCone by blast
hence "\<exists>!h. \<guillemotleft>h : C.dom p0' \<rightarrow> C.dom p0\<guillemotright> \<and>
D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'"
using is_universal by simp
moreover have "\<And>h. \<guillemotleft>h : C.dom p0' \<rightarrow> C.dom p0\<guillemotright> \<Longrightarrow>
D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \<longleftrightarrow>
p0 \<cdot> h = p0' \<and> p1 \<cdot> h = p1'"
proof -
fix h
assume h: "\<guillemotleft>h : C.dom p0' \<rightarrow> C.dom p0\<guillemotright>"
show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1' \<longleftrightarrow>
p0 \<cdot> h = p0' \<and> p1 \<cdot> h = p1'"
proof
assume 1: "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'"
show "p0 \<cdot> h = p0' \<and> p1 \<cdot> h = p1'"
proof
show "p0 \<cdot> h = p0'"
proof -
have "p0' = D.mkCone p0' p1' J.FF"
using D.mkCone_def J.arr_char by simp
also have "... = D.cones_map h (D.mkCone p0 p1) J.FF"
using 1 by simp
also have "... = p0 \<cdot> h"
using h D.mkCone_def J.arr_char cone_\<chi> by auto
finally show ?thesis by auto
qed
show "p1 \<cdot> h = p1'"
proof -
have "p1' = D.mkCone p0' p1' J.TT"
using D.mkCone_def J.arr_char by simp
also have "... = D.cones_map h (D.mkCone p0 p1) J.TT"
using 1 by simp
also have "... = p1 \<cdot> h"
using h D.mkCone_def J.arr_char cone_\<chi> by auto
finally show ?thesis by auto
qed
qed
next
assume 1: "p0 \<cdot> h = p0' \<and> p1 \<cdot> h = p1'"
show "D.cones_map h (D.mkCone p0 p1) = D.mkCone p0' p1'"
using h 1 cone_\<chi> D.mkCone_def by auto
qed
qed
ultimately show ?thesis by blast
qed
lemma induced_arrowI':
assumes "D.is_rendered_commutative_by p0' p1'"
shows "\<guillemotleft>induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \<rightarrow> C.dom p0\<guillemotright>"
and "p0 \<cdot> induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'"
and "p1 \<cdot> induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'"
proof -
interpret A': constant_functor J.comp C \<open>C.dom p0'\<close>
using assms by (unfold_locales, auto)
have cone: "D.cone (C.dom p0') (D.mkCone p0' p1')"
using assms D.cone_mkCone [of p0' p1'] by blast
show 0: "p0 \<cdot> induced_arrow (C.dom p0') (D.mkCone p0' p1') = p0'"
proof -
have "p0 \<cdot> induced_arrow (C.dom p0') (D.mkCone p0' p1') =
D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1'))
(D.mkCone p0 p1) J.FF"
using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\<chi> by force
also have "... = p0'"
proof -
have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1'))
(D.mkCone p0 p1) =
D.mkCone p0' p1'"
using cone induced_arrowI by blast
thus ?thesis
using J.arr_char D.mkCone_def by simp
qed
finally show ?thesis by auto
qed
show "p1 \<cdot> induced_arrow (C.dom p1') (D.mkCone p0' p1') = p1'"
proof -
have "p1 \<cdot> induced_arrow (C.dom p1') (D.mkCone p0' p1') =
D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1'))
(D.mkCone p0 p1) J.TT"
using assms cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\<chi> by fastforce
also have "... = p1'"
proof -
have "D.cones_map (induced_arrow (C.dom p0') (D.mkCone p0' p1'))
(D.mkCone p0 p1) =
D.mkCone p0' p1'"
using cone induced_arrowI by blast
thus ?thesis
using J.arr_char D.mkCone_def by simp
qed
finally show ?thesis by auto
qed
show "\<guillemotleft>induced_arrow (C.dom p0') (D.mkCone p0' p1') : C.dom p0' \<rightarrow> C.dom p0\<guillemotright>"
using 0 cone induced_arrowI by simp
qed
end
context category
begin
definition has_as_binary_product
where "has_as_binary_product a0 a1 p0 p1 \<equiv>
ide a0 \<and> ide a1 \<and> binary_product_diagram.has_as_binary_product C a0 a1 p0 p1"
definition has_binary_products
where "has_binary_products =
(\<forall>a0 a1. ide a0 \<and> ide a1 \<longrightarrow> (\<exists>p0 p1. has_as_binary_product a0 a1 p0 p1))"
end
locale category_with_binary_products =
category +
assumes has_binary_products: has_binary_products
subsection "Elementary Category with Binary Products"
text \<open>
An \emph{elementary category with binary products} is a category equipped with a specific
way of mapping each pair of objects \<open>a\<close> and \<open>b\<close> to a pair of arrows \<open>\<pp>\<^sub>1[a, b]\<close> and \<open>\<pp>\<^sub>0[a, b]\<close>
that comprise a universal span. It is useful to assume that the mappings that produce
\<open>\<pp>\<^sub>1[a, b]\<close> and \<open>\<pp>\<^sub>0[a, b]\<close> from \<open>a\<close> and \<open>b\<close> are extensional; that is, if either \<open>a\<close> or \<open>b\<close>
is not an identity, then \<open>\<pp>\<^sub>1[a, b]\<close> and \<open>\<pp>\<^sub>0[a, b]\<close> are \<open>null\<close>.
\<close>
locale elementary_category_with_binary_products =
category C
for C :: "'a comp" (infixr "\<cdot>" 55)
and pr0 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<pp>\<^sub>0[_, _]")
and pr1 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" ("\<pp>\<^sub>1[_, _]") +
assumes pr0_ext: "\<not> (ide a \<and> ide b) \<Longrightarrow> \<pp>\<^sub>0[a, b] = null"
and pr1_ext: "\<not> (ide a \<and> ide b) \<Longrightarrow> \<pp>\<^sub>1[a, b] = null"
and span_pr: "\<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> span \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]"
and cod_pr0: "\<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> cod \<pp>\<^sub>0[a, b] = b"
and cod_pr1: "\<lbrakk> ide a; ide b \<rbrakk> \<Longrightarrow> cod \<pp>\<^sub>1[a, b] = a"
and universal: "span f g \<Longrightarrow> \<exists>!l. \<pp>\<^sub>1[cod f, cod g] \<cdot> l = f \<and> \<pp>\<^sub>0[cod f, cod g] \<cdot> l = g"
begin
lemma pr0_in_hom':
assumes "ide a" and "ide b"
shows "\<guillemotleft>\<pp>\<^sub>0[a, b] : dom \<pp>\<^sub>0[a, b] \<rightarrow> b\<guillemotright>"
using assms span_pr cod_pr0 by auto
lemma pr1_in_hom':
assumes "ide a" and "ide b"
shows "\<guillemotleft>\<pp>\<^sub>1[a, b] : dom \<pp>\<^sub>0[a, b] \<rightarrow> a\<guillemotright>"
using assms span_pr cod_pr1 by auto
text \<open>
We introduce a notation for tupling, which denotes the arrow into a product that
is induced by a span.
\<close>
definition tuple ("\<langle>_, _\<rangle>")
where "\<langle>f, g\<rangle> \<equiv> if span f g then
THE l. \<pp>\<^sub>1[cod f, cod g] \<cdot> l = f \<and> \<pp>\<^sub>0[cod f, cod g] \<cdot> l = g
else null"
text \<open>
The following defines product of arrows (not just of objects). It will take a little
while before we can prove that it is functorial, but for right now it is nice to have
it as a notation for the apex of a product cone. We have to go through some slightly
unnatural contortions in the development here, though, to avoid having to introduce a
separate preliminary notation just for the product of objects.
\<close>
(* TODO: I want to use \<times> but it has already been commandeered for product types. *)
definition prod (infixr "\<otimes>" 51)
where "f \<otimes> g \<equiv> \<langle>f \<cdot> \<pp>\<^sub>1[dom f, dom g], g \<cdot> \<pp>\<^sub>0[dom f, dom g]\<rangle>"
lemma seq_pr_tuple:
assumes "span f g"
shows "seq \<pp>\<^sub>0[cod f, cod g] \<langle>f, g\<rangle>"
proof -
have "\<pp>\<^sub>0[cod f, cod g] \<cdot> \<langle>f, g\<rangle> = g"
unfolding tuple_def
using assms universal theI [of "\<lambda>l. \<pp>\<^sub>1[cod f, cod g] \<cdot> l = f \<and> \<pp>\<^sub>0[cod f, cod g] \<cdot> l = g"]
by simp meson
thus ?thesis
using assms by simp
qed
lemma tuple_pr_arr:
assumes "ide a" and "ide b" and "seq \<pp>\<^sub>0[a, b] h"
shows "\<langle>\<pp>\<^sub>1[a, b] \<cdot> h, \<pp>\<^sub>0[a, b] \<cdot> h\<rangle> = h"
unfolding tuple_def
using assms span_pr cod_pr0 cod_pr1 universal [of "\<pp>\<^sub>1[a, b] \<cdot> h" "\<pp>\<^sub>0[a, b] \<cdot> h"]
theI_unique [of "\<lambda>l. \<pp>\<^sub>1[a, b] \<cdot> l = \<pp>\<^sub>1[a, b] \<cdot> h \<and> \<pp>\<^sub>0[a, b] \<cdot> l = \<pp>\<^sub>0[a, b] \<cdot> h" h]
by auto
lemma pr_tuple [simp]:
assumes "span f g" and "cod f = a" and "cod g = b"
shows "\<pp>\<^sub>1[a, b] \<cdot> \<langle>f, g\<rangle> = f" and "\<pp>\<^sub>0[a, b] \<cdot> \<langle>f, g\<rangle> = g"
proof -
have 1: "\<pp>\<^sub>1[a, b] \<cdot> \<langle>f, g\<rangle> = f \<and> \<pp>\<^sub>0[a, b] \<cdot> \<langle>f, g\<rangle> = g"
unfolding tuple_def
using assms universal theI [of "\<lambda>l. \<pp>\<^sub>1[a, b] \<cdot> l = f \<and> \<pp>\<^sub>0[a, b] \<cdot> l = g"]
by simp meson
show "\<pp>\<^sub>1[a, b] \<cdot> \<langle>f, g\<rangle> = f" using 1 by simp
show "\<pp>\<^sub>0[a, b] \<cdot> \<langle>f, g\<rangle> = g" using 1 by simp
qed
lemma cod_tuple:
assumes "span f g"
shows "cod \<langle>f, g\<rangle> = cod f \<otimes> cod g"
proof -
have "cod f \<otimes> cod g = \<langle>\<pp>\<^sub>1[cod f, cod g], \<pp>\<^sub>0[cod f, cod g]\<rangle>"
unfolding prod_def
using assms comp_cod_arr span_pr cod_pr0 cod_pr1 by simp
also have "... = \<langle>\<pp>\<^sub>1[cod f, cod g] \<cdot> dom \<pp>\<^sub>0[cod f, cod g],
\<pp>\<^sub>0[cod f, cod g] \<cdot> dom \<pp>\<^sub>0[cod f, cod g]\<rangle>"
using assms span_pr comp_arr_dom by simp
also have "... = dom \<pp>\<^sub>0[cod f, cod g]"
using assms tuple_pr_arr span_pr by simp
also have "... = cod \<langle>f, g\<rangle>"
using assms seq_pr_tuple by blast
finally show ?thesis by simp
qed
lemma tuple_in_hom [intro]:
assumes "\<guillemotleft>f : a \<rightarrow> b\<guillemotright>" and "\<guillemotleft>g : a \<rightarrow> c\<guillemotright>"
shows "\<guillemotleft>\<langle>f, g\<rangle> : a \<rightarrow> b \<otimes> c\<guillemotright>"
using assms pr_tuple dom_comp cod_tuple
apply (elim in_homE, intro in_homI)
apply (metis seqE)
by metis+
lemma tuple_in_hom' [simp]:
assumes "arr f" and "dom f = a" and "cod f = b"
and "arr g" and "dom g = a" and "cod g = c"
shows "\<guillemotleft>\<langle>f, g\<rangle> : a \<rightarrow> b \<otimes> c\<guillemotright>"
using assms by auto
lemma tuple_ext:
assumes "\<not> span f g"
shows "\<langle>f, g\<rangle> = null"
unfolding tuple_def
by (simp add: assms)
lemma tuple_simps [simp]:
assumes "span f g"
shows "arr \<langle>f, g\<rangle>" and "dom \<langle>f, g\<rangle> = dom f" and "cod \<langle>f, g\<rangle> = cod f \<otimes> cod g"
proof -
show "arr \<langle>f, g\<rangle>"
using assms tuple_in_hom by blast
show "dom \<langle>f, g\<rangle> = dom f"
using assms tuple_in_hom
by (metis dom_comp pr_tuple(1))
show "cod \<langle>f, g\<rangle> = cod f \<otimes> cod g"
using assms cod_tuple by auto
qed
lemma tuple_pr [simp]:
assumes "ide a" and "ide b"
shows "\<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle> = a \<otimes> b"
proof -
have 1: "dom \<pp>\<^sub>0[a, b] = a \<otimes> b"
using assms seq_pr_tuple cod_tuple [of "\<pp>\<^sub>1[a, b]" "\<pp>\<^sub>0[a, b]"] span_pr
pr0_in_hom' pr1_in_hom'
by (metis cod_pr0 cod_pr1 seqE)
hence "\<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle> = \<langle>\<pp>\<^sub>1[a, b] \<cdot> (a \<otimes> b), \<pp>\<^sub>0[a, b] \<cdot> (a \<otimes> b)\<rangle>"
using assms pr0_in_hom' pr1_in_hom' comp_arr_dom span_pr by simp
thus ?thesis
using assms 1 tuple_pr_arr span_pr
by (metis comp_arr_dom)
qed
lemma pr_in_hom [intro, simp]:
assumes "ide a" and "ide b"
shows "\<guillemotleft>\<pp>\<^sub>0[a, b] : a \<otimes> b \<rightarrow> b\<guillemotright>" and "\<guillemotleft>\<pp>\<^sub>1[a, b] : a \<otimes> b \<rightarrow> a\<guillemotright>"
proof -
show 0: "\<guillemotleft>\<pp>\<^sub>0[a, b] : a \<otimes> b \<rightarrow> b\<guillemotright>"
using assms pr0_in_hom' seq_pr_tuple [of "\<pp>\<^sub>1[a, b]" "\<pp>\<^sub>0[a, b]"]
cod_tuple [of "\<pp>\<^sub>1[a, b]" "\<pp>\<^sub>0[a, b]"] span_pr cod_pr0 cod_pr1
by (intro in_homI, auto)
show "\<guillemotleft>\<pp>\<^sub>1[a, b] : a \<otimes> b \<rightarrow> a\<guillemotright>"
using assms 0 span_pr pr1_in_hom' by fastforce
qed
lemma pr_simps [simp]:
assumes "ide a" and "ide b"
shows "arr \<pp>\<^sub>0[a, b]" and "dom \<pp>\<^sub>0[a, b] = a \<otimes> b" and "cod \<pp>\<^sub>0[a, b] = b"
and "arr \<pp>\<^sub>1[a, b]" and "dom \<pp>\<^sub>1[a, b] = a \<otimes> b" and "cod \<pp>\<^sub>1[a, b] = a"
using assms pr_in_hom by blast+
lemma arr_pr0_iff [iff]:
shows "arr \<pp>\<^sub>0[a, b] \<longleftrightarrow> ide a \<and> ide b"
proof
show "ide a \<and> ide b \<Longrightarrow> arr \<pp>\<^sub>0[a, b]"
using pr_in_hom by auto
show "arr \<pp>\<^sub>0[a, b] \<Longrightarrow> ide a \<and> ide b"
using pr0_ext not_arr_null by metis
qed
lemma arr_pr1_iff [iff]:
shows "arr \<pp>\<^sub>1[a, b] \<longleftrightarrow> ide a \<and> ide b"
proof
show "ide a \<and> ide b \<Longrightarrow> arr \<pp>\<^sub>1[a, b]"
using pr_in_hom by auto
show "arr \<pp>\<^sub>1[a, b] \<Longrightarrow> ide a \<and> ide b"
using pr1_ext not_arr_null by metis
qed
lemma pr_joint_monic:
assumes "seq \<pp>\<^sub>0[a, b] h"
and "\<pp>\<^sub>0[a, b] \<cdot> h = \<pp>\<^sub>0[a, b] \<cdot> h'" and "\<pp>\<^sub>1[a, b] \<cdot> h = \<pp>\<^sub>1[a, b] \<cdot> h'"
shows "h = h'"
using assms
by (metis arr_pr0_iff seqE tuple_pr_arr)
lemma comp_tuple_arr [simp]:
assumes "span f g" and "arr h" and "dom f = cod h"
shows "\<langle>f, g\<rangle> \<cdot> h = \<langle>f \<cdot> h, g \<cdot> h\<rangle>"
proof (intro pr_joint_monic [where h = "\<langle>f, g\<rangle> \<cdot> h"])
show "seq \<pp>\<^sub>0[cod f, cod g] (\<langle>f, g\<rangle> \<cdot> h)"
using assms by fastforce
show "\<pp>\<^sub>0[cod f, cod g] \<cdot> \<langle>f, g\<rangle> \<cdot> h = \<pp>\<^sub>0[cod f, cod g] \<cdot> \<langle>f \<cdot> h, g \<cdot> h\<rangle>"
proof -
have "\<pp>\<^sub>0[cod f, cod g] \<cdot> \<langle>f, g\<rangle> \<cdot> h = (\<pp>\<^sub>0[cod f, cod g] \<cdot> \<langle>f, g\<rangle>) \<cdot> h"
using comp_assoc by simp
thus ?thesis
using assms by simp
qed
show "\<pp>\<^sub>1[cod f, cod g] \<cdot> \<langle>f, g\<rangle> \<cdot> h = \<pp>\<^sub>1[cod f, cod g] \<cdot> \<langle>f \<cdot> h, g \<cdot> h\<rangle>"
proof -
have "\<pp>\<^sub>1[cod f, cod g] \<cdot> \<langle>f, g\<rangle> \<cdot> h = (\<pp>\<^sub>1[cod f, cod g] \<cdot> \<langle>f, g\<rangle>) \<cdot> h"
using comp_assoc by simp
thus ?thesis
using assms by simp
qed
qed
lemma ide_prod [intro, simp]:
assumes "ide a" and "ide b"
shows "ide (a \<otimes> b)"
using assms pr_simps ide_dom [of "\<pp>\<^sub>0[a, b]"] by simp
lemma prod_in_hom [intro]:
assumes "\<guillemotleft>f : a \<rightarrow> c\<guillemotright>" and "\<guillemotleft>g : b \<rightarrow> d\<guillemotright>"
shows "\<guillemotleft>f \<otimes> g : a \<otimes> b \<rightarrow> c \<otimes> d\<guillemotright>"
using assms prod_def by fastforce
lemma prod_in_hom' [simp]:
assumes "arr f" and "dom f = a" and "cod f = c"
and "arr g" and "dom g = b" and "cod g = d"
shows "\<guillemotleft>f \<otimes> g : a \<otimes> b \<rightarrow> c \<otimes> d\<guillemotright>"
using assms by blast
lemma prod_simps [simp]:
assumes "arr f0" and "arr f1"
shows "arr (f0 \<otimes> f1)"
and "dom (f0 \<otimes> f1) = dom f0 \<otimes> dom f1"
and "cod (f0 \<otimes> f1) = cod f0 \<otimes> cod f1"
using assms prod_in_hom by blast+
end
subsection "Agreement between the Definitions"
text \<open>
We now show that a category with binary products extends (by making a choice)
to an elementary category with binary products, and that the underlying category
of an elementary category with binary products is a category with binary products.
\<close>
context category_with_binary_products
begin
definition pr1
where "pr1 a b \<equiv> if ide a \<and> ide b then
fst (SOME x. has_as_binary_product a b (fst x) (snd x))
else null"
definition pr0
where "pr0 a b \<equiv> if ide a \<and> ide b then
snd (SOME x. has_as_binary_product a b (fst x) (snd x))
else null"
lemma pr_yields_binary_product:
assumes "ide a" and "ide b"
shows "has_as_binary_product a b (pr1 a b) (pr0 a b)"
proof -
have "\<exists>x. has_as_binary_product a b (fst x) (snd x)"
using assms has_binary_products has_binary_products_def has_as_binary_product_def
by simp
thus ?thesis
using assms has_binary_products has_binary_products_def pr0_def pr1_def
someI_ex [of "\<lambda>x. has_as_binary_product a b (fst x) (snd x)"]
by simp
qed
interpretation elementary_category_with_binary_products C pr0 pr1
proof
show "\<And>a b. \<not> (ide a \<and> ide b) \<Longrightarrow> pr0 a b = null"
using pr0_def by auto
show "\<And>a b. \<not> (ide a \<and> ide b) \<Longrightarrow> pr1 a b = null"
using pr1_def by auto
fix a b
assume a: "ide a" and b: "ide b"
interpret J: binary_product_shape .
interpret D: binary_product_diagram C a b
using a b by unfold_locales auto
let ?\<chi> = "D.mkCone (pr1 a b) (pr0 a b)"
interpret \<chi>: limit_cone J.comp C D.map \<open>dom (pr1 a b)\<close> ?\<chi>
using a b pr_yields_binary_product
by (simp add: has_as_binary_product_def)
have 1: "pr1 a b = ?\<chi> J.FF \<and> pr0 a b = ?\<chi> J.TT"
using D.mkCone_def by simp
show "span (pr1 a b) (pr0 a b)"
using 1 \<chi>.preserves_reflects_arr J.seqE J.arr_char J.seq_char J.is_category
D.is_rendered_commutative_by_cone \<chi>.cone_axioms
by metis
show "cod (pr1 a b) = a"
using 1 \<chi>.preserves_cod [of J.FF] J.cod_char J.arr_char by auto
show "cod (pr0 a b) = b"
using 1 \<chi>.preserves_cod [of J.TT] J.cod_char J.arr_char by auto
next
fix f g
assume fg: "span f g"
show "\<exists>!l. pr1 (cod f) (cod g) \<cdot> l = f \<and> pr0 (cod f) (cod g) \<cdot> l = g"
proof -
interpret J: binary_product_shape .
interpret D: binary_product_diagram C \<open>cod f\<close> \<open>cod g\<close>
using fg by unfold_locales auto
let ?\<chi> = "D.mkCone (pr1 (cod f) (cod g)) (pr0 (cod f) (cod g))"
interpret \<chi>: limit_cone J.comp C D.map \<open>dom (pr1 (cod f) (cod g))\<close> ?\<chi>
using fg pr_yields_binary_product [of "cod f" "cod g"] has_as_binary_product_def
by simp
interpret \<chi>: binary_product_cone C \<open>cod f\<close> \<open>cod g\<close>
\<open>pr1 (cod f) (cod g)\<close> \<open>pr0 (cod f) (cod g)\<close> ..
have 1: "pr1 (cod f) (cod g) = ?\<chi> J.FF \<and> pr0 (cod f) (cod g) = ?\<chi> J.TT"
using D.mkCone_def by simp
show "\<exists>!l. pr1 (cod f) (cod g) \<cdot> l = f \<and> pr0 (cod f) (cod g) \<cdot> l = g"
proof -
have "\<exists>!l. \<guillemotleft>l : dom f \<rightarrow> dom (pr1 (cod f) (cod g))\<guillemotright> \<and>
pr1 (cod f) (cod g) \<cdot> l = f \<and> pr0 (cod f) (cod g) \<cdot> l = g"
using fg \<chi>.is_universal' by simp
moreover have "\<And>l. pr1 (cod f) (cod g) \<cdot> l = f
\<Longrightarrow> \<guillemotleft>l : dom f \<rightarrow> dom (pr1 (cod f) (cod g))\<guillemotright>"
using fg dom_comp in_homI seqE seqI by metis
ultimately show ?thesis by auto
qed
qed
qed
proposition extends_to_elementary_category_with_binary_products:
shows "elementary_category_with_binary_products C pr0 pr1"
..
end
context elementary_category_with_binary_products
begin
interpretation category_with_binary_products C
proof
show "has_binary_products"
proof (unfold has_binary_products_def)
have "\<And>a b. ide a \<and> ide b \<Longrightarrow> \<exists>p0 p1. has_as_binary_product a b p0 p1"
proof -
fix a b
assume ab: "ide a \<and> ide b"
interpret J: binary_product_shape .
interpret D: binary_product_diagram C a b
using ab by unfold_locales auto
have 2: "D.is_rendered_commutative_by \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]"
using ab by simp
let ?\<chi> = "D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]"
interpret \<chi>: cone J.comp C D.map \<open>dom \<pp>\<^sub>1[a, b]\<close> ?\<chi>
using D.cone_mkCone 2 by auto
interpret \<chi>: limit_cone J.comp C D.map \<open>dom \<pp>\<^sub>1[a, b]\<close> ?\<chi>
proof
fix a' \<chi>'
assume \<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J.comp C D.map a' \<chi>'
using \<chi>' by simp
show "\<exists>!h. \<guillemotleft>h : a' \<rightarrow> dom \<pp>\<^sub>1[a, b]\<guillemotright> \<and>
D.cones_map h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) = \<chi>'"
proof
let ?h = "\<langle>\<chi>' J.FF, \<chi>' J.TT\<rangle>"
show h': "\<guillemotleft>?h : a' \<rightarrow> dom \<pp>\<^sub>1[a, b]\<guillemotright> \<and>
D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) = \<chi>'"
proof
show h: "\<guillemotleft>?h : a' \<rightarrow> dom \<pp>\<^sub>1[a, b]\<guillemotright>"
using ab tuple_in_hom J.ide_char by fastforce
show "D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) = \<chi>'"
proof -
interpret \<chi>'h: cone J.comp C D.map a'
\<open>D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b])\<close>
proof -
have "D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b] \<in> D.cones (cod \<langle>\<chi>' J.FF, \<chi>' J.TT\<rangle>)"
using ab h D.cone_mkCone D.is_rendered_commutative_by_cone
\<chi>.cone_axioms
by auto
hence "D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) \<in> D.cones a'"
using ab h D.cones_map_mapsto by blast
thus "D.cone a' (D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]))"
by simp
qed
show ?thesis
proof -
have "\<And>j. J.ide j \<Longrightarrow> D.cones_map ?h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) j = \<chi>' j"
using ab h J.ide_char D.mkCone_def \<chi>.cone_axioms by auto
thus ?thesis
using NaturalTransformation.eqI
\<chi>'.natural_transformation_axioms \<chi>'h.natural_transformation_axioms
by blast
qed
qed
qed
show "\<And>h. \<guillemotleft>h : a' \<rightarrow> dom \<pp>\<^sub>1[a, b]\<guillemotright> \<and>
D.cones_map h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) = \<chi>' \<Longrightarrow> h = ?h"
proof -
fix h
assume 1: "\<guillemotleft>h : a' \<rightarrow> dom \<pp>\<^sub>1[a, b]\<guillemotright> \<and>
D.cones_map h (D.mkCone \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]) = \<chi>'"
hence "cod h = dom \<pp>\<^sub>1[a, b]" by auto
show "h = ?h"
using 1 ab \<chi>.cone_axioms D.mkCone_def h' pr_joint_monic [of a b h ?h]
by auto
qed
qed
qed
have "has_as_binary_product a b \<pp>\<^sub>1[a, b] \<pp>\<^sub>0[a, b]"
using ab has_as_binary_product_def \<chi>.limit_cone_axioms by blast
thus "\<exists>p0 p1. has_as_binary_product a b p0 p1"
by blast
qed
thus "\<forall>a b. ide a \<and> ide b \<longrightarrow> (\<exists>p0 p1. has_as_binary_product a b p0 p1)"
by simp
qed
qed
proposition is_category_with_binary_products:
shows "category_with_binary_products C"
..
end
subsection "Further Properties"
context elementary_category_with_binary_products
begin
lemma interchange:
assumes "seq h f" and "seq k g"
shows "(h \<otimes> k) \<cdot> (f \<otimes> g) = h \<cdot> f \<otimes> k \<cdot> g"
using assms prod_def comp_tuple_arr comp_assoc by fastforce
lemma pr_naturality [simp]:
assumes "arr g" and "dom g = b" and "cod g = d"
and "arr f" and "dom f = a" and "cod f = c"
shows "\<pp>\<^sub>0[c, d] \<cdot> (f \<otimes> g) = g \<cdot> \<pp>\<^sub>0[a, b]"
and "\<pp>\<^sub>1[c, d] \<cdot> (f \<otimes> g) = f \<cdot> \<pp>\<^sub>1[a, b]"
using assms prod_def by fastforce+
abbreviation dup ("\<d>[_]")
where "\<d>[f] \<equiv> \<langle>f, f\<rangle>"
lemma dup_in_hom [intro, simp]:
assumes "\<guillemotleft>f : a \<rightarrow> b\<guillemotright>"
shows "\<guillemotleft>\<d>[f] : a \<rightarrow> b \<otimes> b\<guillemotright>"
using assms by fastforce
lemma dup_simps [simp]:
assumes "arr f"
shows "arr \<d>[f]" and "dom \<d>[f] = dom f" and "cod \<d>[f] = cod f \<otimes> cod f"
using assms dup_in_hom by auto
lemma dup_naturality:
assumes "\<guillemotleft>f : a \<rightarrow> b\<guillemotright>"
shows "\<d>[b] \<cdot> f = (f \<otimes> f) \<cdot> \<d>[a]"
using assms prod_def comp_arr_dom comp_cod_arr comp_tuple_arr comp_assoc
by fastforce
lemma pr_dup [simp]:
assumes "ide a"
shows "\<pp>\<^sub>0[a, a] \<cdot> \<d>[a] = a" and "\<pp>\<^sub>1[a, a] \<cdot> \<d>[a] = a"
using assms by simp_all
lemma prod_tuple:
assumes "span f g" and "seq h f" and "seq k g"
shows "(h \<otimes> k) \<cdot> \<langle>f, g\<rangle> = \<langle>h \<cdot> f, k \<cdot> g\<rangle>"
using assms prod_def comp_assoc comp_tuple_arr by fastforce
lemma tuple_eqI:
assumes "seq \<pp>\<^sub>0[b, c] f" and "seq \<pp>\<^sub>1[b, c] f"
and "\<pp>\<^sub>0[b, c] \<cdot> f = f0" and "\<pp>\<^sub>1[b, c] \<cdot> f = f1"
shows "f = \<langle>f1, f0\<rangle>"
using assms pr_joint_monic [of b c "\<langle>f1, f0\<rangle>" f] pr_tuple by auto
definition assoc ("\<a>[_, _, _]")
where "\<a>[a, b, c] \<equiv> \<langle>\<pp>\<^sub>1[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c], \<langle>\<pp>\<^sub>0[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c], \<pp>\<^sub>0[a \<otimes> b, c]\<rangle>\<rangle>"
definition assoc' ("\<a>\<^sup>-\<^sup>1[_, _, _]")
where "\<a>\<^sup>-\<^sup>1[a, b, c] \<equiv> \<langle>\<langle>\<pp>\<^sub>1[a, b \<otimes> c], \<pp>\<^sub>1[b, c] \<cdot> \<pp>\<^sub>0[a, b \<otimes> c]\<rangle>, \<pp>\<^sub>0[b, c] \<cdot> \<pp>\<^sub>0[a, b \<otimes> c]\<rangle>"
lemma assoc_in_hom [intro]:
assumes "ide a" and "ide b" and "ide c"
shows "\<guillemotleft>\<a>[a, b, c] : (a \<otimes> b) \<otimes> c \<rightarrow> a \<otimes> (b \<otimes> c)\<guillemotright>"
using assms assoc_def by auto
lemma assoc_simps [simp]:
assumes "ide a" and "ide b" and "ide c"
shows "arr \<a>[a, b, c]"
and "dom \<a>[a, b, c] = (a \<otimes> b) \<otimes> c"
and "cod \<a>[a, b, c] = a \<otimes> (b \<otimes> c)"
using assms assoc_in_hom by auto
lemma assoc'_in_hom [intro]:
assumes "ide a" and "ide b" and "ide c"
shows "\<guillemotleft>\<a>\<^sup>-\<^sup>1[a, b, c] : a \<otimes> (b \<otimes> c) \<rightarrow> (a \<otimes> b) \<otimes> c\<guillemotright>"
using assms assoc'_def by auto
lemma assoc'_simps [simp]:
assumes "ide a" and "ide b" and "ide c"
shows "arr \<a>\<^sup>-\<^sup>1[a, b, c]"
and "dom \<a>\<^sup>-\<^sup>1[a, b, c] = a \<otimes> (b \<otimes> c)"
and "cod \<a>\<^sup>-\<^sup>1[a, b, c] = (a \<otimes> b) \<otimes> c"
using assms assoc'_in_hom by auto
lemma assoc_naturality:
assumes "\<guillemotleft>f0 : a0 \<rightarrow> b0\<guillemotright>" and "\<guillemotleft>f1 : a1 \<rightarrow> b1\<guillemotright>" and "\<guillemotleft>f2 : a2 \<rightarrow> b2\<guillemotright>"
shows "\<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2) = (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"
proof -
have "\<pp>\<^sub>0[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2) =
\<pp>\<^sub>0[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"
proof -
have "\<pp>\<^sub>0[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2) =
(\<pp>\<^sub>0[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2]) \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
using comp_assoc by simp
also have "... = \<langle>\<pp>\<^sub>0[b0, b1] \<cdot> \<pp>\<^sub>1[b0 \<otimes> b1, b2], \<pp>\<^sub>0[b0 \<otimes> b1, b2]\<rangle> \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
using assms assoc_def by fastforce
also have "... = \<langle>(\<pp>\<^sub>0[b0, b1] \<cdot> \<pp>\<^sub>1[b0 \<otimes> b1, b2]) \<cdot> ((f0 \<otimes> f1) \<otimes> f2),
\<pp>\<^sub>0[b0 \<otimes> b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2)\<rangle>"
using assms comp_tuple_arr by fastforce
also have "... = \<langle>(\<pp>\<^sub>0[b0, b1] \<cdot> (f0 \<otimes> f1)) \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2], f2 \<cdot> \<pp>\<^sub>0[a0 \<otimes> a1, a2]\<rangle>"
using assms comp_assoc by fastforce
also have "... = \<langle>f1 \<cdot> \<pp>\<^sub>0[a0, a1] \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2], f2 \<cdot> \<pp>\<^sub>0[a0 \<otimes> a1, a2]\<rangle>"
using assms comp_assoc
by (metis in_homE pr_naturality(1))
also have "... = \<pp>\<^sub>0[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"
using assms comp_assoc assoc_def prod_tuple by fastforce
finally show ?thesis by blast
qed
moreover have "\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2) =
\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"
proof -
have "\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2) =
(\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> \<a>[b0, b1, b2]) \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
using comp_assoc by simp
also have "... = (\<pp>\<^sub>1[b0, b1] \<cdot> \<pp>\<^sub>1[b0 \<otimes> b1, b2]) \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
using assms assoc_def by fastforce
also have "... = (\<pp>\<^sub>1[b0, b1] \<cdot> (f0 \<otimes> f1)) \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2]"
using assms comp_assoc by fastforce
also have "... = f0 \<cdot> \<pp>\<^sub>1[a0, a1] \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2]"
using assms comp_assoc
by (metis in_homE pr_naturality(2))
also have "... = \<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"
proof -
have "\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2] =
(\<pp>\<^sub>1[b0, b1 \<otimes> b2] \<cdot> (f0 \<otimes> (f1 \<otimes> f2))) \<cdot> \<a>[a0, a1, a2]"
using comp_assoc by simp
also have "... = f0 \<cdot> \<pp>\<^sub>1[a0, a1 \<otimes> a2] \<cdot> \<a>[a0, a1, a2]"
using assms comp_assoc by fastforce
also have "... = f0 \<cdot> \<pp>\<^sub>1[a0, a1] \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2]"
using assms assoc_def by fastforce
finally show ?thesis by simp
qed
finally show ?thesis by blast
qed
ultimately show ?thesis
using assms pr_joint_monic [of b0 "b1 \<otimes> b2" "\<a>[b0, b1, b2] \<cdot> ((f0 \<otimes> f1) \<otimes> f2)"
"(f0 \<otimes> (f1 \<otimes> f2)) \<cdot> \<a>[a0, a1, a2]"]
by fastforce
qed
lemma pentagon:
assumes "ide a" and "ide b" and "ide c" and "ide d"
shows "((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) = \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
proof (intro pr_joint_monic
[where h = "((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d)"
and h' = "\<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"])
show "seq \<pp>\<^sub>0[a, b \<otimes> (c \<otimes> d)] (((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d))"
using assms by simp
show "\<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> ((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) =
\<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
proof -
have "\<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> ((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) =
((\<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> (a \<otimes> \<a>[b, c, d])) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d)"
using comp_assoc by simp
also have "... = (\<pp>\<^sub>1[a, (b \<otimes> c) \<otimes> d] \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d)"
using assms pr_naturality(2) comp_cod_arr by force
also have "... = \<pp>\<^sub>1[a, b \<otimes> c] \<cdot> \<pp>\<^sub>1[a \<otimes> b \<otimes> c, d] \<cdot> (\<a>[a, b, c] \<otimes> d)"
using assms assoc_def comp_assoc by simp
also have "... = (\<pp>\<^sub>1[a, b \<otimes> c] \<cdot> \<a>[a, b, c]) \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d]"
using assms pr_naturality(2) comp_assoc by fastforce
also have "... = \<pp>\<^sub>1[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d]"
using assms assoc_def comp_assoc by simp
finally have "\<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> ((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) =
\<pp>\<^sub>1[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d]"
by blast
also have "... = \<pp>\<^sub>1[a, b \<otimes> c \<otimes> d] \<cdot> \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
using assms assoc_def comp_assoc by auto
finally show ?thesis by blast
qed
show "\<pp>\<^sub>0[a, b \<otimes> (c \<otimes> d)] \<cdot> ((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) =
\<pp>\<^sub>0[a, b \<otimes> (c \<otimes> d)] \<cdot> \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
proof -
have "\<pp>\<^sub>0[a, b \<otimes> (c \<otimes> d)] \<cdot> ((a \<otimes> \<a>[b, c, d]) \<cdot> \<a>[a, b \<otimes> c, d]) \<cdot> (\<a>[a, b, c] \<otimes> d) =
\<pp>\<^sub>0[a, b \<otimes> c \<otimes> d] \<cdot>
((a \<otimes> \<langle>\<pp>\<^sub>1[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<langle>\<pp>\<^sub>0[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<pp>\<^sub>0[b \<otimes> c, d]\<rangle>\<rangle>) \<cdot>
\<langle>\<pp>\<^sub>1[a, b \<otimes> c] \<cdot> \<pp>\<^sub>1[a \<otimes> b \<otimes> c, d],
\<langle>\<pp>\<^sub>0[a, b \<otimes> c] \<cdot> \<pp>\<^sub>1[a \<otimes> b \<otimes> c, d], \<pp>\<^sub>0[a \<otimes> b \<otimes> c, d]\<rangle>\<rangle>) \<cdot>
(\<langle>\<pp>\<^sub>1[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c], \<langle>\<pp>\<^sub>0[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c], \<pp>\<^sub>0[a \<otimes> b, c]\<rangle>\<rangle> \<otimes> d)"
using assms assoc_def by simp
also have "... = \<langle>\<pp>\<^sub>1[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d],
\<langle>\<pp>\<^sub>0[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<pp>\<^sub>0[b \<otimes> c, d]\<rangle>\<rangle> \<cdot> (\<pp>\<^sub>0[a, (b \<otimes> c) \<otimes> d] \<cdot>
\<langle>\<pp>\<^sub>1[a, b \<otimes> c] \<cdot> \<pp>\<^sub>1[a \<otimes> b \<otimes> c, d],
\<langle>\<pp>\<^sub>0[a, b \<otimes> c] \<cdot> \<pp>\<^sub>1[a \<otimes> b \<otimes> c, d], \<pp>\<^sub>0[a \<otimes> b \<otimes> c, d]\<rangle>\<rangle>) \<cdot>
(\<langle>\<pp>\<^sub>1[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c],
\<langle>\<pp>\<^sub>0[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c], \<pp>\<^sub>0[a \<otimes> b, c]\<rangle>\<rangle> \<otimes> d)"
proof -
have "\<pp>\<^sub>0[a, b \<otimes> c \<otimes> d] \<cdot>
(a \<otimes> \<langle>\<pp>\<^sub>1[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<langle>\<pp>\<^sub>0[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<pp>\<^sub>0[b \<otimes> c, d]\<rangle>\<rangle>) =
\<langle>\<pp>\<^sub>1[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<langle>\<pp>\<^sub>0[b, c] \<cdot> \<pp>\<^sub>1[b \<otimes> c, d], \<pp>\<^sub>0[b \<otimes> c, d]\<rangle>\<rangle> \<cdot>
\<pp>\<^sub>0[a, (b \<otimes> c) \<otimes> d]"
using assms assoc_def ide_in_hom pr_naturality(1) by auto
thus ?thesis using comp_assoc by metis
qed
also have "... = \<langle>\<pp>\<^sub>0[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d],
\<langle>\<pp>\<^sub>0[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d], d \<cdot> \<pp>\<^sub>0[(a \<otimes> b) \<otimes> c, d]\<rangle>\<rangle>"
using assms comp_assoc by simp
also have "... = \<langle>\<pp>\<^sub>0[a, b] \<cdot> \<pp>\<^sub>1[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d],
\<langle>\<pp>\<^sub>0[a \<otimes> b, c] \<cdot> \<pp>\<^sub>1[(a \<otimes> b) \<otimes> c, d], \<pp>\<^sub>0[(a \<otimes> b) \<otimes> c, d]\<rangle>\<rangle>"
using assms comp_cod_arr by simp
also have "... = \<pp>\<^sub>0[a, b \<otimes> (c \<otimes> d)] \<cdot> \<a>[a, b, c \<otimes> d] \<cdot> \<a>[a \<otimes> b, c, d]"
using assms assoc_def comp_assoc by simp
finally show ?thesis by simp
qed
qed
lemma inverse_arrows_assoc:
assumes "ide a" and "ide b" and "ide c"
shows "inverse_arrows \<a>[a, b, c] \<a>\<^sup>-\<^sup>1[a, b, c]"
using assms assoc_def assoc'_def comp_assoc
by (auto simp add: tuple_pr_arr)
interpretation CC: product_category C C ..
abbreviation Prod
where "Prod fg \<equiv> fst fg \<otimes> snd fg"
abbreviation Prod'
where "Prod' fg \<equiv> snd fg \<otimes> fst fg"
interpretation \<Pi>: binary_functor C C C Prod
using tuple_ext CC.comp_char interchange
apply unfold_locales
apply auto
by (metis prod_def seqE)+
interpretation Prod': binary_functor C C C Prod'
using tuple_ext CC.comp_char interchange
apply unfold_locales
apply auto
by (metis prod_def seqE)+
lemma binary_functor_Prod:
shows "binary_functor C C C Prod" and "binary_functor C C C Prod'"
..
definition sym ("\<s>[_, _]")
where "\<s>[a1, a0] \<equiv> if ide a0 \<and> ide a1 then \<langle>\<pp>\<^sub>0[a1, a0], \<pp>\<^sub>1[a1, a0]\<rangle> else null"
lemma sym_in_hom [intro]:
assumes "ide a" and "ide b"
shows "\<guillemotleft>\<s>[a, b] : a \<otimes> b \<rightarrow> b \<otimes> a\<guillemotright>"
using assms sym_def by auto
lemma sym_simps [simp]:
assumes "ide a" and "ide b"
shows "arr \<s>[a, b]" and "dom \<s>[a, b] = a \<otimes> b" and "cod \<s>[a, b] = b \<otimes> a"
using assms sym_in_hom by auto
lemma comp_sym_tuple [simp]:
assumes "\<guillemotleft>f0 : a \<rightarrow> b0\<guillemotright>" and "\<guillemotleft>f1 : a \<rightarrow> b1\<guillemotright>"
shows "\<s>[b0, b1] \<cdot> \<langle>f0, f1\<rangle> = \<langle>f1, f0\<rangle>"
using assms sym_def comp_tuple_arr by fastforce
lemma prj_sym [simp]:
assumes "ide a0" and "ide a1"
shows "\<pp>\<^sub>0[a1, a0] \<cdot> \<s>[a0, a1] = \<pp>\<^sub>1[a0, a1]"
and "\<pp>\<^sub>1[a1, a0] \<cdot> \<s>[a0, a1] = \<pp>\<^sub>0[a0, a1]"
using assms sym_def by auto
lemma comp_sym_sym [simp]:
assumes "ide a0" and "ide a1"
shows "\<s>[a1, a0] \<cdot> \<s>[a0, a1] = (a0 \<otimes> a1)"
using assms sym_def comp_tuple_arr by auto
lemma sym_inverse_arrows:
assumes "ide a0" and "ide a1"
shows "inverse_arrows \<s>[a0, a1] \<s>[a1, a0]"
using assms sym_in_hom comp_sym_sym by auto
lemma sym_assoc_coherence:
assumes "ide a" and "ide b" and "ide c"
shows "\<a>[b, c, a] \<cdot> \<s>[a, b \<otimes> c] \<cdot> \<a>[a, b, c] = (b \<otimes> \<s>[a, c]) \<cdot> \<a>[b, a, c] \<cdot> (\<s>[a, b] \<otimes> c)"
using assms sym_def assoc_def comp_assoc prod_tuple comp_cod_arr by simp
lemma sym_naturality:
assumes "\<guillemotleft>f0 : a0 \<rightarrow> b0\<guillemotright>" and "\<guillemotleft>f1 : a1 \<rightarrow> b1\<guillemotright>"
shows "\<s>[b0, b1] \<cdot> (f0 \<otimes> f1) = (f1 \<otimes> f0) \<cdot> \<s>[a0, a1]"
using assms sym_def comp_assoc prod_tuple by fastforce
abbreviation \<sigma>
where "\<sigma> fg \<equiv> \<s>[cod (fst fg), cod (snd fg)] \<cdot> (fst fg \<otimes> snd fg)"
interpretation \<sigma>: natural_transformation CC.comp C Prod Prod' \<sigma>
using sym_def CC.arr_char CC.null_char comp_arr_dom comp_cod_arr
apply unfold_locales
apply auto
using arr_cod_iff_arr ideD(1)
apply metis
using arr_cod_iff_arr ideD(1)
apply metis
using prod_tuple by simp
lemma \<sigma>_is_natural_transformation:
shows "natural_transformation CC.comp C Prod Prod' \<sigma>"
..
abbreviation Diag
where "Diag f \<equiv> if arr f then (f, f) else CC.null"
interpretation \<Delta>: "functor" C CC.comp Diag
by (unfold_locales, auto)
lemma functor_Diag:
shows "functor C CC.comp Diag"
..
interpretation \<Delta>o\<Pi>: composite_functor CC.comp C CC.comp Prod Diag ..
interpretation \<Pi>o\<Delta>: composite_functor C CC.comp C Diag Prod ..
abbreviation \<pi>
where "\<pi> \<equiv> \<lambda>(f, g). (\<pp>\<^sub>1[cod f, cod g] \<cdot> (f \<otimes> g), \<pp>\<^sub>0[cod f, cod g] \<cdot> (f \<otimes> g))"
interpretation \<pi>: transformation_by_components CC.comp CC.comp \<Delta>o\<Pi>.map CC.map \<pi>
using pr_naturality comp_arr_dom comp_cod_arr
by unfold_locales auto
lemma \<pi>_is_natural_transformation:
shows "natural_transformation CC.comp CC.comp \<Delta>o\<Pi>.map CC.map \<pi>"
proof -
have "\<pi>.map = \<pi>"
using \<pi>.map_def ext \<Pi>.is_extensional comp_arr_dom comp_cod_arr by auto
thus "natural_transformation CC.comp CC.comp \<Delta>o\<Pi>.map CC.map \<pi>"
using \<pi>.natural_transformation_axioms by simp
qed
interpretation \<delta>: natural_transformation C C map \<Pi>o\<Delta>.map dup
using dup_naturality comp_arr_dom comp_cod_arr prod_tuple tuple_ext
by unfold_locales auto
lemma dup_is_natural_transformation:
shows "natural_transformation C C map \<Pi>o\<Delta>.map dup"
..
interpretation \<Delta>o\<Pi>o\<Delta>: composite_functor C CC.comp CC.comp Diag \<Delta>o\<Pi>.map ..
interpretation \<Pi>o\<Delta>o\<Pi>: composite_functor CC.comp C C Prod \<Pi>o\<Delta>.map ..
interpretation \<Delta>o\<delta>: natural_transformation C CC.comp Diag \<Delta>o\<Pi>o\<Delta>.map \<open>Diag \<circ> dup\<close>
proof -
have "Diag \<circ> map = Diag"
by auto
thus "natural_transformation C CC.comp Diag \<Delta>o\<Pi>o\<Delta>.map (Diag \<circ> dup)"
using \<Delta>.natural_transformation_axioms \<delta>.natural_transformation_axioms o_assoc
horizontal_composite [of C C map \<Pi>o\<Delta>.map dup CC.comp Diag Diag Diag]
by metis
qed
interpretation \<delta>o\<Pi>: natural_transformation CC.comp C Prod \<Pi>o\<Delta>o\<Pi>.map \<open>dup \<circ> Prod\<close>
using \<delta>.natural_transformation_axioms \<Pi>.natural_transformation_axioms o_assoc
horizontal_composite [of CC.comp C Prod Prod Prod C map \<Pi>o\<Delta>.map dup]
by simp
interpretation \<pi>o\<Delta>: natural_transformation C CC.comp \<Delta>o\<Pi>o\<Delta>.map Diag \<open>\<pi>.map \<circ> Diag\<close>
using \<pi>.natural_transformation_axioms \<Delta>.natural_transformation_axioms
horizontal_composite
[of C CC.comp Diag Diag Diag CC.comp \<Delta>o\<Pi>.map CC.map \<pi>.map]
by simp
interpretation \<Pi>o\<pi>: natural_transformation CC.comp C \<Pi>o\<Delta>o\<Pi>.map Prod \<open>Prod \<circ> \<pi>.map\<close>
proof -
have "Prod \<circ> \<Delta>o\<Pi>.map = \<Pi>o\<Delta>o\<Pi>.map"
by auto
thus "natural_transformation CC.comp C \<Pi>o\<Delta>o\<Pi>.map Prod (Prod \<circ> \<pi>.map)"
using \<pi>.natural_transformation_axioms \<Pi>.natural_transformation_axioms o_assoc
horizontal_composite
[of CC.comp CC.comp \<Delta>o\<Pi>.map CC.map \<pi>.map C Prod Prod Prod]
by simp
qed
interpretation \<Delta>o\<delta>_\<pi>o\<Delta>: vertical_composite C CC.comp Diag \<Delta>o\<Pi>o\<Delta>.map Diag
\<open>Diag \<circ> dup\<close> \<open>\<pi>.map \<circ> Diag\<close>
..
interpretation \<Pi>o\<pi>_\<delta>o\<Pi>: vertical_composite CC.comp C Prod \<Pi>o\<Delta>o\<Pi>.map Prod
\<open>dup \<circ> Prod\<close> \<open>Prod \<circ> \<pi>.map\<close>
..
interpretation \<Delta>\<Pi>: unit_counit_adjunction CC.comp C Diag Prod dup \<pi>.map
proof
show "\<Delta>o\<delta>_\<pi>o\<Delta>.map = Diag"
proof
fix f
have "\<not> arr f \<Longrightarrow> \<Delta>o\<delta>_\<pi>o\<Delta>.map f = Diag f"
by (simp add: \<Delta>o\<delta>_\<pi>o\<Delta>.is_extensional)
moreover have "arr f \<Longrightarrow> \<Delta>o\<delta>_\<pi>o\<Delta>.map f = Diag f"
using comp_cod_arr comp_assoc \<Delta>o\<delta>_\<pi>o\<Delta>.map_def by auto
ultimately show "\<Delta>o\<delta>_\<pi>o\<Delta>.map f = Diag f" by blast
qed
show "\<Pi>o\<pi>_\<delta>o\<Pi>.map = Prod"
proof
fix fg
show "\<Pi>o\<pi>_\<delta>o\<Pi>.map fg = Prod fg"
proof -
have "\<not> CC.arr fg \<Longrightarrow> ?thesis"
by (simp add: \<Pi>.is_extensional \<Pi>o\<pi>_\<delta>o\<Pi>.is_extensional)
moreover have "CC.arr fg \<Longrightarrow> ?thesis"
proof -
assume fg: "CC.arr fg"
have 1: "dup (Prod fg) = \<langle>cod (fst fg) \<otimes> cod (snd fg), cod (fst fg) \<otimes> cod (snd fg)\<rangle> \<cdot>
(fst fg \<otimes> snd fg)"
using fg \<delta>.is_natural_2
apply simp
by (metis (no_types, lifting) prod_simps(1) prod_simps(3))
have "\<Pi>o\<pi>_\<delta>o\<Pi>.map fg =
(\<pp>\<^sub>1[cod (fst fg), cod (snd fg)] \<otimes> \<pp>\<^sub>0[cod (fst fg), cod (snd fg)]) \<cdot>
\<langle>cod (fst fg) \<otimes> cod (snd fg), cod (fst fg) \<otimes> cod (snd fg)\<rangle> \<cdot>
(fst fg \<otimes> snd fg)"
using fg 1 \<Pi>o\<pi>_\<delta>o\<Pi>.map_def comp_cod_arr by simp
also have "... = ((\<pp>\<^sub>1[cod (fst fg), cod (snd fg)] \<otimes> \<pp>\<^sub>0[cod (fst fg), cod (snd fg)]) \<cdot>
\<langle>cod (fst fg) \<otimes> cod (snd fg), cod (fst fg) \<otimes> cod (snd fg)\<rangle>) \<cdot>
(fst fg \<otimes> snd fg)"
using comp_assoc by simp
also have "... = \<langle>\<pp>\<^sub>1[cod (fst fg), cod (snd fg)] \<cdot> (cod (fst fg) \<otimes> cod (snd fg)),
\<pp>\<^sub>0[cod (fst fg), cod (snd fg)] \<cdot> (cod (fst fg) \<otimes> cod (snd fg))\<rangle> \<cdot>
(fst fg \<otimes> snd fg)"
using fg prod_tuple by simp
also have "... = Prod fg"
using fg comp_arr_dom \<Pi>.is_natural_2 by auto
finally show ?thesis by simp
qed
ultimately show ?thesis by blast
qed
qed
qed
proposition induces_unit_counit_adjunction:
shows "unit_counit_adjunction CC.comp C Diag Prod dup \<pi>.map"
using \<Delta>\<Pi>.unit_counit_adjunction_axioms by simp
end
section "Category with Terminal Object"
locale category_with_terminal_object =
category +
assumes has_terminal: "\<exists>t. terminal t"
locale elementary_category_with_terminal_object =
category C
for C :: "'a comp" (infixr "\<cdot>" 55)
and one :: "'a" ("\<one>")
and trm :: "'a \<Rightarrow> 'a" ("\<t>[_]") +
assumes ide_one: "ide \<one>"
and trm_in_hom: "ide a \<Longrightarrow> \<guillemotleft>\<t>[a] : a \<rightarrow> \<one>\<guillemotright>"
and trm_eqI: "\<lbrakk> ide a; \<guillemotleft>f : a \<rightarrow> \<one>\<guillemotright> \<rbrakk> \<Longrightarrow> f = \<t>[a]"
begin
lemma trm_simps:
assumes "ide a"
shows "arr \<t>[a]" and "dom \<t>[a] = a" and "cod \<t>[a] = \<one>"
using assms trm_in_hom by auto
lemma trm_one:
shows "\<t>[\<one>] = \<one>"
using ide_one trm_in_hom trm_eqI ide_in_hom by auto
lemma terminal_one:
shows "terminal \<one>"
using ide_one trm_in_hom trm_eqI terminal_def by metis
lemma trm_naturality:
assumes "arr f"
shows "\<t>[cod f] \<cdot> f = \<t>[dom f]"
using assms trm_eqI
by (metis comp_in_homI' ide_cod ide_dom in_homE trm_in_hom)
proposition is_category_with_terminal_object:
shows "category_with_terminal_object C"
apply unfold_locales
using terminal_one by auto
end
context category_with_terminal_object
begin
definition some_terminal ("\<one>")
where "some_terminal \<equiv> SOME t. terminal t"
definition "trm" ("\<t>[_]")
where "\<t>[f] \<equiv> if arr f then THE t. \<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright> else null"
lemma terminal_some_terminal [intro]:
shows "terminal \<one>"
using some_terminal_def has_terminal someI_ex [of "\<lambda>t. terminal t"] by presburger
lemma ide_some_terminal:
shows "ide \<one>"
using terminal_def by blast
lemma trm_in_hom [intro]:
assumes "arr f"
shows "\<guillemotleft>\<t>[f] : dom f \<rightarrow> \<one>\<guillemotright>"
proof -
have "ide (dom f)" using assms by fastforce
hence "\<exists>!t. \<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>"
using assms trm_def terminal_def terminal_some_terminal by simp
thus ?thesis
using assms trm_def [of f] theI' [of "\<lambda>t. \<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>"] by auto
qed
lemma trm_simps [simp]:
assumes "arr f"
shows "arr \<t>[f]" and "dom \<t>[f] = dom f" and "cod \<t>[f] = \<one>"
using assms trm_in_hom by auto
lemma trm_eqI:
assumes "\<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>"
shows "t = \<t>[f]"
proof -
have "ide (dom f)" using assms
by (metis ide_dom in_homE)
hence "\<exists>!t. \<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>"
using terminal_def [of \<one>] terminal_some_terminal by auto
moreover have "\<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>"
using assms by simp
ultimately show ?thesis
using assms trm_def the1_equality [of "\<lambda>t. \<guillemotleft>t : dom f \<rightarrow> \<one>\<guillemotright>" t]
\<open>ide (dom f)\<close> arr_dom_iff_arr
by fastforce
qed
sublocale elementary_category_with_terminal_object C \<one> trm
using ide_some_terminal trm_eqI
by unfold_locales auto
proposition extends_to_elementary_category_with_terminal_object:
shows "elementary_category_with_terminal_object C \<one> trm"
..
end
section "Cartesian Category"
locale cartesian_category =
category_with_binary_products +
category_with_terminal_object
locale elementary_cartesian_category =
elementary_category_with_binary_products +
elementary_category_with_terminal_object
begin
proposition is_cartesian_category:
shows "cartesian_category C"
using cartesian_category.intro is_category_with_binary_products
is_category_with_terminal_object
by auto
end
context cartesian_category
begin
proposition extends_to_elementary_cartesian_category:
shows "elementary_cartesian_category C pr0 pr1 \<one> trm"
by (simp add: elementary_cartesian_category_def
elementary_category_with_terminal_object_axioms
extends_to_elementary_category_with_binary_products)
sublocale elementary_cartesian_category C pr0 pr1 \<one> trm
using extends_to_elementary_cartesian_category by simp
end
text \<open>
Here we prove some facts that will later allow us to show that an elementary cartesian
category is a monoidal category.
\<close>
context elementary_cartesian_category
begin
abbreviation \<iota>
where "\<iota> \<equiv> \<pp>\<^sub>0[\<one>, \<one>]"
lemma pr_coincidence:
shows "\<iota> = \<pp>\<^sub>1[\<one>, \<one>]"
using ide_one
by (simp add: terminal_arr_unique terminal_one)
lemma \<iota>_is_terminal_arr:
shows "terminal_arr \<iota>"
using ide_one
by (simp add: terminal_one)
lemma inverse_arrows_\<iota>:
shows "inverse_arrows \<iota> \<langle>\<one>, \<one>\<rangle>"
using ide_one
by (metis (no_types, lifting) dup_is_natural_transformation \<iota>_is_terminal_arr cod_pr0
comp_cod_arr pr_dup(1) ide_dom inverse_arrows_def map_simp
natural_transformation.is_natural_2 pr_simps(2) pr1_in_hom' trm_eqI trm_naturality
trm_one tuple_pr)
lemma \<iota>_is_iso:
shows "iso \<iota>"
using inverse_arrows_\<iota> by auto
lemma trm_tensor:
assumes "ide a" and "ide b"
shows "\<t>[a \<otimes> b] = \<iota> \<cdot> (\<t>[a] \<otimes> \<t>[b])"
proof -
have "\<t>[a \<otimes> b] = \<t>[a] \<cdot> \<pp>\<^sub>1[a, b]"
by (metis assms(1-2) cod_pr1 pr_simps(4-6) trm_naturality)
moreover have "\<guillemotleft>\<t>[b] : b \<rightarrow> \<one>\<guillemotright>"
using assms(2) trm_in_hom by blast
ultimately show ?thesis
using assms(1) pr_coincidence trm_in_hom by fastforce
qed
abbreviation runit ("\<r>[_]")
where "\<r>[a] \<equiv> \<pp>\<^sub>1[a, \<one>]"
abbreviation runit' ("\<r>\<^sup>-\<^sup>1[_]")
where "\<r>\<^sup>-\<^sup>1[a] \<equiv> \<langle>a, \<t>[a]\<rangle>"
abbreviation lunit ("\<l>[_]")
where "\<l>[a] \<equiv> \<pp>\<^sub>0[\<one>, a]"
abbreviation lunit' ("\<l>\<^sup>-\<^sup>1[_]")
where "\<l>\<^sup>-\<^sup>1[a] \<equiv> \<langle>\<t>[a], a\<rangle>"
lemma runit_in_hom:
assumes "ide a"
shows "\<guillemotleft>\<r>[a] : a \<otimes> \<one> \<rightarrow> a\<guillemotright>"
using assms ide_one by simp
lemma runit'_in_hom:
assumes "ide a"
shows "\<guillemotleft>\<r>\<^sup>-\<^sup>1[a] : a \<rightarrow> a \<otimes> \<one>\<guillemotright>"
using assms ide_in_hom trm_in_hom by blast
lemma lunit_in_hom:
assumes "ide a"
shows "\<guillemotleft>\<l>[a] : \<one> \<otimes> a \<rightarrow> a\<guillemotright>"
using assms ide_one by simp
lemma lunit'_in_hom:
assumes "ide a"
shows "\<guillemotleft>\<l>\<^sup>-\<^sup>1[a] : a \<rightarrow> \<one> \<otimes> a\<guillemotright>"
using assms ide_in_hom trm_in_hom by blast
lemma runit_naturality:
assumes "ide a"
shows "\<r>[cod a] \<cdot> (a \<otimes> \<one>) = a \<cdot> \<r>[dom a]"
using assms pr_naturality(2) ide_char ide_one by blast
lemma inverse_arrows_runit:
assumes "ide a"
shows "inverse_arrows \<r>[a] \<r>\<^sup>-\<^sup>1[a]"
proof
show "ide (\<r>[a] \<cdot> \<r>\<^sup>-\<^sup>1[a])"
proof -
have "\<r>[a] \<cdot> \<r>\<^sup>-\<^sup>1[a] = a"
using assms
by (metis in_homE ide_char pr_tuple(1) trm_in_hom)
thus ?thesis
using assms by presburger
qed
show "ide (\<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a])"
proof -
have "ide (a \<otimes> \<one>)"
using assms ide_one by blast
moreover have "\<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a] = a \<otimes> \<one>"
proof (intro pr_joint_monic [of a \<one> "\<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a]" "a \<otimes> \<one>"])
show "seq \<pp>\<^sub>0[a, \<one>] (\<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a])"
using assms ide_one runit'_in_hom [of a]
by (intro seqI) auto
show "\<pp>\<^sub>0[a, \<one>] \<cdot> \<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a] = \<pp>\<^sub>0[a, \<one>] \<cdot> (a \<otimes> \<one>)"
proof -
have "\<pp>\<^sub>0[a, \<one>] \<cdot> \<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a] = (\<pp>\<^sub>0[a, \<one>] \<cdot> \<r>\<^sup>-\<^sup>1[a]) \<cdot> \<r>[a]"
using comp_assoc by simp
also have "... = \<t>[a] \<cdot> \<r>[a]"
using assms ide_one
by (metis in_homE pr_tuple(2) ide_char trm_in_hom)
also have "... = \<t>[a \<otimes> \<one>]"
using assms ide_one trm_naturality [of "\<r>[a]"] by simp
also have "... = \<pp>\<^sub>0[a, \<one>] \<cdot> (a \<otimes> \<one>)"
using assms comp_arr_dom ide_one trm_naturality trm_one by fastforce
finally show ?thesis by blast
qed
show "\<pp>\<^sub>1[a, \<one>] \<cdot> \<r>\<^sup>-\<^sup>1[a] \<cdot> \<r>[a] = \<pp>\<^sub>1[a, \<one>] \<cdot> (a \<otimes> \<one>)"
using assms
by (metis \<open>ide (\<r>[a] \<cdot> \<r>\<^sup>-\<^sup>1[a])\<close> cod_comp cod_pr1 dom_comp ide_compE ide_one
comp_assoc runit_naturality)
qed
ultimately show ?thesis by simp
qed
qed
lemma lunit_naturality:
assumes "arr f"
shows "C \<l>[cod f] (\<one> \<otimes> f) = C f \<l>[dom f]"
using assms pr_naturality(1) ide_char ide_one by blast
lemma inverse_arrows_lunit:
assumes "ide a"
shows "inverse_arrows \<l>[a] \<l>\<^sup>-\<^sup>1[a]"
proof
show "ide (C \<l>[a] \<l>\<^sup>-\<^sup>1[a])"
proof -
have "C \<l>[a] \<l>\<^sup>-\<^sup>1[a] = a"
using assms
by (metis ide_char in_homE pr_tuple(2) trm_in_hom)
thus ?thesis
using assms by simp
qed
show "ide (\<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a])"
proof -
have "\<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a] = \<one> \<otimes> a"
proof (intro pr_joint_monic [of \<one> a "\<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a]" "\<one> \<otimes> a"])
show "seq \<l>[a] (\<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a])"
using assms \<open>ide (\<l>[a] \<cdot> \<l>\<^sup>-\<^sup>1[a])\<close> by blast
show "\<l>[a] \<cdot> \<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a] = \<l>[a] \<cdot> (\<one> \<otimes> a)"
using assms
by (metis \<open>ide (\<l>[a] \<cdot> \<l>\<^sup>-\<^sup>1[a])\<close> cod_comp cod_pr0 dom_cod ide_compE ide_one
comp_assoc lunit_naturality)
show "\<pp>\<^sub>1[\<one>, a] \<cdot> \<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a] = \<pp>\<^sub>1[\<one>, a] \<cdot> (\<one> \<otimes> a)"
proof -
have "\<pp>\<^sub>1[\<one>, a] \<cdot> \<l>\<^sup>-\<^sup>1[a] \<cdot> \<l>[a] = (\<pp>\<^sub>1[\<one>, a] \<cdot> \<l>\<^sup>-\<^sup>1[a]) \<cdot> \<l>[a]"
using comp_assoc by simp
also have "... = \<t>[a] \<cdot> \<l>[a]"
using assms ide_one
by (metis pr_tuple(1) ide_char in_homE trm_in_hom)
also have "... = \<t>[\<one> \<otimes> a]"
using assms ide_one trm_naturality [of "\<l>[a]"] by simp
also have "... = \<pp>\<^sub>1[\<one>, a] \<cdot> (\<one> \<otimes> a)"
using assms comp_arr_dom ide_one trm_naturality trm_one by fastforce
finally show ?thesis by simp
qed
qed
moreover have "ide (\<one> \<otimes> a)"
using assms ide_one by simp
finally show ?thesis by blast
qed
qed
lemma comp_lunit_term_dup:
assumes "ide a"
shows "\<l>[a] \<cdot> (\<t>[a] \<otimes> a) \<cdot> \<d>[a] = a"
proof -
have "\<guillemotleft>\<t>[a] : a \<rightarrow> \<one>\<guillemotright>"
using assms trm_in_hom by blast
hence "\<l>[a] \<cdot> (\<t>[a] \<otimes> a) = a \<cdot> \<pp>\<^sub>0[a, a]"
by (metis assms pr_naturality(1) ide_char in_homE)
thus ?thesis
by (metis (no_types) assms comp_assoc comp_ide_self pr_dup(1))
qed
lemma comp_runit_term_dup:
assumes "ide a"
shows "\<r>[a] \<cdot> (a \<otimes> \<t>[a]) \<cdot> \<d>[a] = a"
proof -
have "\<guillemotleft>\<t>[a] : a \<rightarrow> \<one>\<guillemotright>"
using assms trm_in_hom by blast
hence "\<r>[a] \<cdot> (a \<otimes> \<t>[a]) = a \<cdot> \<pp>\<^sub>1[a, a]"
using assms by auto
thus ?thesis
using assms
by (metis comp_ide_arr pr_dup(2) ide_char comp_assoc seqI)
qed
lemma comp_proj_assoc:
assumes "ide a0" and "ide a1" and "ide a2"
shows "\<pp>\<^sub>1[a0, a1 \<otimes> a2] \<cdot> \<a>[a0, a1, a2] = \<pp>\<^sub>1[a0, a1] \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2]"
and "\<pp>\<^sub>0[a0, a1 \<otimes> a2] \<cdot> \<a>[a0, a1, a2] = \<langle>\<pp>\<^sub>0[a0, a1] \<cdot> \<pp>\<^sub>1[a0 \<otimes> a1, a2], \<pp>\<^sub>0[a0 \<otimes> a1, a2]\<rangle>"
using assms assoc_def by auto
lemma dup_coassoc:
assumes "ide a"
shows "\<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a] = (a \<otimes> \<d>[a]) \<cdot> \<d>[a]"
proof (intro pr_joint_monic
[of a "a \<otimes> a" "\<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a]" "(a \<otimes> \<d>[a]) \<cdot> \<d>[a]"])
show "seq \<pp>\<^sub>0[a, a \<otimes> a] (\<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a])"
using assms by simp
show "\<pp>\<^sub>0[a, a \<otimes> a] \<cdot> \<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a] = \<pp>\<^sub>0[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a]) \<cdot> \<d>[a]"
proof -
have "\<pp>\<^sub>0[a, a \<otimes> a] \<cdot> \<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a] =
((\<pp>\<^sub>0[a, a \<otimes> a] \<cdot> \<a>[a, a, a]) \<cdot> (\<d>[a] \<otimes> a)) \<cdot> \<d>[a]"
using comp_assoc by simp
also have "... = \<langle>((\<pp>\<^sub>0[a, a] \<cdot> \<pp>\<^sub>1[a \<otimes> a, a]) \<cdot> (\<d>[a] \<otimes> a)) \<cdot> \<d>[a], (a \<cdot> \<pp>\<^sub>0[a, a]) \<cdot> \<d>[a]\<rangle>"
using assms assoc_def by simp
also have "... = \<d>[a]"
using assms comp_assoc by simp
also have "... = (\<pp>\<^sub>0[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a])) \<cdot> \<d>[a]"
using assms assoc_def comp_assoc by simp
also have "... = \<pp>\<^sub>0[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a]) \<cdot> \<d>[a]"
using comp_assoc by simp
finally show ?thesis by blast
qed
show "\<pp>\<^sub>1[a, a \<otimes> a] \<cdot> \<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a] = \<pp>\<^sub>1[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a]) \<cdot> \<d>[a]"
proof -
have "\<pp>\<^sub>1[a, a \<otimes> a] \<cdot> \<a>[a, a, a] \<cdot> (\<d>[a] \<otimes> a) \<cdot> \<d>[a] =
((\<pp>\<^sub>1[a, a \<otimes> a] \<cdot> \<a>[a, a, a]) \<cdot> (\<d>[a] \<otimes> a)) \<cdot> \<d>[a]"
using comp_assoc by simp
also have "... = ((\<pp>\<^sub>1[a, a] \<cdot> \<pp>\<^sub>1[a \<otimes> a, a]) \<cdot> (\<d>[a] \<otimes> a)) \<cdot> \<d>[a]"
using assms assoc_def by simp
also have "... = a"
using assms comp_assoc by simp
also have "... = (a \<cdot> \<pp>\<^sub>1[a, a]) \<cdot> \<d>[a]"
using assms comp_assoc by simp
also have "... = (\<pp>\<^sub>1[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a])) \<cdot> \<d>[a]"
using assms by simp
also have "... = \<pp>\<^sub>1[a, a \<otimes> a] \<cdot> (a \<otimes> \<d>[a]) \<cdot> \<d>[a]"
using comp_assoc by simp
finally show ?thesis by blast
qed
qed
lemma comp_assoc_tuple:
assumes "\<guillemotleft>f0 : a \<rightarrow> b0\<guillemotright>" and "\<guillemotleft>f1 : a \<rightarrow> b1\<guillemotright>" and "\<guillemotleft>f2 : a \<rightarrow> b2\<guillemotright>"
shows "\<a>[b0, b1, b2] \<cdot> \<langle>\<langle>f0, f1\<rangle>, f2\<rangle> = \<langle>f0, \<langle>f1, f2\<rangle>\<rangle>"
and "\<a>\<^sup>-\<^sup>1[b0, b1, b2] \<cdot> \<langle>f0, \<langle>f1, f2\<rangle>\<rangle> = \<langle>\<langle>f0, f1\<rangle>, f2\<rangle>"
using assms assoc_def assoc'_def comp_assoc by fastforce+
lemma dup_tensor:
assumes "ide a" and "ide b"
shows "\<d>[a \<otimes> b] = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b])"
proof (intro pr_joint_monic [of "a \<otimes> b" "a \<otimes> b" "\<d>[a \<otimes> b]"])
show "seq \<pp>\<^sub>0[a \<otimes> b, a \<otimes> b] (\<d>[a \<otimes> b])"
using assms by simp
have 1: "\<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b]) =
\<langle>a \<otimes> b, a \<otimes> b\<rangle>"
proof -
have "\<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b])
= \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<langle>\<pp>\<^sub>1[a, b], \<d>[b] \<cdot> \<pp>\<^sub>0[a, b]\<rangle>\<rangle>"
proof -
have "\<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b]) = \<langle>\<pp>\<^sub>1[a, b], \<langle>\<pp>\<^sub>1[a, b], \<d>[b] \<cdot> \<pp>\<^sub>0[a, b]\<rangle>\<rangle>"
using assms assoc_def comp_assoc pr_naturality comp_cod_arr by simp
thus ?thesis by presburger
qed
also have "... = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot>
\<langle>a \<cdot> a \<cdot> a \<cdot> \<pp>\<^sub>1[a, b], \<a>[b, a, b] \<cdot> (\<s>[a, b] \<cdot> (a \<otimes> b) \<otimes> b) \<cdot>
\<a>\<^sup>-\<^sup>1[a, b, b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<d>[b \<cdot> \<pp>\<^sub>0[a, b]]\<rangle>\<rangle>"
using assms prod_tuple by simp
also have "... = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot>
\<langle>\<pp>\<^sub>1[a, b], \<a>[b, a, b] \<cdot> (\<s>[a, b] \<otimes> b) \<cdot> \<a>\<^sup>-\<^sup>1[a, b, b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<d>[\<pp>\<^sub>0[a, b]]\<rangle>\<rangle>"
proof -
have "a \<cdot> a \<cdot> a \<cdot> \<pp>\<^sub>1[a, b] = \<pp>\<^sub>1[a, b]"
using assms comp_cod_arr by simp
moreover have "b \<cdot> \<pp>\<^sub>0[a, b] = \<pp>\<^sub>0[a, b]"
using assms comp_cod_arr by simp
moreover have "\<s>[a, b] \<cdot> (a \<otimes> b) \<otimes> b = \<s>[a, b] \<otimes> b"
using assms comp_arr_dom by simp
ultimately show ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<a>[b, a, b] \<cdot> (\<s>[a, b] \<otimes> b) \<cdot>
\<langle>\<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle>, \<pp>\<^sub>0[a, b]\<rangle>\<rangle>"
proof -
have "\<a>\<^sup>-\<^sup>1[a, b, b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<d>[\<pp>\<^sub>0[a, b]]\<rangle> = \<langle>\<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle>, \<pp>\<^sub>0[a, b]\<rangle>"
using assms comp_assoc_tuple(2) by blast
thus ?thesis by simp
qed
also have "... = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<a>[b, a, b] \<cdot> \<langle>\<s>[a, b], \<pp>\<^sub>0[a, b]\<rangle>\<rangle>"
using assms prod_tuple comp_arr_dom comp_cod_arr by simp
also have "... = \<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> \<langle>\<pp>\<^sub>1[a, b], \<langle>\<pp>\<^sub>0[a, b], \<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle>\<rangle>\<rangle>"
using assms comp_assoc_tuple(1)
by (metis sym_def pr_in_hom)
also have "... = \<langle>\<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle>, \<langle>\<pp>\<^sub>1[a, b], \<pp>\<^sub>0[a, b]\<rangle>\<rangle>"
using assms comp_assoc_tuple(2) by force
also have "... = \<d>[a \<otimes> b]"
using assms by simp
finally show ?thesis by simp
qed
show "\<pp>\<^sub>0[a \<otimes> b, a \<otimes> b] \<cdot> \<d>[a \<otimes> b]
= \<pp>\<^sub>0[a \<otimes> b, a \<otimes> b] \<cdot>
\<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b])"
using assms 1 by force
show "\<pp>\<^sub>1[a \<otimes> b, a \<otimes> b] \<cdot> \<d>[a \<otimes> b]
= \<pp>\<^sub>1[a \<otimes> b, a \<otimes> b] \<cdot>
\<a>\<^sup>-\<^sup>1[a, b, a \<otimes> b] \<cdot> (a \<otimes> \<a>[b, a, b]) \<cdot> (a \<otimes> \<sigma> (a, b) \<otimes> b) \<cdot>
(a \<otimes> \<a>\<^sup>-\<^sup>1[a, b, b]) \<cdot> \<a>[a, a, b \<otimes> b] \<cdot> (\<d>[a] \<otimes> \<d>[b])"
using assms 1 by force
qed
(* TODO: Not sure if the remaining facts are useful. *)
lemma \<iota>_eq_trm:
shows "\<iota> = \<t>[\<one> \<otimes> \<one>]"
proof (intro terminal_arr_unique)
show "par \<iota> \<t>[\<one> \<otimes> \<one>]"
by (simp add: ide_one trm_one trm_tensor)
show "terminal_arr \<t>[\<one> \<otimes> \<one>]"
using ide_one \<iota>_is_terminal_arr \<open>par \<iota> \<t>[\<one> \<otimes> \<one>]\<close> by auto
show "terminal_arr \<iota>"
using \<iota>_is_terminal_arr by blast
qed
lemma terminal_tensor_one_one:
shows "terminal (\<one> \<otimes> \<one>)"
proof
show "ide (\<one> \<otimes> \<one>)"
using ide_one by simp
show "\<And>a. ide a \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>"
proof -
fix a
assume a: "ide a"
show "\<exists>!f. \<guillemotleft>f : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>"
proof
show "\<guillemotleft>inv \<iota> \<cdot> \<t>[a] : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>"
using a ide_one inverse_arrows_\<iota> inverse_unique trm_in_hom by fastforce
show "\<And>f. \<guillemotleft>f : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright> \<Longrightarrow> f = inv \<iota> \<cdot> \<t>[a]"
proof -
fix f
assume f: "\<guillemotleft>f : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>"
have "\<iota> \<cdot> f = \<t>[a]"
proof (intro terminal_arr_unique)
show "par (\<iota> \<cdot> f) \<t>[a]"
using a f
by (metis \<iota>_is_iso \<iota>_is_terminal_arr \<open>\<guillemotleft>inv \<iota> \<cdot> \<t>[a] : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>\<close>
cod_comp dom_comp dom_inv ide_one in_homE pr_simps(2-3) seqE seqI)
show "terminal_arr (\<iota> \<cdot> f)"
using a f \<iota>_is_terminal_arr cod_comp by force
show "terminal_arr \<t>[a]"
using a \<open>par (\<iota> \<cdot> f) \<t>[a]\<close> \<open>terminal_arr (\<iota> \<cdot> f)\<close> by auto
qed
thus "f = inv \<iota> \<cdot> \<t>[a]"
using a f \<iota>_is_iso invert_side_of_triangle(1)
\<open>\<guillemotleft>inv \<iota> \<cdot> \<t>[a] : a \<rightarrow> \<one> \<otimes> \<one>\<guillemotright>\<close>
by blast
qed
qed
qed
qed
end
section "Category with Finite Products"
text \<open>
In this last section, we show that the notion ``cartesian category'', which we defined
to be a category with binary products and terminal object, coincides with the notion
``category with finite products''. Due to the inability to quantify over types in HOL,
we content ourselves with defining the latter notion as "has \<open>I\<close>-indexed products
for every finite set \<open>I\<close> of natural numbers." We can transfer this property to finite
sets at other types using the fact that products are preserved under bijections of
the index sets.
\<close>
locale category_with_finite_products =
category C
for C :: "'c comp" +
assumes has_finite_products: "finite (I :: nat set) \<Longrightarrow> has_products I"
begin
lemma has_finite_products':
assumes "I \<noteq> UNIV"
shows "finite I \<Longrightarrow> has_products I"
proof -
assume I: "finite I"
obtain n \<phi> where \<phi>: "bij_betw \<phi> {k. k < (n :: nat)} I"
using I finite_imp_nat_seg_image_inj_on inj_on_imp_bij_betw by fastforce
show "has_products I"
using assms(1) \<phi> has_finite_products has_products_preserved_by_bijection
category_with_finite_products.has_finite_products
by blast
qed
end
lemma (in category) has_binary_products_if:
assumes "has_products ({0, 1} :: nat set)"
shows "has_binary_products"
proof (unfold has_binary_products_def)
show "\<forall>a0 a1. ide a0 \<and> ide a1 \<longrightarrow> (\<exists>p0 p1. has_as_binary_product a0 a1 p0 p1)"
proof (intro allI impI)
fix a0 a1
assume 1: "ide a0 \<and> ide a1"
show "\<exists>p0 p1. has_as_binary_product a0 a1 p0 p1"
proof -
interpret J: binary_product_shape
by unfold_locales
interpret D: binary_product_diagram C a0 a1
using 1 by unfold_locales auto
interpret discrete_diagram J.comp C D.map
using J.is_discrete
by unfold_locales auto
show "\<exists>p0 p1. has_as_binary_product a0 a1 p0 p1"
proof (unfold has_as_binary_product_def)
text \<open>
Here we have to work around the fact that \<open>has_finite_products\<close> is defined
in terms of @{typ "nat set"}, whereas \<open>has_as_binary_product\<close> is defined
in terms of \<open>J.arr set\<close>.
\<close>
let ?\<phi> = "(\<lambda>x :: nat. if x = 0 then J.FF else J.TT)"
let ?\<psi> = "\<lambda>j. if j = J.FF then 0 else 1"
have "bij_betw ?\<phi> ({0, 1} :: nat set) {J.FF, J.TT}"
using bij_betwI [of ?\<phi> "{0, 1} :: nat set" "{J.FF, J.TT}" ?\<psi>] by fastforce
hence "has_products {J.FF, J.TT}"
using assms has_products_def [of "{J.FF, J.TT}"]
has_products_preserved_by_bijection
[of "{0, 1} :: nat set" ?\<phi> "{J.FF, J.TT}"]
by blast
hence "\<exists>a. has_as_product J.comp D.map a"
using has_products_def [of "{J.FF, J.TT}"]
discrete_diagram_axioms J.arr_char
by blast
hence "\<exists>a \<pi>. product_cone J.comp C D.map a \<pi>"
using has_as_product_def by blast
hence 2: "\<exists>a \<pi>. D.limit_cone a \<pi>"
unfolding product_cone_def by simp
obtain a \<pi> where \<pi>: "D.limit_cone a \<pi>"
using 2 by auto
interpret \<pi>: limit_cone J.comp C D.map a \<pi>
using \<pi> by auto
have "\<pi> = D.mkCone (\<pi> J.FF) (\<pi> J.TT)"
proof -
have "\<And>a. J.ide a \<Longrightarrow> \<pi> a = D.mkCone (\<pi> J.FF) (\<pi> J.TT) a"
using D.mkCone_def J.ide_char by auto
moreover have "a = dom (\<pi> J.FF)"
by simp
moreover have "D.cone a (D.mkCone (\<pi> (J.MkIde False)) (\<pi> (J.MkIde True)))"
using 1 D.cone_mkCone [of "\<pi> J.FF" "\<pi> J.TT"] by auto
ultimately show ?thesis
using D.mkCone_def \<pi>.natural_transformation_axioms
D.cone_mkCone [of "\<pi> J.FF" "\<pi> J.TT"]
NaturalTransformation.eqI
[of "J.comp" C \<pi>.A.map "D.map" \<pi> "D.mkCone (\<pi> J.FF) (\<pi> J.TT)"]
cone_def [of J.comp C D.map a "D.mkCone (\<pi> J.FF) (\<pi> J.TT)"] J.ide_char
by blast
qed
hence "D.limit_cone (dom (\<pi> J.FF)) (D.mkCone (\<pi> J.FF) (\<pi> J.TT))"
using \<pi>.limit_cone_axioms by simp
thus "\<exists>p0 p1. ide a0 \<and> ide a1 \<and> D.has_as_binary_product p0 p1"
using 1 by blast
qed
qed
qed
qed
sublocale category_with_finite_products \<subseteq> category_with_binary_products C
using has_binary_products_if has_finite_products
by (unfold_locales, unfold has_binary_products_def) simp
proposition (in category_with_finite_products) is_category_with_binary_products:
shows "category_with_binary_products C"
..
sublocale category_with_finite_products \<subseteq> category_with_terminal_object C
proof
interpret J: discrete_category "{} :: nat set"
by unfold_locales auto
interpret D: empty_diagram J.comp C "\<lambda>j. null"
by unfold_locales auto
interpret D: discrete_diagram J.comp C "\<lambda>j. null"
using J.is_discrete by unfold_locales auto
have "\<And>a. D.has_as_limit a \<longleftrightarrow> has_as_product J.comp (\<lambda>j. null) a"
using product_cone_def J.category_axioms category_axioms D.discrete_diagram_axioms
has_as_product_def product_cone_def
by metis
moreover have "\<exists>a. has_as_product J.comp (\<lambda>j. null) a"
using has_finite_products [of "{} :: nat set"] has_products_def [of "{} :: nat set"]
D.discrete_diagram_axioms
by blast
ultimately have "\<exists>a. D.has_as_limit a" by blast
thus "\<exists>a. terminal a" using D.has_as_limit_iff_terminal by blast
qed
proposition (in category_with_finite_products) is_category_with_terminal_object:
shows "category_with_terminal_object C"
..
sublocale category_with_finite_products \<subseteq> cartesian_category ..
proposition (in category_with_finite_products) is_cartesian_category:
shows "cartesian_category C"
..
context category
begin
lemma binary_product_of_products_is_product:
assumes "has_as_product J0 D0 a0" and "has_as_product J1 D1 a1"
and "has_as_binary_product a0 a1 p0 p1"
and "Collect (partial_magma.arr J0) \<inter> Collect (partial_magma.arr J1) = {}"
and "partial_magma.null J0 = partial_magma.null J1"
shows "has_as_product
(discrete_category.comp
(Collect (partial_magma.arr J0) \<union> Collect (partial_magma.arr J1))
(partial_magma.null J0))
(\<lambda>i. if i \<in> Collect (partial_magma.arr J0) then D0 i
else if i \<in> Collect (partial_magma.arr J1) then D1 i
else null)
(dom p0)"
proof -
obtain \<pi>0 where \<pi>0: "product_cone J0 (\<cdot>) D0 a0 \<pi>0"
using assms(1) has_as_product_def by blast
obtain \<pi>1 where \<pi>1: "product_cone J1 (\<cdot>) D1 a1 \<pi>1"
using assms(2) has_as_product_def by blast
interpret J0: category J0
using \<pi>0 product_cone.axioms(1) by metis
interpret J1: category J1
using \<pi>1 product_cone.axioms(1) by metis
interpret D0: discrete_diagram J0 C D0
using \<pi>0 product_cone.axioms(3) by metis
interpret D1: discrete_diagram J1 C D1
using \<pi>1 product_cone.axioms(3) by metis
interpret \<pi>0: product_cone J0 C D0 a0 \<pi>0
using \<pi>0 by auto
interpret \<pi>1: product_cone J1 C D1 a1 \<pi>1
using \<pi>1 by auto
interpret J: discrete_category \<open>Collect J0.arr \<union> Collect J1.arr\<close> J0.null
using J0.not_arr_null assms(5) by unfold_locales auto
interpret X: binary_product_shape .
interpret a0xa1: binary_product_diagram C a0 a1
using assms(3) has_as_binary_product_def
by (simp add: binary_product_diagram.intro binary_product_diagram_axioms.intro
category_axioms)
have p0p1: "a0xa1.has_as_binary_product p0 p1"
using assms(3) has_as_binary_product_def [of a0 a1 p0 p1] by simp
let ?D = "(\<lambda>i. if i \<in> Collect J0.arr then D0 i
else if i \<in> Collect J1.arr then D1 i
else null)"
let ?a = "dom p0"
let ?\<pi> = "\<lambda>i. if i \<in> Collect J0.arr then \<pi>0 i \<cdot> p0
else if i \<in> Collect J1.arr then \<pi>1 i \<cdot> p1
else null"
let ?p0p1 = "a0xa1.mkCone p0 p1"
interpret p0p1: limit_cone X.comp C a0xa1.map ?a ?p0p1
using p0p1 by simp
have a: "ide ?a"
using p0p1.ide_apex by simp
have p0: "\<guillemotleft>p0 : ?a \<rightarrow> a0\<guillemotright>"
using a0xa1.mkCone_def p0p1.preserves_hom [of X.FF X.FF X.FF] X.ide_char X.ide_in_hom
by auto
have p1: "\<guillemotleft>p1 : ?a \<rightarrow> a1\<guillemotright>"
using a0xa1.mkCone_def p0p1.preserves_hom [of X.TT X.TT X.TT] X.ide_char X.ide_in_hom
by auto
interpret D: discrete_diagram J.comp C ?D
using assms J.arr_char J.dom_char J.cod_char J.is_discrete D0.is_discrete D1.is_discrete
J.cod_comp J.seq_char
by unfold_locales auto
interpret A: constant_functor J.comp C ?a
using p0p1.ide_apex by unfold_locales simp
interpret \<pi>: natural_transformation J.comp C A.map ?D ?\<pi>
proof
fix j
show "\<not> J.arr j \<Longrightarrow> ?\<pi> j = null"
by simp
assume j: "J.arr j"
have \<pi>0j: "J0.arr j \<Longrightarrow> \<guillemotleft>\<pi>0 j : a0 \<rightarrow> D0 j\<guillemotright>"
using D0.is_discrete by auto
have \<pi>1j: "J1.arr j \<Longrightarrow> \<guillemotleft>\<pi>1 j : a1 \<rightarrow> D1 j\<guillemotright>"
using D1.is_discrete by auto
show "dom (?\<pi> j) = A.map (J.dom j)"
using j J.arr_char p0 p1 \<pi>0j \<pi>1j
by fastforce
show "cod (?\<pi> j) = ?D (J.cod j)"
using j J.arr_char p0 p1 \<pi>0j \<pi>1j
by fastforce
show "?D j \<cdot> ?\<pi> (J.dom j) = ?\<pi> j"
proof -
have 0: "J0.arr j \<Longrightarrow> D0 j \<cdot> \<pi>0 j \<cdot> p0 = \<pi>0 j \<cdot> p0"
proof -
have "J0.arr j \<Longrightarrow> (D0 j \<cdot> \<pi>0 j) \<cdot> p0 = \<pi>0 j \<cdot> p0"
using p0 \<pi>0.is_natural_1 \<pi>0.is_natural_2 D0.is_discrete by simp
thus "J0.arr j \<Longrightarrow> D0 j \<cdot> \<pi>0 j \<cdot> p0 = \<pi>0 j \<cdot> p0"
using comp_assoc by simp
qed
have 1: "J1.arr j \<Longrightarrow> D1 j \<cdot> \<pi>1 j \<cdot> p1 = \<pi>1 j \<cdot> p1"
proof -
have "J1.arr j \<Longrightarrow> (D1 j \<cdot> \<pi>1 j) \<cdot> p1 = \<pi>1 j \<cdot> p1"
using p1 \<pi>1.is_natural_1 \<pi>1.is_natural_2 D1.is_discrete by simp
thus "J1.arr j \<Longrightarrow> D1 j \<cdot> \<pi>1 j \<cdot> p1 = \<pi>1 j \<cdot> p1"
using comp_assoc by simp
qed
show ?thesis
using 0 1 by auto
qed
show "?\<pi> (J.cod j) \<cdot> A.map j = ?\<pi> j"
using j comp_arr_dom p0 p1 comp_assoc by auto
qed
interpret \<pi>: cone J.comp C ?D ?a ?\<pi> ..
interpret \<pi>: product_cone J.comp C ?D ?a ?\<pi>
proof
show "\<And>a' \<chi>'. D.cone a' \<chi>' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> ?a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
proof -
fix a' \<chi>'
assume \<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J.comp C ?D a' \<chi>'
using \<chi>' by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> ?a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
proof
let ?\<chi>0' = "\<lambda>i. if i \<in> Collect J0.arr then \<chi>' i else null"
let ?\<chi>1' = "\<lambda>i. if i \<in> Collect J1.arr then \<chi>' i else null"
have 0: "\<And>i. i \<in> Collect J0.arr \<Longrightarrow> \<chi>' i \<in> hom a' (D0 i)"
using J.arr_char by auto
have 1: "\<And>i. i \<in> Collect J1.arr \<Longrightarrow> \<chi>' i \<in> hom a' (D1 i)"
- using J.arr_char `Collect J0.arr \<inter> Collect J1.arr = {}` by force
+ using J.arr_char \<open>Collect J0.arr \<inter> Collect J1.arr = {}\<close> by force
interpret A0': constant_functor J0 C a'
apply unfold_locales using \<chi>'.ide_apex by auto
interpret A1': constant_functor J1 C a'
apply unfold_locales using \<chi>'.ide_apex by auto
interpret \<chi>0': cone J0 C D0 a' ?\<chi>0'
proof (unfold_locales)
fix j
show "\<not> J0.arr j \<Longrightarrow> (if j \<in> Collect J0.arr then \<chi>' j else null) = null"
by simp
assume j: "J0.arr j"
show 0: "dom (?\<chi>0' j) = A0'.map (J0.dom j)"
using j by simp
show 1: "cod (?\<chi>0' j) = D0 (J0.cod j)"
using j J.arr_char J.cod_char D0.is_discrete by simp
show "D0 j \<cdot> (?\<chi>0' (J0.dom j)) = ?\<chi>0' j"
using 1 j J.arr_char D0.is_discrete comp_cod_arr by simp
show "?\<chi>0' (J0.cod j) \<cdot> A0'.map j = ?\<chi>0' j"
using 0 j J.arr_char D0.is_discrete comp_arr_dom by simp
qed
interpret \<chi>1': cone J1 C D1 a' ?\<chi>1'
proof (unfold_locales)
fix j
show "\<not> J1.arr j \<Longrightarrow> (if j \<in> Collect J1.arr then \<chi>' j else null) = null"
by simp
assume j: "J1.arr j"
show 0: "dom (?\<chi>1' j) = A1'.map (J1.dom j)"
using j by simp
show 1: "cod (?\<chi>1' j) = D1 (J1.cod j)"
using assms(4) j J.arr_char J.cod_char D1.is_discrete by auto
show "D1 j \<cdot> (?\<chi>1' (J1.dom j)) = ?\<chi>1' j"
using 1 j J.arr_char D1.is_discrete comp_cod_arr by simp
show "?\<chi>1' (J1.cod j) \<cdot> A1'.map j = ?\<chi>1' j"
using 0 j J.arr_char D1.is_discrete comp_arr_dom by simp
qed
define f0 where "f0 = \<pi>0.induced_arrow a' ?\<chi>0'"
define f1 where "f1 = \<pi>1.induced_arrow a' ?\<chi>1'"
have f0: "\<guillemotleft>f0 : a' \<rightarrow> a0\<guillemotright>"
using f0_def \<pi>0.induced_arrowI \<chi>0'.cone_axioms by simp
have f1: "\<guillemotleft>f1 : a' \<rightarrow> a1\<guillemotright>"
using f1_def \<pi>1.induced_arrowI \<chi>1'.cone_axioms by simp
have 2: "a0xa1.is_rendered_commutative_by f0 f1"
using f0 f1 by auto
interpret p0p1: binary_product_cone C a0 a1 p0 p1 ..
interpret f0f1: cone X.comp C a0xa1.map a' \<open>a0xa1.mkCone f0 f1\<close>
using 2 f0 f1 a0xa1.cone_mkCone [of f0 f1] by auto
define f where "f = p0p1.induced_arrow a' (a0xa1.mkCone f0 f1)"
have f: "\<guillemotleft>f : a' \<rightarrow> ?a\<guillemotright>"
using f_def 2 f0 f1 p0p1.induced_arrowI'(1) by auto
moreover have \<chi>': "D.cones_map f ?\<pi> = \<chi>'"
proof
fix j
show "D.cones_map f ?\<pi> j = \<chi>' j"
proof (cases "J0.arr j", cases "J1.arr j")
show "\<lbrakk>J0.arr j; J1.arr j\<rbrakk> \<Longrightarrow> D.cones_map f ?\<pi> j = \<chi>' j"
using assms(4) by auto
show "\<lbrakk>J0.arr j; \<not> J1.arr j\<rbrakk> \<Longrightarrow> D.cones_map f ?\<pi> j = \<chi>' j"
proof -
assume J0: "J0.arr j" and J1: "\<not> J1.arr j"
have "D.cones_map f ?\<pi> j = (\<pi>0 j \<cdot> p0) \<cdot> f"
using f J0 J1 \<pi>.cone_axioms by auto
also have "... = \<pi>0 j \<cdot> p0 \<cdot> f"
using comp_assoc by simp
also have "... = \<pi>0 j \<cdot> f0"
using 2 f0 f1 f_def p0p1.induced_arrowI' by auto
also have "... = \<chi>' j"
proof -
have "\<pi>0 j \<cdot> f0 = \<pi>0 j \<cdot> \<pi>0.induced_arrow' a' \<chi>'"
unfolding f0_def by simp
also have "... = (\<lambda>j. if J0.arr j then
\<pi>0 j \<cdot> \<pi>0.induced_arrow a'
(\<lambda>i. if i \<in> Collect J0.arr then \<chi>' i else null)
else null) j"
using J0 by simp
also have "... = D0.mkCone \<chi>' j"
proof -
have "(\<lambda>j. if J0.arr j then
\<pi>0 j \<cdot> \<pi>0.induced_arrow a'
(\<lambda>i. if i \<in> Collect J0.arr then \<chi>' i else null)
else null) =
D0.mkCone \<chi>'"
using f0 f0_def \<pi>0.induced_arrowI(2) [of ?\<chi>0' a'] J0
D0.mkCone_cone \<chi>0'.cone_axioms \<pi>0.cone_axioms J0
by auto
thus ?thesis by meson
qed
also have "... = \<chi>' j"
using J0 by simp
finally show ?thesis by blast
qed
finally show ?thesis by simp
qed
show "\<not> J0.arr j \<Longrightarrow> D.cones_map f ?\<pi> j = \<chi>' j"
proof (cases "J1.arr j")
show "\<lbrakk>\<not> J0.arr j; \<not> J1.arr j\<rbrakk> \<Longrightarrow> D.cones_map f ?\<pi> j = \<chi>' j"
using f \<pi>.cone_axioms \<chi>'.is_extensional by auto
show "\<lbrakk>\<not> J0.arr j; J1.arr j\<rbrakk> \<Longrightarrow> D.cones_map f ?\<pi> j = \<chi>' j"
proof -
assume J0: "\<not> J0.arr j" and J1: "J1.arr j"
have "D.cones_map f ?\<pi> j = (\<pi>1 j \<cdot> p1) \<cdot> f"
using J0 J1 f \<pi>.cone_axioms by auto
also have "... = \<pi>1 j \<cdot> p1 \<cdot> f"
using comp_assoc by simp
also have "... = \<pi>1 j \<cdot> f1"
using 2 f0 f1 f_def p0p1.induced_arrowI' by auto
also have "... = \<chi>' j"
proof -
have "\<pi>1 j \<cdot> f1 = \<pi>1 j \<cdot> \<pi>1.induced_arrow' a' \<chi>'"
unfolding f1_def by simp
also have "... = (\<lambda>j. if J1.arr j then
\<pi>1 j \<cdot> \<pi>1.induced_arrow a'
(\<lambda>i. if i \<in> Collect J1.arr
then \<chi>' i else null)
else null) j"
using J1 by simp
also have "... = D1.mkCone \<chi>' j"
proof -
have "(\<lambda>j. if J1.arr j then
\<pi>1 j \<cdot> \<pi>1.induced_arrow a'
(\<lambda>i. if i \<in> Collect J1.arr then \<chi>' i else null)
else null) =
D1.mkCone \<chi>'"
using f1 f1_def \<pi>1.induced_arrowI(2) [of ?\<chi>1' a'] J1
D1.mkCone_cone [of a' \<chi>'] \<chi>1'.cone_axioms \<pi>1.cone_axioms J1
by auto
thus ?thesis by meson
qed
also have "... = \<chi>' j"
using J1 by simp
finally show ?thesis by blast
qed
finally show ?thesis by simp
qed
qed
qed
qed
ultimately show "\<guillemotleft>f : a' \<rightarrow> ?a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'" by blast
show "\<And>f'. \<guillemotleft>f' : a' \<rightarrow> ?a\<guillemotright> \<and> D.cones_map f' ?\<pi> = \<chi>' \<Longrightarrow> f' = f"
proof -
fix f'
assume f': "\<guillemotleft>f' : a' \<rightarrow> ?a\<guillemotright> \<and> D.cones_map f' ?\<pi> = \<chi>'"
let ?f0' = "p0 \<cdot> f'"
let ?f1' = "p1 \<cdot> f'"
have 1: "a0xa1.is_rendered_commutative_by ?f0' ?f1'"
using f' p0 p1 p0p1.renders_commutative seqI' by auto
have f0': "\<guillemotleft>?f0' : a' \<rightarrow> a0\<guillemotright>"
using f' p0 by auto
have f1': "\<guillemotleft>?f1' : a' \<rightarrow> a1\<guillemotright>"
using f' p1 by auto
have "p0 \<cdot> f = p0 \<cdot> f'"
proof -
have "D0.cones_map (p0 \<cdot> f) \<pi>0 = ?\<chi>0'"
using f p0 \<pi>0.cone_axioms \<chi>' \<pi>.cone_axioms comp_assoc assms(4) seqI'
by fastforce
moreover have "D0.cones_map (p0 \<cdot> f') \<pi>0 = ?\<chi>0'"
using f' p0 \<pi>0.cone_axioms \<pi>.cone_axioms comp_assoc assms(4) seqI'
by fastforce
moreover have "p0 \<cdot> f = f0"
using 2 f0 f_def p0p1.induced_arrowI'(2) by blast
ultimately show ?thesis
using f0 f0' \<chi>0'.cone_axioms \<pi>0.is_universal [of a'] by auto
qed
moreover have "p1 \<cdot> f = p1 \<cdot> f'"
proof -
have "D1.cones_map (p1 \<cdot> f) \<pi>1 = ?\<chi>1'"
proof
fix j
show "D1.cones_map (p1 \<cdot> f) \<pi>1 j = ?\<chi>1' j"
using f p1 \<pi>1.cone_axioms \<chi>' \<pi>.cone_axioms comp_assoc assms(4) seqI'
apply auto
by auto
qed
moreover have "D1.cones_map (p1 \<cdot> f') \<pi>1 = ?\<chi>1'"
proof
fix j
show "D1.cones_map (p1 \<cdot> f') \<pi>1 j = ?\<chi>1' j"
using f' p1 \<pi>1.cone_axioms \<pi>.cone_axioms comp_assoc assms(4) seqI'
apply auto
by auto
qed
moreover have "p1 \<cdot> f = f1"
using 2 f1 f_def p0p1.induced_arrowI'(3) by blast
ultimately show ?thesis
using f1 f1' \<chi>1'.cone_axioms \<pi>1.is_universal [of a'] by auto
qed
ultimately show "f' = f"
using f f' p0p1.is_universal' [of a']
by (metis (no_types, lifting) "1" dom_comp in_homE p0p1.is_universal' p1 seqI')
qed
qed
qed
qed
show "has_as_product J.comp ?D ?a"
unfolding has_as_product_def
using \<pi>.product_cone_axioms by auto
qed
end
sublocale cartesian_category \<subseteq> category_with_finite_products
proof
obtain t where t: "terminal t" using has_terminal by blast
{ fix n :: nat
have "\<And>I :: nat set. finite I \<and> card I = n \<Longrightarrow> has_products I"
proof (induct n)
show "\<And>I :: nat set. finite I \<and> card I = 0 \<Longrightarrow> has_products I"
proof -
fix I :: "nat set"
assume "finite I \<and> card I = 0"
hence I: "I = {}" by force
thus "has_products I"
proof -
interpret J: discrete_category I 0
apply unfold_locales using I by auto
have "\<And>D. discrete_diagram J.comp C D \<Longrightarrow> \<exists>a. has_as_product J.comp D a"
proof -
fix D
assume D: "discrete_diagram J.comp C D"
interpret D: discrete_diagram J.comp C D using D by auto
interpret D: empty_diagram J.comp C D
using I J.arr_char by unfold_locales simp
have "has_as_product J.comp D t"
using t D.has_as_limit_iff_terminal has_as_product_def product_cone_def
J.category_axioms category_axioms D.discrete_diagram_axioms
by metis
thus "\<exists>a. has_as_product J.comp D a" by blast
qed
moreover have "I \<noteq> UNIV"
using I by blast
ultimately show ?thesis
using I has_products_def
by (metis category_with_terminal_object.has_terminal discrete_diagram.product_coneI
discrete_diagram_def empty_diagram.has_as_limit_iff_terminal empty_diagram.intro
empty_diagram_axioms.intro empty_iff has_as_product_def
is_category_with_terminal_object mem_Collect_eq)
qed
qed
show "\<And>n I :: nat set.
\<lbrakk> (\<And>I :: nat set. finite I \<and> card I = n \<Longrightarrow> has_products I);
finite I \<and> card I = Suc n \<rbrakk>
\<Longrightarrow> has_products I"
proof -
fix n :: nat
fix I :: "nat set"
assume IH: "\<And>I :: nat set. finite I \<and> card I = n \<Longrightarrow> has_products I"
assume I: "finite I \<and> card I = Suc n"
show "has_products I"
proof -
have "card I = 1 \<Longrightarrow> has_products I"
using I has_unary_products by blast
moreover have "card I \<noteq> 1 \<Longrightarrow> has_products I"
proof -
assume "card I \<noteq> 1"
hence cardI: "card I > 1" using I by simp
obtain i where i: "i \<in> I" using cardI by fastforce
let ?I0 = "{i}" and ?I1 = "I - {i}"
have 1: "I = ?I0 \<union> ?I1 \<and> ?I0 \<inter> ?I1 = {} \<and> card ?I0 = 1 \<and> card ?I1 = n"
using i I cardI by auto
show "has_products I"
proof (unfold has_products_def, intro conjI allI impI)
show "I \<noteq> UNIV"
using I by auto
fix J D
assume D: "discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I"
interpret D: discrete_diagram J C D
using D by simp
have Null: "D.J.null \<notin> ?I0 \<and> D.J.null \<notin> ?I1"
using D D.J.not_arr_null i by blast
interpret J0: discrete_category ?I0 D.J.null
using 1 Null D by unfold_locales auto
interpret J1: discrete_category ?I1 D.J.null
using Null by unfold_locales auto
interpret J0uJ1: discrete_category \<open>Collect J0.arr \<union> Collect J1.arr\<close> J0.null
using Null 1 J0.null_char J1.null_char by unfold_locales auto
interpret D0: discrete_diagram_from_map ?I0 C D D.J.null
using 1 J0.ide_char D.preserves_ide D D.is_discrete i by unfold_locales auto
interpret D1: discrete_diagram_from_map ?I1 C D D.J.null
using 1 J1.ide_char D.preserves_ide D D.is_discrete i by unfold_locales auto
obtain a0 where a0: "has_as_product J0.comp D0.map a0"
using 1 has_unary_products [of ?I0] has_products_def [of ?I0]
D0.discrete_diagram_axioms
by fastforce
obtain a1 where a1: "has_as_product J1.comp D1.map a1"
using 1 I IH [of ?I1] has_products_def [of ?I1] D1.discrete_diagram_axioms
by blast
have 2: "\<exists>p0 p1. has_as_binary_product a0 a1 p0 p1"
proof -
have "ide a0 \<and> ide a1"
using a0 a1 product_is_ide by auto
thus ?thesis
using a0 a1 has_binary_products has_binary_products_def by simp
qed
obtain p0 p1 where a: "has_as_binary_product a0 a1 p0 p1"
using 2 by auto
let ?a = "dom p0"
have "has_as_product J D ?a"
proof -
have "D = (\<lambda>j. if j \<in> Collect J0.arr then D0.map j
else if j \<in> Collect J1.arr then D1.map j
else null)"
proof
fix j
show "D j = (if j \<in> Collect J0.arr then D0.map j
else if j \<in> Collect J1.arr then D1.map j
else null)"
using 1 D0.map_def D1.map_def D.is_extensional D J0.arr_char J1.arr_char
by auto
qed
moreover have "J = J0uJ1.comp"
proof -
have "\<And>j j'. J j j' = J0uJ1.comp j j'"
proof -
fix j j'
show "J j j' = J0uJ1.comp j j'"
using D J0uJ1.arr_char J0.arr_char J1.arr_char D.is_discrete i
apply (cases "j \<in> ?I0", cases "j' \<in> ?I0")
apply simp_all
apply auto[1]
apply (metis D.J.comp_arr_ide D.J.comp_ide_arr D.J.ext D.J.seqE
D.is_discrete J0.null_char J0uJ1.null_char)
by (metis D.J.comp_arr_ide D.J.comp_ide_arr D.J.comp_ide_self
D.J.ext D.J.seqE D.is_discrete J0.null_char J0uJ1.null_char
mem_Collect_eq)
qed
thus ?thesis by blast
qed
moreover have "Collect J0.arr \<inter> Collect J1.arr = {}"
by auto
moreover have "J0.null = J1.null"
using J0.null_char J1.null_char by simp
ultimately show "has_as_product J D ?a"
using binary_product_of_products_is_product
[of J0.comp D0.map a0 J1.comp D1.map a1 p0 p1]
J0.arr_char J1.arr_char
1 a0 a1 a
by simp
qed
thus "\<exists>a. has_as_product J D a" by blast
qed
qed
ultimately show "has_products I" by blast
qed
qed
qed
}
hence 1: "\<And>n I :: nat set. finite I \<and> card I = n \<Longrightarrow> has_products I" by simp
thus "\<And>I :: nat set. finite I \<Longrightarrow> has_products I" by blast
qed
proposition (in cartesian_category) is_category_with_finite_products:
shows "category_with_finite_products C"
..
end
diff --git a/thys/Category3/Limit.thy b/thys/Category3/Limit.thy
--- a/thys/Category3/Limit.thy
+++ b/thys/Category3/Limit.thy
@@ -1,6472 +1,6472 @@
(* Title: Limit
Author: Eugene W. Stark <stark@cs.stonybrook.edu>, 2016
Maintainer: Eugene W. Stark <stark@cs.stonybrook.edu>
*)
chapter Limit
theory Limit
imports FreeCategory DiscreteCategory Adjunction
begin
text\<open>
This theory defines the notion of limit in terms of diagrams and cones and relates
it to the concept of a representation of a functor. The diagonal functor associated
with a diagram shape @{term J} is defined and it is shown that a right adjoint to
the diagonal functor gives limits of shape @{term J} and that a category has limits
of shape @{term J} if and only if the diagonal functor is a left adjoint functor.
Products and equalizers are defined as special cases of limits, and it is shown
that a category with equalizers has limits of shape @{term J} if it has products
indexed by the sets of objects and arrows of @{term J}.
The existence of limits in a set category is investigated, and it is shown that
every set category has equalizers and that a set category @{term S} has @{term I}-indexed
products if and only if the universe of @{term S} ``admits @{term I}-indexed tupling.''
The existence of limits in functor categories is also developed, showing that
limits in functor categories are ``determined pointwise'' and that a functor category
@{term "[A, B]"} has limits of shape @{term J} if @{term B} does.
Finally, it is shown that the Yoneda functor preserves limits.
This theory concerns itself only with limits; I have made no attempt to consider colimits.
Although it would be possible to rework the entire development in dual form,
it is possible that there is a more efficient way to dualize at least parts of it without
repeating all the work. This is something that deserves further thought.
\<close>
section "Representations of Functors"
text\<open>
A representation of a contravariant functor \<open>F: Cop \<rightarrow> S\<close>, where @{term S}
is a replete set category that is the target of a hom-functor for @{term C}, consists of
an object @{term a} of @{term C} and a natural isomorphism @{term "\<Phi>: Y a \<rightarrow> F"},
where \<open>Y: C \<rightarrow> [Cop, S]\<close> is the Yoneda functor.
\<close>
locale representation_of_functor =
C: category C +
Cop: dual_category C +
S: replete_set_category S +
F: "functor" Cop.comp S F +
Hom: hom_functor C S \<phi> +
Ya: yoneda_functor_fixed_object C S \<phi> a +
natural_isomorphism Cop.comp S \<open>Ya.Y a\<close> F \<Phi>
for C :: "'c comp" (infixr "\<cdot>" 55)
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and F :: "'c \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
begin
abbreviation Y where "Y \<equiv> Ya.Y"
abbreviation \<psi> where "\<psi> \<equiv> Hom.\<psi>"
end
text\<open>
Two representations of the same functor are uniquely isomorphic.
\<close>
locale two_representations_one_functor =
C: category C +
Cop: dual_category C +
S: replete_set_category S +
F: set_valued_functor Cop.comp S F +
yoneda_functor C S \<phi> +
Ya: yoneda_functor_fixed_object C S \<phi> a +
Ya': yoneda_functor_fixed_object C S \<phi> a' +
\<Phi>: representation_of_functor C S \<phi> F a \<Phi> +
\<Phi>': representation_of_functor C S \<phi> F a' \<Phi>'
for C :: "'c comp" (infixr "\<cdot>" 55)
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and F :: "'c \<Rightarrow> 's"
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
and a' :: 'c
and \<Phi>' :: "'c \<Rightarrow> 's"
begin
interpretation \<Psi>: inverse_transformation Cop.comp S \<open>Y a\<close> F \<Phi> ..
interpretation \<Psi>': inverse_transformation Cop.comp S \<open>Y a'\<close> F \<Phi>' ..
interpretation \<Phi>\<Psi>': vertical_composite Cop.comp S \<open>Y a\<close> F \<open>Y a'\<close> \<Phi> \<Psi>'.map ..
interpretation \<Phi>'\<Psi>: vertical_composite Cop.comp S \<open>Y a'\<close> F \<open>Y a\<close> \<Phi>' \<Psi>.map ..
lemma are_uniquely_isomorphic:
shows "\<exists>!\<phi>. \<guillemotleft>\<phi> : a \<rightarrow> a'\<guillemotright> \<and> C.iso \<phi> \<and> map \<phi> = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
proof -
have "natural_isomorphism Cop.comp S (Y a) F \<Phi>" ..
moreover have "natural_isomorphism Cop.comp S F (Y a') \<Psi>'.map" ..
ultimately have 1: "natural_isomorphism Cop.comp S (Y a) (Y a') \<Phi>\<Psi>'.map"
using NaturalTransformation.natural_isomorphisms_compose by blast
interpret \<Phi>\<Psi>': natural_isomorphism Cop.comp S \<open>Y a\<close> \<open>Y a'\<close> \<Phi>\<Psi>'.map
using 1 by auto
have "natural_isomorphism Cop.comp S (Y a') F \<Phi>'" ..
moreover have "natural_isomorphism Cop.comp S F (Y a) \<Psi>.map" ..
ultimately have 2: "natural_isomorphism Cop.comp S (Y a') (Y a) \<Phi>'\<Psi>.map"
using NaturalTransformation.natural_isomorphisms_compose by blast
interpret \<Phi>'\<Psi>: natural_isomorphism Cop.comp S \<open>Y a'\<close> \<open>Y a\<close> \<Phi>'\<Psi>.map
using 2 by auto
interpret \<Phi>\<Psi>'_\<Phi>'\<Psi>: inverse_transformations Cop.comp S \<open>Y a\<close> \<open>Y a'\<close> \<Phi>\<Psi>'.map \<Phi>'\<Psi>.map
proof
fix x
assume X: "Cop.ide x"
show "S.inverse_arrows (\<Phi>\<Psi>'.map x) (\<Phi>'\<Psi>.map x)"
proof
have 1: "S.arr (\<Phi>\<Psi>'.map x) \<and> \<Phi>\<Psi>'.map x = \<Psi>'.map x \<cdot>\<^sub>S \<Phi> x"
using X \<Phi>\<Psi>'.preserves_reflects_arr [of x]
by (simp add: \<Phi>\<Psi>'.map_simp_2)
have 2: "S.arr (\<Phi>'\<Psi>.map x) \<and> \<Phi>'\<Psi>.map x = \<Psi>.map x \<cdot>\<^sub>S \<Phi>' x"
using X \<Phi>'\<Psi>.preserves_reflects_arr [of x]
by (simp add: \<Phi>'\<Psi>.map_simp_1)
show "S.ide (\<Phi>\<Psi>'.map x \<cdot>\<^sub>S \<Phi>'\<Psi>.map x)"
using 1 2 X \<Psi>.is_natural_2 \<Psi>'.inverts_components \<Psi>.inverts_components
by (metis S.inverse_arrows_def S.inverse_arrows_compose)
show "S.ide (\<Phi>'\<Psi>.map x \<cdot>\<^sub>S \<Phi>\<Psi>'.map x)"
using 1 2 X \<Psi>'.inverts_components \<Psi>.inverts_components
by (metis S.inverse_arrows_def S.inverse_arrows_compose)
qed
qed
have "Cop_S.inverse_arrows (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)
(Cop_S.MkArr (Y a') (Y a) \<Phi>'\<Psi>.map)"
proof -
have Ya: "functor Cop.comp S (Y a)" ..
have Ya': "functor Cop.comp S (Y a')" ..
have \<Phi>\<Psi>': "natural_transformation Cop.comp S (Y a) (Y a') \<Phi>\<Psi>'.map" ..
have \<Phi>'\<Psi>: "natural_transformation Cop.comp S (Y a') (Y a) \<Phi>'\<Psi>.map" ..
show ?thesis
proof (intro Cop_S.inverse_arrowsI)
have 0: "inverse_transformations Cop.comp S (Y a) (Y a') \<Phi>\<Psi>'.map \<Phi>'\<Psi>.map" ..
have 1: "Cop_S.antipar (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)
(Cop_S.MkArr (Y a') (Y a) \<Phi>'\<Psi>.map)"
using Ya Ya' \<Phi>\<Psi>' \<Phi>'\<Psi> Cop_S.dom_char Cop_S.cod_char Cop_S.seqI
Cop_S.arr_MkArr Cop_S.cod_MkArr Cop_S.dom_MkArr
by presburger
show "Cop_S.ide (Cop_S.comp (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)
(Cop_S.MkArr (Y a') (Y a) \<Phi>'\<Psi>.map))"
using 0 1 NaturalTransformation.inverse_transformations_inverse(2) Cop_S.comp_MkArr
by (metis Cop_S.cod_MkArr Cop_S.ide_char' Cop_S.seqE)
show "Cop_S.ide (Cop_S.comp (Cop_S.MkArr (Y a') (Y a) \<Phi>'\<Psi>.map)
(Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map))"
using 0 1 NaturalTransformation.inverse_transformations_inverse(1) Cop_S.comp_MkArr
by (metis Cop_S.cod_MkArr Cop_S.ide_char' Cop_S.seqE)
qed
qed
hence 3: "Cop_S.iso (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)" using Cop_S.isoI by blast
hence "Cop_S.arr (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map)" using Cop_S.iso_is_arr by blast
hence "Cop_S.in_hom (Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map) (map a) (map a')"
using Ya.ide_a Ya'.ide_a Cop_S.dom_char Cop_S.cod_char by auto
hence "\<exists>f. \<guillemotleft>f : a \<rightarrow> a'\<guillemotright> \<and> map f = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
using Ya.ide_a Ya'.ide_a is_full Y_def Cop_S.iso_is_arr full_functor.is_full
by auto
from this obtain \<phi>
where \<phi>: "\<guillemotleft>\<phi> : a \<rightarrow> a'\<guillemotright> \<and> map \<phi> = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
by blast
from \<phi> have "C.iso \<phi>"
using 3 reflects_iso [of \<phi> a a'] by simp
hence EX: "\<exists>\<phi>. \<guillemotleft>\<phi> : a \<rightarrow> a'\<guillemotright> \<and> C.iso \<phi> \<and> map \<phi> = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
using \<phi> by blast
have
UN: "\<And>\<phi>'. \<guillemotleft>\<phi>' : a \<rightarrow> a'\<guillemotright> \<and> map \<phi>' = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map \<Longrightarrow> \<phi>' = \<phi>"
proof -
fix \<phi>'
assume \<phi>': "\<guillemotleft>\<phi>' : a \<rightarrow> a'\<guillemotright> \<and> map \<phi>' = Cop_S.MkArr (Y a) (Y a') \<Phi>\<Psi>'.map"
have "C.par \<phi> \<phi>' \<and> map \<phi> = map \<phi>'" using \<phi> \<phi>' by auto
thus "\<phi>' = \<phi>" using is_faithful by fast
qed
from EX UN show ?thesis by auto
qed
end
section "Diagrams and Cones"
text\<open>
A \emph{diagram} in a category @{term C} is a functor \<open>D: J \<rightarrow> C\<close>.
We refer to the category @{term J} as the diagram \emph{shape}.
Note that in the usual expositions of category theory that use set theory
as their foundations, the shape @{term J} of a diagram is required to be
a ``small'' category, where smallness means that the collection of objects
of @{term J}, as well as each of the ``homs,'' is a set.
However, in HOL there is no class of all sets, so it is not meaningful
to speak of @{term J} as ``small'' in any kind of absolute sense.
There is likely a meaningful notion of smallness of @{term J}
\emph{relative to} @{term C} (the result below that states that a set
category has @{term I}-indexed products if and only if its universe
``admits @{term I}-indexed tuples'' is suggestive of how this might
be defined), but I haven't fully explored this idea at present.
\<close>
locale diagram =
C: category C +
J: category J +
"functor" J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
begin
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
end
lemma comp_diagram_functor:
assumes "diagram J C D" and "functor J' J F"
shows "diagram J' C (D o F)"
by (meson assms(1) assms(2) diagram_def functor.axioms(1) functor_comp)
text\<open>
A \emph{cone} over a diagram \<open>D: J \<rightarrow> C\<close> is a natural transformation
from a constant functor to @{term D}. The value of the constant functor is
the \emph{apex} of the cone.
\<close>
locale cone =
C: category C +
J: category J +
D: diagram J C D +
A: constant_functor J C a +
natural_transformation J C A.map D \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<chi> :: "'j \<Rightarrow> 'c"
begin
lemma ide_apex:
shows "C.ide a"
using A.value_is_ide by auto
lemma component_in_hom:
assumes "J.arr j"
shows "\<guillemotleft>\<chi> j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using assms by auto
end
text\<open>
A cone over diagram @{term D} is transformed into a cone over diagram @{term "D o F"}
by pre-composing with @{term F}.
\<close>
lemma comp_cone_functor:
assumes "cone J C D a \<chi>" and "functor J' J F"
shows "cone J' C (D o F) a (\<chi> o F)"
proof -
interpret \<chi>: cone J C D a \<chi> using assms(1) by auto
interpret F: "functor" J' J F using assms(2) by auto
interpret A': constant_functor J' C a
apply unfold_locales using \<chi>.A.value_is_ide by auto
have 1: "\<chi>.A.map o F = A'.map"
using \<chi>.A.map_def A'.map_def \<chi>.J.not_arr_null by auto
interpret \<chi>': natural_transformation J' C A'.map \<open>D o F\<close> \<open>\<chi> o F\<close>
using 1 horizontal_composite F.natural_transformation_axioms
\<chi>.natural_transformation_axioms
by fastforce
show "cone J' C (D o F) a (\<chi> o F)" ..
qed
text\<open>
A cone over diagram @{term D} can be transformed into a cone over a diagram @{term D'}
by post-composing with a natural transformation from @{term D} to @{term D'}.
\<close>
lemma vcomp_transformation_cone:
assumes "cone J C D a \<chi>"
and "natural_transformation J C D D' \<tau>"
shows "cone J C D' a (vertical_composite.map J C \<chi> \<tau>)"
proof -
interpret \<chi>: cone J C D a \<chi> using assms(1) by auto
interpret \<tau>: natural_transformation J C D D' \<tau> using assms(2) by auto
interpret \<tau>o\<chi>: vertical_composite J C \<chi>.A.map D D' \<chi> \<tau> ..
interpret \<tau>o\<chi>: cone J C D' a \<tau>o\<chi>.map ..
show ?thesis ..
qed
context "functor"
begin
lemma preserves_diagrams:
fixes J :: "'j comp"
assumes "diagram J A D"
shows "diagram J B (F o D)"
proof -
interpret D: diagram J A D using assms by auto
interpret FoD: composite_functor J A B D F ..
show "diagram J B (F o D)" ..
qed
lemma preserves_cones:
fixes J :: "'j comp"
assumes "cone J A D a \<chi>"
shows "cone J B (F o D) (F a) (F o \<chi>)"
proof -
interpret \<chi>: cone J A D a \<chi> using assms by auto
interpret Fa: constant_functor J B \<open>F a\<close>
apply unfold_locales using \<chi>.ide_apex by auto
have 1: "F o \<chi>.A.map = Fa.map"
proof
fix f
show "(F \<circ> \<chi>.A.map) f = Fa.map f"
using is_extensional Fa.is_extensional \<chi>.A.is_extensional
by (cases "\<chi>.J.arr f", simp_all)
qed
interpret \<chi>': natural_transformation J B Fa.map \<open>F o D\<close> \<open>F o \<chi>\<close>
using 1 horizontal_composite \<chi>.natural_transformation_axioms
natural_transformation_axioms
by fastforce
show "cone J B (F o D) (F a) (F o \<chi>)" ..
qed
end
context diagram
begin
abbreviation cone
where "cone a \<chi> \<equiv> Limit.cone J C D a \<chi>"
abbreviation cones :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) set"
where "cones a \<equiv> { \<chi>. cone a \<chi> }"
text\<open>
An arrow @{term "f \<in> C.hom a' a"} induces by composition a transformation from
cones with apex @{term a} to cones with apex @{term a'}. This transformation
is functorial in @{term f}.
\<close>
abbreviation cones_map :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> ('j \<Rightarrow> 'c)"
where "cones_map f \<equiv> (\<lambda>\<chi> \<in> cones (C.cod f). \<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null)"
lemma cones_map_mapsto:
assumes "C.arr f"
shows "cones_map f \<in>
extensional (cones (C.cod f)) \<inter> (cones (C.cod f) \<rightarrow> cones (C.dom f))"
proof
show "cones_map f \<in> extensional (cones (C.cod f))" by blast
show "cones_map f \<in> cones (C.cod f) \<rightarrow> cones (C.dom f)"
proof
fix \<chi>
assume "\<chi> \<in> cones (C.cod f)"
hence \<chi>: "cone (C.cod f) \<chi>" by auto
interpret \<chi>: cone J C D \<open>C.cod f\<close> \<chi> using \<chi> by auto
interpret B: constant_functor J C \<open>C.dom f\<close>
apply unfold_locales using assms by auto
have "cone (C.dom f) (\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null)"
using assms B.value_is_ide \<chi>.is_natural_1 \<chi>.is_natural_2
apply (unfold_locales, auto)
using \<chi>.is_natural_1
apply (metis C.comp_assoc)
using \<chi>.is_natural_2 C.comp_arr_dom
by (metis J.arr_cod_iff_arr J.cod_cod C.comp_assoc)
thus "(\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null) \<in> cones (C.dom f)" by auto
qed
qed
lemma cones_map_ide:
assumes "\<chi> \<in> cones a"
shows "cones_map a \<chi> = \<chi>"
proof -
interpret \<chi>: cone J C D a \<chi> using assms by auto
show ?thesis
proof
fix j
show "cones_map a \<chi> j = \<chi> j"
using assms \<chi>.A.value_is_ide \<chi>.preserves_hom C.comp_arr_dom \<chi>.is_extensional
by (cases "J.arr j", auto)
qed
qed
lemma cones_map_comp:
assumes "C.seq f g"
shows "cones_map (f \<cdot> g) = restrict (cones_map g o cones_map f) (cones (C.cod f))"
proof (intro restr_eqI)
show "cones (C.cod (f \<cdot> g)) = cones (C.cod f)" using assms by simp
show "\<And>\<chi>. \<chi> \<in> cones (C.cod (f \<cdot> g)) \<Longrightarrow>
(\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null) = (cones_map g o cones_map f) \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> cones (C.cod (f \<cdot> g))"
show "(\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null) = (cones_map g o cones_map f) \<chi>"
proof -
have "((cones_map g) o (cones_map f)) \<chi> = cones_map g (cones_map f \<chi>)"
by force
also have "... = (\<lambda>j. if J.arr j then
(\<lambda>j. if J.arr j then \<chi> j \<cdot> f else C.null) j \<cdot> g else C.null)"
proof
fix j
have "cone (C.dom f) (cones_map f \<chi>)"
using assms \<chi> cones_map_mapsto by (elim C.seqE, force)
thus "cones_map g (cones_map f \<chi>) j =
(if J.arr j then C (if J.arr j then \<chi> j \<cdot> f else C.null) g else C.null)"
using \<chi> assms by auto
qed
also have "... = (\<lambda>j. if J.arr j then \<chi> j \<cdot> f \<cdot> g else C.null)"
proof -
have "\<And>j. J.arr j \<Longrightarrow> (\<chi> j \<cdot> f) \<cdot> g = \<chi> j \<cdot> f \<cdot> g"
proof -
interpret \<chi>: cone J C D \<open>C.cod f\<close> \<chi> using assms \<chi> by auto
fix j
assume j: "J.arr j"
show "(\<chi> j \<cdot> f) \<cdot> g = \<chi> j \<cdot> f \<cdot> g"
using assms C.comp_assoc by simp
qed
thus ?thesis by auto
qed
finally show ?thesis by auto
qed
qed
qed
end
text\<open>
Changing the apex of a cone by pre-composing with an arrow @{term f} commutes
with changing the diagram of a cone by post-composing with a natural transformation.
\<close>
lemma cones_map_vcomp:
assumes "diagram J C D" and "diagram J C D'"
and "natural_transformation J C D D' \<tau>"
and "cone J C D a \<chi>"
and f: "partial_magma.in_hom C f a' a"
shows "diagram.cones_map J C D' f (vertical_composite.map J C \<chi> \<tau>)
= vertical_composite.map J C (diagram.cones_map J C D f \<chi>) \<tau>"
proof -
interpret D: diagram J C D using assms(1) by auto
interpret D': diagram J C D' using assms(2) by auto
interpret \<tau>: natural_transformation J C D D' \<tau> using assms(3) by auto
interpret \<chi>: cone J C D a \<chi> using assms(4) by auto
interpret \<tau>o\<chi>: vertical_composite J C \<chi>.A.map D D' \<chi> \<tau> ..
interpret \<tau>o\<chi>: cone J C D' a \<tau>o\<chi>.map ..
interpret \<chi>f: cone J C D a' \<open>D.cones_map f \<chi>\<close>
using f \<chi>.cone_axioms D.cones_map_mapsto by blast
interpret \<tau>o\<chi>f: vertical_composite J C \<chi>f.A.map D D' \<open>D.cones_map f \<chi>\<close> \<tau> ..
interpret \<tau>o\<chi>_f: cone J C D' a' \<open>D'.cones_map f \<tau>o\<chi>.map\<close>
using f \<tau>o\<chi>.cone_axioms D'.cones_map_mapsto [of f] by blast
write C (infixr "\<cdot>" 55)
show "D'.cones_map f \<tau>o\<chi>.map = \<tau>o\<chi>f.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J C \<chi>f.A.map D' (D'.cones_map f \<tau>o\<chi>.map)" ..
show "natural_transformation J C \<chi>f.A.map D' \<tau>o\<chi>f.map" ..
show "\<And>j. D.J.ide j \<Longrightarrow> D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>f.map j"
proof -
fix j
assume j: "D.J.ide j"
have "D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>.map j \<cdot> f"
using f \<tau>o\<chi>.cone_axioms \<tau>o\<chi>.map_simp_2 \<tau>o\<chi>.is_extensional by auto
also have "... = (\<tau> j \<cdot> \<chi> (D.J.dom j)) \<cdot> f"
using j \<tau>o\<chi>.map_simp_2 by simp
also have "... = \<tau> j \<cdot> \<chi> (D.J.dom j) \<cdot> f"
using D.C.comp_assoc by simp
also have "... = \<tau>o\<chi>f.map j"
using j f \<chi>.cone_axioms \<tau>o\<chi>f.map_simp_2 by auto
finally show "D'.cones_map f \<tau>o\<chi>.map j = \<tau>o\<chi>f.map j" by auto
qed
qed
qed
text\<open>
Given a diagram @{term D}, we can construct a contravariant set-valued functor,
which takes each object @{term a} of @{term C} to the set of cones over @{term D}
with apex @{term a}, and takes each arrow @{term f} of @{term C} to the function
on cones over @{term D} induced by pre-composition with @{term f}.
For this, we need to introduce a set category @{term S} whose universe is large
enough to contain all the cones over @{term D}, and we need to have an explicit
correspondence between cones and elements of the universe of @{term S}.
A replete set category @{term S} equipped with an injective mapping
@{term_type "\<iota> :: ('j => 'c) => 's"} serves this purpose.
\<close>
locale cones_functor =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
begin
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>S _\<guillemotright>")
abbreviation \<o> where "\<o> \<equiv> S.\<o>"
definition map :: "'c \<Rightarrow> 's"
where "map = (\<lambda>f. if C.arr f then
S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom f))
(\<iota> o D.cones_map f o \<o>)
else S.null)"
lemma map_simp [simp]:
assumes "C.arr f"
shows "map f = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom f))
(\<iota> o D.cones_map f o \<o>)"
using assms map_def by auto
lemma arr_map:
assumes "C.arr f"
shows "S.arr (map f)"
proof -
have "\<iota> o D.cones_map f o \<o> \<in> \<iota> ` D.cones (C.cod f) \<rightarrow> \<iota> ` D.cones (C.dom f)"
using assms D.cones_map_mapsto by force
thus ?thesis using assms S.\<iota>_mapsto S.arr_mkArr by auto
qed
lemma map_ide:
assumes "C.ide a"
shows "map a = S.mkIde (\<iota> ` D.cones a)"
proof -
have "map a = S.mkArr (\<iota> ` D.cones a) (\<iota> ` D.cones a) (\<iota> o D.cones_map a o \<o>)"
using assms map_simp by force
also have "... = S.mkArr (\<iota> ` D.cones a) (\<iota> ` D.cones a) (\<lambda>x. x)"
using S.\<iota>_mapsto D.cones_map_ide S.arr_mkArr by force
also have "... = S.mkIde (\<iota> ` D.cones a)"
using assms S.mkIde_as_mkArr S.\<iota>_mapsto by blast
finally show ?thesis by auto
qed
lemma map_preserves_dom:
assumes "Cop.arr f"
shows "map (Cop.dom f) = S.dom (map f)"
using assms arr_map map_ide by auto
lemma map_preserves_cod:
assumes "Cop.arr f"
shows "map (Cop.cod f) = S.cod (map f)"
using assms arr_map map_ide by auto
lemma map_preserves_comp:
assumes "Cop.seq g f"
shows "map (g \<cdot>\<^sup>o\<^sup>p f) = map g \<cdot>\<^sub>S map f"
proof -
have "map (g \<cdot>\<^sup>o\<^sup>p f) = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
((\<iota> o D.cones_map g o \<o>) o (\<iota> o D.cones_map f o \<o>))"
proof -
have 1: "S.arr (map (g \<cdot>\<^sup>o\<^sup>p f))"
using assms arr_map [of "C f g"] by simp
have "map (g \<cdot>\<^sup>o\<^sup>p f) = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
(\<iota> o D.cones_map (C f g) o \<o>)"
using assms map_simp [of "C f g"] by simp
also have "... = S.mkArr (\<iota> ` D.cones (C.cod f)) (\<iota> ` D.cones (C.dom g))
((\<iota> o D.cones_map g o \<o>) o (\<iota> o D.cones_map f o \<o>))"
using assms 1 calculation D.cones_map_mapsto D.cones_map_comp by auto
finally show ?thesis by blast
qed
also have "... = map g \<cdot>\<^sub>S map f"
using assms arr_map [of f] arr_map [of g] map_simp S.comp_mkArr by auto
finally show ?thesis by auto
qed
lemma is_functor:
shows "functor Cop.comp S map"
apply (unfold_locales)
using map_def arr_map map_preserves_dom map_preserves_cod map_preserves_comp
by auto
end
sublocale cones_functor \<subseteq> "functor" Cop.comp S map using is_functor by auto
sublocale cones_functor \<subseteq> set_valued_functor Cop.comp S map ..
section Limits
subsection "Limit Cones"
text\<open>
A \emph{limit cone} for a diagram @{term D} is a cone @{term \<chi>} over @{term D}
with the universal property that any other cone @{term \<chi>'} over the diagram @{term D}
factors uniquely through @{term \<chi>}.
\<close>
locale limit_cone =
C: category C +
J: category J +
D: diagram J C D +
cone J C D a \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<chi> :: "'j \<Rightarrow> 'c" +
assumes is_universal: "cone J C D a' \<chi>' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
begin
definition induced_arrow :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> 'c"
where "induced_arrow a' \<chi>' = (THE f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>')"
lemma induced_arrowI:
assumes \<chi>': "\<chi>' \<in> D.cones a'"
shows "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright>"
and "D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
proof -
have "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
using assms \<chi>' is_universal by simp
hence 1: "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"] induced_arrow_def
by presburger
show "\<guillemotleft>induced_arrow a' \<chi>' : a' \<rightarrow> a\<guillemotright>" using 1 by simp
show "D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'" using 1 by simp
qed
lemma cones_map_induced_arrow:
shows "induced_arrow a' \<in> D.cones a' \<rightarrow> C.hom a' a"
and "\<And>\<chi>'. \<chi>' \<in> D.cones a' \<Longrightarrow> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using induced_arrowI by auto
lemma induced_arrow_cones_map:
assumes "C.ide a'"
shows "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
and "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> induced_arrow a' (D.cones_map f \<chi>) = f"
proof -
have a': "C.ide a'" using assms by (simp add: cone.ide_apex)
have cone_\<chi>: "cone J C D a \<chi>" ..
show "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
using cone_\<chi> D.cones_map_mapsto by blast
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
show "induced_arrow a' (D.cones_map f \<chi>) = f"
proof -
have "D.cones_map f \<chi> \<in> D.cones a'"
using f cone_\<chi> D.cones_map_mapsto by blast
hence "\<exists>!f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = D.cones_map f \<chi>"
using assms is_universal by auto
thus ?thesis
using f induced_arrow_def
the1_equality [of "\<lambda>f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = D.cones_map f \<chi>"]
by presburger
qed
qed
text\<open>
For a limit cone @{term \<chi>} with apex @{term a}, for each object @{term a'} the
hom-set @{term "C.hom a' a"} is in bijective correspondence with the set of cones
with apex @{term a'}.
\<close>
lemma bij_betw_hom_and_cones:
assumes "C.ide a'"
shows "bij_betw (\<lambda>f. D.cones_map f \<chi>) (C.hom a' a) (D.cones a')"
proof (intro bij_betwI)
show "(\<lambda>f. D.cones_map f \<chi>) \<in> C.hom a' a \<rightarrow> D.cones a'"
using assms induced_arrow_cones_map by blast
show "induced_arrow a' \<in> D.cones a' \<rightarrow> C.hom a' a"
using assms cones_map_induced_arrow by blast
show "\<And>f. f \<in> C.hom a' a \<Longrightarrow> induced_arrow a' (D.cones_map f \<chi>) = f"
using assms induced_arrow_cones_map by blast
show "\<And>\<chi>'. \<chi>' \<in> D.cones a' \<Longrightarrow> D.cones_map (induced_arrow a' \<chi>') \<chi> = \<chi>'"
using assms cones_map_induced_arrow by blast
qed
lemma induced_arrow_eqI:
assumes "D.cone a' \<chi>'" and "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" and "D.cones_map f \<chi> = \<chi>'"
shows "induced_arrow a' \<chi>' = f"
using assms is_universal induced_arrow_def
the1_equality [of "\<lambda>f. f \<in> C.hom a' a \<and> D.cones_map f \<chi> = \<chi>'" f]
by simp
lemma induced_arrow_self:
shows "induced_arrow a \<chi> = a"
proof -
have "\<guillemotleft>a : a \<rightarrow> a\<guillemotright> \<and> D.cones_map a \<chi> = \<chi>"
using ide_apex cone_axioms D.cones_map_ide by force
thus ?thesis using induced_arrow_eqI cone_axioms by auto
qed
end
context diagram
begin
abbreviation limit_cone
where "limit_cone a \<chi> \<equiv> Limit.limit_cone J C D a \<chi>"
text\<open>
A diagram @{term D} has object @{term a} as a limit if @{term a} is the apex
of some limit cone over @{term D}.
\<close>
abbreviation has_as_limit :: "'c \<Rightarrow> bool"
where "has_as_limit a \<equiv> (\<exists>\<chi>. limit_cone a \<chi>)"
abbreviation has_limit
where "has_limit \<equiv> (\<exists>a \<chi>. limit_cone a \<chi>)"
definition some_limit :: 'c
where "some_limit = (SOME a. \<exists>\<chi>. limit_cone a \<chi>)"
definition some_limit_cone :: "'j \<Rightarrow> 'c"
where "some_limit_cone = (SOME \<chi>. limit_cone some_limit \<chi>)"
lemma limit_cone_some_limit_cone:
assumes has_limit
shows "limit_cone some_limit some_limit_cone"
proof -
have "\<exists>a. has_as_limit a" using assms by simp
hence "has_as_limit some_limit"
using some_limit_def someI_ex [of "\<lambda>a. \<exists>\<chi>. limit_cone a \<chi>"] by simp
thus "limit_cone some_limit some_limit_cone"
using assms some_limit_cone_def someI_ex [of "\<lambda>\<chi>. limit_cone some_limit \<chi>"]
by simp
qed
lemma ex_limitE:
assumes "\<exists>a. has_as_limit a"
obtains a \<chi> where "limit_cone a \<chi>"
using assms someI_ex by blast
end
subsection "Limits by Representation"
text\<open>
A limit for a diagram D can also be given by a representation \<open>(a, \<Phi>)\<close>
of the cones functor.
\<close>
locale representation_of_cones_functor =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota> +
Cones: cones_functor J C D S \<iota> +
Hom: hom_functor C S \<phi> +
representation_of_functor C S \<phi> Cones.map a \<Phi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
subsection "Putting it all Together"
text\<open>
A ``limit situation'' combines and connects the ways of presenting a limit.
\<close>
locale limit_situation =
C: category C +
Cop: dual_category C +
J: category J +
D: diagram J C D +
S: replete_concrete_set_category S UNIV \<iota> +
Cones: cones_functor J C D S \<iota> +
Hom: hom_functor C S \<phi> +
\<Phi>: representation_of_functor C S \<phi> Cones.map a \<Phi> +
\<chi>: limit_cone J C D a \<chi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and S :: "'s comp" (infixr "\<cdot>\<^sub>S" 55)
and \<phi> :: "'c * 'c \<Rightarrow> 'c \<Rightarrow> 's"
and \<iota> :: "('j \<Rightarrow> 'c) \<Rightarrow> 's"
and a :: 'c
and \<Phi> :: "'c \<Rightarrow> 's"
and \<chi> :: "'j \<Rightarrow> 'c" +
assumes \<chi>_in_terms_of_\<Phi>: "\<chi> = S.\<o> (S.Fun (\<Phi> a) (\<phi> (a, a) a))"
and \<Phi>_in_terms_of_\<chi>:
"Cop.ide a' \<Longrightarrow> \<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a')
(\<lambda>x. \<iota> (D.cones_map (Hom.\<psi> (a', a) x) \<chi>))"
text (in limit_situation) \<open>
The assumption @{prop \<chi>_in_terms_of_\<Phi>} states that the universal cone @{term \<chi>} is obtained
by applying the function @{term "S.Fun (\<Phi> a)"} to the identity @{term a} of
@{term[source=true] C} (after taking into account the necessary coercions).
\<close>
text (in limit_situation) \<open>
The assumption @{prop \<Phi>_in_terms_of_\<chi>} states that the component of @{term \<Phi>} at @{term a'}
is the arrow of @{term[source=true] S} corresponding to the function that takes an arrow
@{term "f \<in> C.hom a' a"} and produces the cone with vertex @{term a'} obtained
by transforming the universal cone @{term \<chi>} by @{term f}.
\<close>
subsection "Limit Cones Induce Limit Situations"
text\<open>
To obtain a limit situation from a limit cone, we need to introduce a set category
that is large enough to contain the hom-sets of @{term C} as well as the cones
over @{term D}. We use the category of all @{typ "('c + ('j \<Rightarrow> 'c))"}-sets for this.
\<close>
context limit_cone
begin
interpretation Cop: dual_category C ..
interpretation CopxC: product_category Cop.comp C ..
interpretation S: replete_setcat \<open>undefined :: 'c + ('j \<Rightarrow> 'c)\<close> .
notation S.comp (infixr "\<cdot>\<^sub>S" 55)
interpretation Sr: replete_concrete_set_category S.comp UNIV \<open>S.UP o Inr\<close>
apply unfold_locales
using S.UP_mapsto
apply auto[1]
using S.inj_UP inj_Inr inj_compose
by metis
interpretation Cones: cones_functor J C D S.comp \<open>S.UP o Inr\<close> ..
interpretation Hom: hom_functor C S.comp \<open>\<lambda>_. S.UP o Inl\<close>
apply (unfold_locales)
using S.UP_mapsto
apply auto[1]
using S.inj_UP injD inj_onI inj_Inl inj_compose
by (metis (no_types, lifting))
interpretation Y: yoneda_functor C S.comp \<open>\<lambda>_. S.UP o Inl\<close> ..
interpretation Ya: yoneda_functor_fixed_object C S.comp \<open>\<lambda>_. S.UP o Inl\<close> a
apply (unfold_locales) using ide_apex by auto
abbreviation inl :: "'c \<Rightarrow> 'c + ('j \<Rightarrow> 'c)" where "inl \<equiv> Inl"
abbreviation inr :: "('j \<Rightarrow> 'c) \<Rightarrow> 'c + ('j \<Rightarrow> 'c)" where "inr \<equiv> Inr"
abbreviation \<iota> where "\<iota> \<equiv> S.UP o inr"
abbreviation \<o> where "\<o> \<equiv> Cones.\<o>"
abbreviation \<phi> where "\<phi> \<equiv> \<lambda>_. S.UP o inl"
abbreviation \<psi> where "\<psi> \<equiv> Hom.\<psi>"
abbreviation Y where "Y \<equiv> Y.Y"
lemma Ya_ide:
assumes a': "C.ide a'"
shows "Y a a' = S.mkIde (Hom.set (a', a))"
using assms ide_apex Y.Y_simp Hom.map_ide by simp
lemma Ya_arr:
assumes g: "C.arr g"
shows "Y a g = S.mkArr (Hom.set (C.cod g, a)) (Hom.set (C.dom g, a))
(\<phi> (C.dom g, a) o Cop.comp g o \<psi> (C.cod g, a))"
using ide_apex g Y.Y_ide_arr [of a g "C.dom g" "C.cod g"] by auto
lemma cone_\<chi> [simp]:
shows "\<chi> \<in> D.cones a"
using cone_axioms by simp
text\<open>
For each object @{term a'} of @{term[source=true] C} we have a function mapping
@{term "C.hom a' a"} to the set of cones over @{term D} with apex @{term a'},
which takes @{term "f \<in> C.hom a' a"} to \<open>\<chi>f\<close>, where \<open>\<chi>f\<close> is the cone obtained by
composing @{term \<chi>} with @{term f} (after accounting for coercions to and from the
universe of @{term S}). The corresponding arrows of @{term S} are the
components of a natural isomorphism from @{term "Y a"} to \<open>Cones\<close>.
\<close>
definition \<Phi>o :: "'c \<Rightarrow> ('c + ('j \<Rightarrow> 'c)) setcat.arr"
where
"\<Phi>o a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>))"
lemma \<Phi>o_in_hom:
assumes a': "C.ide a'"
shows "\<guillemotleft>\<Phi>o a' : S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
proof -
have " \<guillemotleft>S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)) :
S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
proof -
have "(\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)) \<in> Hom.set (a', a) \<rightarrow> \<iota> ` D.cones a'"
proof
fix x
assume x: "x \<in> Hom.set (a', a)"
hence "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using ide_apex a' Hom.\<psi>_mapsto by auto
hence "D.cones_map (\<psi> (a', a) x) \<chi> \<in> D.cones a'"
using ide_apex a' x D.cones_map_mapsto cone_\<chi> by force
thus "\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) \<in> \<iota> ` D.cones a'" by simp
qed
moreover have "Hom.set (a', a) \<subseteq> S.Univ"
using ide_apex a' Hom.set_subset_Univ by auto
moreover have "\<iota> ` D.cones a' \<subseteq> S.Univ"
using S.UP_mapsto by auto
ultimately show ?thesis using S.mkArr_in_hom by simp
qed
thus ?thesis using \<Phi>o_def [of a'] by auto
qed
interpretation \<Phi>: transformation_by_components Cop.comp S.comp \<open>Y a\<close> Cones.map \<Phi>o
proof
fix a'
assume A': "Cop.ide a'"
show "\<guillemotleft>\<Phi>o a' : Y a a' \<rightarrow>\<^sub>S Cones.map a'\<guillemotright>"
using A' Ya_ide \<Phi>o_in_hom Cones.map_ide by auto
next
fix g
assume g: "Cop.arr g"
show "\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g)"
proof -
let ?A = "Hom.set (C.cod g, a)"
let ?B = "Hom.set (C.dom g, a)"
let ?B' = "\<iota> ` D.cones (C.cod g)"
let ?C = "\<iota> ` D.cones (C.dom g)"
let ?F = "\<phi> (C.dom g, a) o Cop.comp g o \<psi> (C.cod g, a)"
let ?F' = "\<iota> o D.cones_map g o \<o>"
let ?G = "\<lambda>x. \<iota> (D.cones_map (\<psi> (C.dom g, a) x) \<chi>)"
let ?G' = "\<lambda>x. \<iota> (D.cones_map (\<psi> (C.cod g, a) x) \<chi>)"
have "S.arr (Y a g) \<and> Y a g = S.mkArr ?A ?B ?F"
using ide_apex g Ya.preserves_arr Ya_arr by fastforce
moreover have "S.arr (\<Phi>o (Cop.cod g))"
using g \<Phi>o_in_hom [of "Cop.cod g"] by auto
moreover have "\<Phi>o (Cop.cod g) = S.mkArr ?B ?C ?G"
using g \<Phi>o_def [of "C.dom g"] by auto
moreover have "S.seq (\<Phi>o (Cop.cod g)) (Y a g)"
using ide_apex g \<Phi>o_in_hom [of "Cop.cod g"] by auto
ultimately have 1: "S.seq (\<Phi>o (Cop.cod g)) (Y a g) \<and>
\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)"
using S.comp_mkArr [of ?A ?B ?F ?C ?G] by argo
have "Cones.map g = S.mkArr (\<iota> ` D.cones (C.cod g)) (\<iota> ` D.cones (C.dom g)) ?F'"
using g Cones.map_simp by fastforce
moreover have "\<Phi>o (Cop.dom g) = S.mkArr ?A ?B' ?G'"
using g \<Phi>o_def by fastforce
moreover have "S.seq (Cones.map g) (\<Phi>o (Cop.dom g))"
using g Cones.preserves_hom [of g "C.cod g" "C.dom g"] \<Phi>o_in_hom [of "Cop.dom g"]
by force
ultimately have
2: "S.seq (Cones.map g) (\<Phi>o (Cop.dom g)) \<and>
Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g) = S.mkArr ?A ?C (?F' o ?G')"
using S.seqI' [of "\<Phi>o (Cop.dom g)" "Cones.map g"] S.comp_mkArr by auto
have "\<Phi>o (Cop.cod g) \<cdot>\<^sub>S Y a g = S.mkArr ?A ?C (?G o ?F)"
using 1 by auto
also have "... = S.mkArr ?A ?C (?F' o ?G')"
proof (intro S.mkArr_eqI')
show "S.arr (S.mkArr ?A ?C (?G o ?F))" using 1 by force
show "\<And>x. x \<in> ?A \<Longrightarrow> (?G o ?F) x = (?F' o ?G') x"
proof -
fix x
assume x: "x \<in> ?A"
hence 1: "\<guillemotleft>\<psi> (C.cod g, a) x : C.cod g \<rightarrow> a\<guillemotright>"
using ide_apex g Hom.\<psi>_mapsto [of "C.cod g" a] by auto
have "(?G o ?F) x = \<iota> (D.cones_map (\<psi> (C.dom g, a)
(\<phi> (C.dom g, a) (\<psi> (C.cod g, a) x \<cdot> g))) \<chi>)"
proof - (* Why is it so balky with this proof? *)
have "(?G o ?F) x = ?G (?F x)" by simp
also have "... = \<iota> (D.cones_map (\<psi> (C.dom g, a)
(\<phi> (C.dom g, a) (\<psi> (C.cod g, a) x \<cdot> g))) \<chi>)"
proof -
have "?F x = \<phi> (C.dom g, a) (\<psi> (C.cod g, a) x \<cdot> g)" by simp
thus ?thesis by presburger (* presburger 5ms, metis 797ms! Why? *)
qed
finally show ?thesis by auto
qed
also have "... = \<iota> (D.cones_map (\<psi> (C.cod g, a) x \<cdot> g) \<chi>)"
proof -
have "\<guillemotleft>\<psi> (C.cod g, a) x \<cdot> g : C.dom g \<rightarrow> a\<guillemotright>" using g 1 by auto
thus ?thesis using Hom.\<psi>_\<phi> by presburger
qed
also have "... = \<iota> (D.cones_map g (D.cones_map (\<psi> (C.cod g, a) x) \<chi>))"
using g x 1 cone_\<chi> D.cones_map_comp [of "\<psi> (C.cod g, a) x" g] by fastforce
also have "... = \<iota> (D.cones_map g (\<o> (\<iota> (D.cones_map (\<psi> (C.cod g, a) x) \<chi>))))"
using 1 cone_\<chi> D.cones_map_mapsto Sr.\<o>_\<iota> by auto
also have "... = (?F' o ?G') x" by simp
finally show "(?G o ?F) x = (?F' o ?G') x" by auto
qed
qed
also have "... = Cones.map g \<cdot>\<^sub>S \<Phi>o (Cop.dom g)"
using 2 by auto
finally show ?thesis by auto
qed
qed
interpretation \<Phi>: set_valued_transformation Cop.comp S.comp \<open>Y a\<close> Cones.map \<Phi>.map ..
interpretation \<Phi>: natural_isomorphism Cop.comp S.comp \<open>Y a\<close> Cones.map \<Phi>.map
proof
fix a'
assume a': "Cop.ide a'"
show "S.iso (\<Phi>.map a')"
proof -
let ?F = "\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
have bij: "bij_betw ?F (Hom.set (a', a)) (\<iota> ` D.cones a')"
proof -
have "\<And>x x'. \<lbrakk> x \<in> Hom.set (a', a); x' \<in> Hom.set (a', a);
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>) \<rbrakk>
\<Longrightarrow> x = x'"
proof -
fix x x'
assume x: "x \<in> Hom.set (a', a)" and x': "x' \<in> Hom.set (a', a)"
and xx': "\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)"
have \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>" using x ide_apex a' Hom.\<psi>_mapsto by auto
have \<psi>x': "\<guillemotleft>\<psi> (a', a) x' : a' \<rightarrow> a\<guillemotright>" using x' ide_apex a' Hom.\<psi>_mapsto by auto
have 1: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
have "D.cones_map (\<psi> (a', a) x) \<chi> \<in> D.cones a'"
using \<psi>x a' cone_\<chi> D.cones_map_mapsto by force
hence 2: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>"
using a' is_universal by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
have "\<And>f. \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)
\<longleftrightarrow> D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>"
proof -
fix f :: 'c
have "D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>
\<longrightarrow> \<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
by simp
thus "(\<iota> (D.cones_map f \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>))
= (D.cones_map f \<chi> = D.cones_map (\<psi> (a', a) x) \<chi>)"
by (meson Sr.inj_\<iota> injD)
qed
thus ?thesis using 2 by auto
qed
qed
have 2: "\<exists>!x''. x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof -
from 1 obtain f'' where
f'': "\<guillemotleft>f'' : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f'' \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
by blast
have "\<phi> (a', a) f'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) (\<phi> (a', a) f'')) \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
proof
show "\<phi> (a', a) f'' \<in> Hom.set (a', a)" using f'' Hom.set_def by auto
show "\<iota> (D.cones_map (\<psi> (a', a) (\<phi> (a', a) f'')) \<chi>) =
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
using f'' Hom.\<psi>_\<phi> by presburger
qed
moreover have
"\<And>x''. x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)
\<Longrightarrow> x'' = \<phi> (a', a) f''"
proof -
fix x''
assume x'': "x'' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
hence "\<guillemotleft>\<psi> (a', a) x'' : a' \<rightarrow> a\<guillemotright> \<and>
\<iota> (D.cones_map (\<psi> (a', a) x'') \<chi>) = \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
using ide_apex a' Hom.set_def Hom.\<psi>_mapsto [of a' a] by auto
hence "\<phi> (a', a) (\<psi> (a', a) x'') = \<phi> (a', a) f''"
using 1 f'' by auto
thus "x'' = \<phi> (a', a) f''"
using ide_apex a' x'' Hom.\<phi>_\<psi> by simp
qed
ultimately show ?thesis
using ex1I [of "\<lambda>x'. x' \<in> Hom.set (a', a) \<and>
\<iota> (D.cones_map (\<psi> (a', a) x') \<chi>) =
\<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
"\<phi> (a', a) f''"]
by simp
qed
thus "x = x'" using x x' xx' by auto
qed
hence "inj_on ?F (Hom.set (a', a))"
using inj_onI [of "Hom.set (a', a)" ?F] by auto
moreover have "?F ` Hom.set (a', a) = \<iota> ` D.cones a'"
proof
show "?F ` Hom.set (a', a) \<subseteq> \<iota> ` D.cones a'"
proof
fix X'
assume X': "X' \<in> ?F ` Hom.set (a', a)"
from this obtain x' where x': "x' \<in> Hom.set (a', a) \<and> ?F x' = X'" by blast
show "X' \<in> \<iota> ` D.cones a'"
proof -
have "X' = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)" using x' by blast
hence "X' = \<iota> (D.cones_map (\<psi> (a', a) x') \<chi>)" using x' by force
moreover have "\<guillemotleft>\<psi> (a', a) x' : a' \<rightarrow> a\<guillemotright>"
using ide_apex a' x' Hom.set_def Hom.\<psi>_\<phi> by auto
ultimately show ?thesis
using x' cone_\<chi> D.cones_map_mapsto by force
qed
qed
show "\<iota> ` D.cones a' \<subseteq> ?F ` Hom.set (a', a)"
proof
fix X'
assume X': "X' \<in> \<iota> ` D.cones a'"
hence "\<o> X' \<in> \<o> ` \<iota> ` D.cones a'" by simp
with Sr.\<o>_\<iota> have "\<o> X' \<in> D.cones a'"
by auto
hence "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<o> X'"
using a' is_universal by simp
from this obtain f where "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<o> X'"
by auto
hence f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> \<iota> (D.cones_map f \<chi>) = X'"
using X' Sr.\<iota>_\<o> by auto
have "X' = ?F (\<phi> (a', a) f)"
using f Hom.\<psi>_\<phi> by presburger
thus "X' \<in> ?F ` Hom.set (a', a)"
using f Hom.set_def by force
qed
qed
ultimately show ?thesis
using bij_betw_def [of ?F "Hom.set (a', a)" "\<iota> ` D.cones a'"] inj_on_def by auto
qed
let ?f = "S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
have iso: "S.iso ?f"
proof -
have "?F \<in> Hom.set (a', a) \<rightarrow> \<iota> ` D.cones a'"
using bij bij_betw_imp_funcset by fast
hence 1: "S.arr ?f"
using ide_apex a' Hom.set_subset_Univ S.\<iota>_mapsto S.arr_mkArr by auto
thus ?thesis using bij S.iso_char S.arr_mkArr S.set_mkIde by fastforce
qed
moreover have "?f = \<Phi>.map a'"
using a' \<Phi>o_def by force
finally show ?thesis by auto
qed
qed
interpretation R: representation_of_functor C S.comp \<phi> Cones.map a \<Phi>.map ..
lemma \<chi>_in_terms_of_\<Phi>:
shows "\<chi> = \<o> (\<Phi>.FUN a (\<phi> (a, a) a))"
proof -
have "\<Phi>.FUN a (\<phi> (a, a) a) =
(\<lambda>x \<in> Hom.set (a, a). \<iota> (D.cones_map (\<psi> (a, a) x) \<chi>)) (\<phi> (a, a) a)"
using ide_apex S.Fun_mkArr \<Phi>.map_simp_ide \<Phi>o_def \<Phi>.preserves_reflects_arr [of a]
by simp
also have "... = \<iota> (D.cones_map a \<chi>)"
proof -
have "\<phi> (a, a) a \<in> Hom.set (a, a)"
using ide_apex Hom.\<phi>_mapsto by fastforce
hence "(\<lambda>x \<in> Hom.set (a, a). \<iota> (D.cones_map (\<psi> (a, a) x) \<chi>)) (\<phi> (a, a) a)
= \<iota> (D.cones_map (\<psi> (a, a) (\<phi> (a, a) a)) \<chi>)"
using restrict_apply' [of "\<phi> (a, a) a" "Hom.set (a, a)"] by blast
also have "... = \<iota> (D.cones_map a \<chi>)"
proof -
have "\<psi> (a, a) (\<phi> (a, a) a) = a"
using ide_apex Hom.\<psi>_\<phi> [of a a a] by fastforce
thus ?thesis by metis
qed
finally show ?thesis by auto
qed
finally have "\<Phi>.FUN a (\<phi> (a, a) a) = \<iota> (D.cones_map a \<chi>)" by auto
also have "... = \<iota> \<chi>"
using ide_apex D.cones_map_ide [of \<chi> a] cone_\<chi> by simp
finally have "\<Phi>.FUN a (\<phi> (a, a) a) = \<iota> \<chi>" by blast
hence "\<o> (\<Phi>.FUN a (\<phi> (a, a) a)) = \<o> (\<iota> \<chi>)" by simp
thus ?thesis using cone_\<chi> Sr.\<o>_\<iota> by simp
qed
abbreviation Hom
where "Hom \<equiv> Hom.map"
abbreviation \<Phi>
where "\<Phi> \<equiv> \<Phi>.map"
lemma induces_limit_situation:
shows "limit_situation J C D S.comp \<phi> \<iota> a \<Phi> \<chi>"
proof
show "\<chi> = \<o> (\<Phi>.FUN a (\<phi> (a, a) a))" using \<chi>_in_terms_of_\<Phi> by auto
fix a'
show "Cop.ide a' \<Longrightarrow> \<Phi>.map a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a')
(\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>))"
using \<Phi>.map_simp_ide \<Phi>o_def [of a'] by force
qed
no_notation S.comp (infixr "\<cdot>\<^sub>S" 55)
end
sublocale limit_cone \<subseteq> limit_situation J C D replete_setcat.comp \<phi> \<iota> a \<Phi> \<chi>
using induces_limit_situation by auto
subsection "Representations of the Cones Functor Induce Limit Situations"
context representation_of_cones_functor
begin
interpretation \<Phi>: set_valued_transformation Cop.comp S \<open>Y a\<close> Cones.map \<Phi> ..
interpretation \<Psi>: inverse_transformation Cop.comp S \<open>Y a\<close> Cones.map \<Phi> ..
interpretation \<Psi>: set_valued_transformation Cop.comp S Cones.map \<open>Y a\<close> \<Psi>.map ..
abbreviation \<o>
where "\<o> \<equiv> Cones.\<o>"
abbreviation \<chi>
where "\<chi> \<equiv> \<o> (S.Fun (\<Phi> a) (\<phi> (a, a) a))"
lemma Cones_SET_eq_\<iota>_img_cones:
assumes "C.ide a'"
shows "Cones.SET a' = \<iota> ` D.cones a'"
proof -
have "\<iota> ` D.cones a' \<subseteq> S.Univ" using S.\<iota>_mapsto by auto
thus ?thesis using assms Cones.map_ide S.set_mkIde by auto
qed
lemma \<iota>\<chi>:
shows "\<iota> \<chi> = S.Fun (\<Phi> a) (\<phi> (a, a) a)"
proof -
have "S.Fun (\<Phi> a) (\<phi> (a, a) a) \<in> Cones.SET a"
using Ya.ide_a Hom.\<phi>_mapsto S.Fun_mapsto [of "\<Phi> a"] Hom.set_map by fastforce
thus ?thesis
using Ya.ide_a Cones_SET_eq_\<iota>_img_cones by auto
qed
interpretation \<chi>: cone J C D a \<chi>
proof -
have "\<iota> \<chi> \<in> \<iota> ` D.cones a"
using Ya.ide_a \<iota>\<chi> S.Fun_mapsto [of "\<Phi> a"] Hom.\<phi>_mapsto Hom.set_map
Cones_SET_eq_\<iota>_img_cones by fastforce
thus "D.cone a \<chi>"
by (metis S.\<o>_\<iota> UNIV_I imageE mem_Collect_eq)
qed
lemma cone_\<chi>:
shows "D.cone a \<chi>" ..
lemma \<Phi>_FUN_simp:
assumes a': "C.ide a'" and x: "x \<in> Hom.set (a', a)"
shows "\<Phi>.FUN a' x = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)"
proof -
have \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using Ya.ide_a a' x Hom.\<psi>_mapsto by blast
have \<phi>a: "\<phi> (a, a) a \<in> Hom.set (a, a)" using Ya.ide_a Hom.\<phi>_mapsto by fastforce
have "\<Phi>.FUN a' x = (\<Phi>.FUN a' o Ya.FUN (\<psi> (a', a) x)) (\<phi> (a, a) a)"
proof -
have "\<phi> (a', a) (a \<cdot> \<psi> (a', a) x) = x"
using Ya.ide_a a' x \<psi>x Hom.\<phi>_\<psi> C.comp_cod_arr by fastforce
moreover have "S.arr (S.mkArr (Hom.set (a, a)) (Hom.set (a', a))
(\<phi> (a', a) \<circ> Cop.comp (\<psi> (a', a) x) \<circ> \<psi> (a, a)))"
by (metis (no_types) Cop.hom_char Ya.Y_ide_arr(2) Ya.preserves_reflects_arr
\<chi>.ide_apex \<psi>x Cop.in_homE)
ultimately show ?thesis
using Ya.ide_a a' x Ya.Y_ide_arr \<psi>x \<phi>a C.ide_in_hom by auto
qed
also have "... = (Cones.FUN (\<psi> (a', a) x) o \<Phi>.FUN a) (\<phi> (a, a) a)"
proof -
have "(\<Phi>.FUN a' o Ya.FUN (\<psi> (a', a) x)) (\<phi> (a, a) a)
= S.Fun (\<Phi> a' \<cdot>\<^sub>S Y a (\<psi> (a', a) x)) (\<phi> (a, a) a)"
using \<psi>x a' \<phi>a Ya.ide_a Ya.map_simp Hom.set_map by (elim C.in_homE, auto)
also have "... = S.Fun (S (Cones.map (\<psi> (a', a) x)) (\<Phi> a)) (\<phi> (a, a) a)"
using \<psi>x is_natural_1 [of "\<psi> (a', a) x"] is_natural_2 [of "\<psi> (a', a) x"] by auto
also have "... = (Cones.FUN (\<psi> (a', a) x) o \<Phi>.FUN a) (\<phi> (a, a) a)"
proof -
have "S.seq (Cones.map (\<psi> (a', a) x)) (\<Phi> a)"
using Ya.ide_a \<psi>x Cones.map_preserves_dom [of "\<psi> (a', a) x"]
apply (intro S.seqI)
apply auto[2]
by fastforce
thus ?thesis
using Ya.ide_a \<phi>a Hom.set_map by auto
qed
finally show ?thesis by simp
qed
also have "... = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)" using \<iota>\<chi> by simp
finally show ?thesis by auto
qed
lemma \<chi>_is_universal:
assumes "D.cone a' \<chi>'"
shows "\<guillemotleft>\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>')) : a' \<rightarrow> a\<guillemotright>"
and "D.cones_map (\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))) \<chi> = \<chi>'"
and "\<lbrakk> \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D.cones_map f' \<chi> = \<chi>' \<rbrakk> \<Longrightarrow> f' = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
proof -
interpret \<chi>': cone J C D a' \<chi>' using assms by auto
have a': "C.ide a'" using \<chi>'.ide_apex by simp
have \<iota>\<chi>': "\<iota> \<chi>' \<in> Cones.SET a'" using assms a' Cones_SET_eq_\<iota>_img_cones by auto
let ?f = "\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
have A: "\<Psi>.FUN a' (\<iota> \<chi>') \<in> Hom.set (a', a)"
proof -
have "\<Psi>.FUN a' \<in> Cones.SET a' \<rightarrow> Ya.SET a'"
using a' \<Psi>.preserves_hom [of a' a' a'] S.Fun_mapsto [of "\<Psi>.map a'"] by fastforce
thus ?thesis using a' \<iota>\<chi>' Ya.ide_a Hom.set_map by auto
qed
show f: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using A a' Ya.ide_a Hom.\<psi>_mapsto [of a' a] by auto
have E: "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> Cones.FUN f (\<iota> \<chi>) = \<Phi>.FUN a' (\<phi> (a', a) f)"
proof -
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
have "\<phi> (a', a) f \<in> Hom.set (a', a)"
using a' Ya.ide_a f Hom.\<phi>_mapsto by auto
thus "Cones.FUN f (\<iota> \<chi>) = \<Phi>.FUN a' (\<phi> (a', a) f)"
using a' f \<Phi>_FUN_simp by simp
qed
have I: "\<Phi>.FUN a' (\<Psi>.FUN a' (\<iota> \<chi>')) = \<iota> \<chi>'"
proof -
have "\<Phi>.FUN a' (\<Psi>.FUN a' (\<iota> \<chi>')) =
compose (\<Psi>.DOM a') (\<Phi>.FUN a') (\<Psi>.FUN a') (\<iota> \<chi>')"
using a' \<iota>\<chi>' Cones.map_ide \<Psi>.preserves_hom [of a' a' a'] by force
also have "... = (\<lambda>x \<in> \<Psi>.DOM a'. x) (\<iota> \<chi>')"
using a' \<Psi>.inverts_components S.inverse_arrows_char by force
also have "... = \<iota> \<chi>'"
using a' \<iota>\<chi>' Cones.map_ide \<Psi>.preserves_hom [of a' a' a'] by force
finally show ?thesis by auto
qed
show f\<chi>: "D.cones_map ?f \<chi> = \<chi>'"
proof -
have "D.cones_map ?f \<chi> = (\<o> o Cones.FUN ?f o \<iota>) \<chi>"
using f Cones.preserves_arr [of ?f] cone_\<chi>
by (cases "D.cone a \<chi>", auto)
also have "... = \<chi>'"
using f Ya.ide_a a' A E I by auto
finally show ?thesis by auto
qed
show "\<lbrakk> \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D.cones_map f' \<chi> = \<chi>' \<rbrakk> \<Longrightarrow> f' = ?f"
proof -
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>" and f'\<chi>: "D.cones_map f' \<chi> = \<chi>'"
show "f' = ?f"
proof -
have 1: "\<phi> (a', a) f' \<in> Hom.set (a', a) \<and> \<phi> (a', a) ?f \<in> Hom.set (a', a)"
using Ya.ide_a a' f f' Hom.\<phi>_mapsto by auto
have "S.iso (\<Phi> a')" using \<chi>'.ide_apex components_are_iso by auto
hence 2: "S.arr (\<Phi> a') \<and> bij_betw (\<Phi>.FUN a') (Hom.set (a', a)) (Cones.SET a')"
using Ya.ide_a a' S.iso_char Hom.set_map by auto
have "\<Phi>.FUN a' (\<phi> (a', a) f') = \<Phi>.FUN a' (\<phi> (a', a) ?f)"
proof -
have "\<Phi>.FUN a' (\<phi> (a', a) ?f) = \<iota> \<chi>'"
using A I Hom.\<phi>_\<psi> Ya.ide_a a' by simp
also have "... = Cones.FUN f' (\<iota> \<chi>)"
using f f' A E cone_\<chi> Cones.preserves_arr f\<chi> f'\<chi> by (elim C.in_homE, auto)
also have "... = \<Phi>.FUN a' (\<phi> (a', a) f')"
using f' E by simp
finally show ?thesis by argo
qed
moreover have "inj_on (\<Phi>.FUN a') (Hom.set (a', a))"
using 2 bij_betw_imp_inj_on by blast
ultimately have 3: "\<phi> (a', a) f' = \<phi> (a', a) ?f"
using 1 inj_on_def [of "\<Phi>.FUN a'" "Hom.set (a', a)"] by blast
show ?thesis
proof -
have "f' = \<psi> (a', a) (\<phi> (a', a) f')"
using Ya.ide_a a' f' Hom.\<psi>_\<phi> by simp
also have "... = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
using Ya.ide_a a' Hom.\<psi>_\<phi> A 3 by simp
finally show ?thesis by blast
qed
qed
qed
qed
interpretation \<chi>: limit_cone J C D a \<chi>
proof
show "\<And>a' \<chi>'. D.cone a' \<chi>' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
proof -
fix a' \<chi>'
assume 1: "D.cone a' \<chi>'"
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>'"
proof
show "\<guillemotleft>\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>')) : a' \<rightarrow> a\<guillemotright> \<and>
D.cones_map (\<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))) \<chi> = \<chi>'"
using 1 \<chi>_is_universal by blast
show "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' \<Longrightarrow> f = \<psi> (a', a) (\<Psi>.FUN a' (\<iota> \<chi>'))"
using 1 \<chi>_is_universal by blast
qed
qed
qed
lemma \<chi>_is_limit_cone:
shows "D.limit_cone a \<chi>" ..
lemma induces_limit_situation:
shows "limit_situation J C D S \<phi> \<iota> a \<Phi> \<chi>"
proof
show "\<chi> = \<chi>" by simp
fix a'
assume a': "Cop.ide a'"
let ?F = "\<lambda>x. \<iota> (D.cones_map (\<psi> (a', a) x) \<chi>)"
show "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
proof -
have 1: "\<guillemotleft>\<Phi> a' : S.mkIde (Hom.set (a', a)) \<rightarrow>\<^sub>S S.mkIde (\<iota> ` D.cones a')\<guillemotright>"
using a' Cones.map_ide Ya.ide_a by auto
moreover have "\<Phi>.DOM a' = Hom.set (a', a)"
using 1 Hom.set_subset_Univ a' Ya.ide_a Hom.set_map by simp
moreover have "\<Phi>.COD a' = \<iota> ` D.cones a'"
using a' Cones_SET_eq_\<iota>_img_cones by fastforce
ultimately have 2: "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<Phi>.FUN a')"
using S.mkArr_Fun [of "\<Phi> a'"] by fastforce
also have "... = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F"
proof
show "S.arr (S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') (\<Phi>.FUN a'))"
using 1 2 by auto
show "\<And>x. x \<in> Hom.set (a', a) \<Longrightarrow> \<Phi>.FUN a' x = ?F x"
proof -
fix x
assume x: "x \<in> Hom.set (a', a)"
hence \<psi>x: "\<guillemotleft>\<psi> (a', a) x : a' \<rightarrow> a\<guillemotright>"
using a' Ya.ide_a Hom.\<psi>_mapsto by auto
show "\<Phi>.FUN a' x = ?F x"
proof -
have "\<Phi>.FUN a' x = Cones.FUN (\<psi> (a', a) x) (\<iota> \<chi>)"
using a' x \<Phi>_FUN_simp by simp
also have "... = restrict (\<iota> o D.cones_map (\<psi> (a', a) x) o \<o>) (\<iota> ` D.cones a) (\<iota> \<chi>)"
using \<psi>x Cones.map_simp Cones.preserves_arr [of "\<psi> (a', a) x"] S.Fun_mkArr
by (elim C.in_homE, auto)
also have "... = ?F x" using cone_\<chi> by simp
ultimately show ?thesis by simp
qed
qed
qed
finally show "\<Phi> a' = S.mkArr (Hom.set (a', a)) (\<iota> ` D.cones a') ?F" by auto
qed
qed
end
sublocale representation_of_cones_functor \<subseteq> limit_situation J C D S \<phi> \<iota> a \<Phi> \<chi>
using induces_limit_situation by auto
section "Categories with Limits"
context category
begin
text\<open>
A category @{term[source=true] C} has limits of shape @{term J} if every diagram of shape
@{term J} admits a limit cone.
\<close>
definition has_limits_of_shape
where "has_limits_of_shape J \<equiv> \<forall>D. diagram J C D \<longrightarrow> (\<exists>a \<chi>. limit_cone J C D a \<chi>)"
text\<open>
A category has limits at a type @{typ 'j} if it has limits of shape @{term J}
for every category @{term J} whose arrows are of type @{typ 'j}.
\<close>
definition has_limits
where "has_limits (_ :: 'j) \<equiv> \<forall>J :: 'j comp. category J \<longrightarrow> has_limits_of_shape J"
text\<open>
Whether a category has limits of shape \<open>J\<close> truly depends only on the ``shape''
(\emph{i.e.}~isomorphism class) of \<open>J\<close> and not on details of its construction.
\<close>
lemma has_limits_preserved_by_isomorphism:
assumes "has_limits_of_shape J" and "isomorphic_categories J J'"
shows "has_limits_of_shape J'"
proof -
interpret J: category J
using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto
interpret J': category J'
using assms(2) isomorphic_categories_def isomorphic_categories_axioms_def by auto
from assms(2) obtain \<phi> \<psi> where IF: "inverse_functors J' J \<phi> \<psi>"
using isomorphic_categories_def isomorphic_categories_axioms_def
inverse_functors_sym
by blast
interpret IF: inverse_functors J' J \<phi> \<psi> using IF by auto
have \<psi>\<phi>: "\<psi> o \<phi> = J.map" using IF.inv by metis
have \<phi>\<psi>: "\<phi> o \<psi> = J'.map" using IF.inv' by metis
have "\<And>D'. diagram J' C D' \<Longrightarrow> \<exists>a \<chi>. limit_cone J' C D' a \<chi>"
proof -
fix D'
assume D': "diagram J' C D'"
interpret D': diagram J' C D' using D' by auto
interpret D: composite_functor J J' C \<phi> D' ..
interpret D: diagram J C \<open>D' o \<phi>\<close> ..
have D: "diagram J C (D' o \<phi>)" ..
from assms(1) obtain a \<chi> where \<chi>: "D.limit_cone a \<chi>"
using D has_limits_of_shape_def by blast
interpret \<chi>: limit_cone J C \<open>D' o \<phi>\<close> a \<chi> using \<chi> by auto
interpret A': constant_functor J' C a
using \<chi>.ide_apex by (unfold_locales, auto)
have \<chi>o\<psi>: "cone J' C (D' o \<phi> o \<psi>) a (\<chi> o \<psi>)"
using comp_cone_functor IF.G.functor_axioms \<chi>.cone_axioms by fastforce
hence \<chi>o\<psi>: "cone J' C D' a (\<chi> o \<psi>)"
using \<phi>\<psi> by (metis D'.functor_axioms Fun.comp_assoc comp_functor_identity)
interpret \<chi>o\<psi>: cone J' C D' a \<open>\<chi> o \<psi>\<close> using \<chi>o\<psi> by auto
interpret \<chi>o\<psi>: limit_cone J' C D' a \<open>\<chi> o \<psi>\<close>
proof
fix a' \<chi>'
assume \<chi>': "D'.cone a' \<chi>'"
interpret \<chi>': cone J' C D' a' \<chi>' using \<chi>' by auto
have \<chi>'o\<phi>: "cone J C (D' o \<phi>) a' (\<chi>' o \<phi>)"
using \<chi>' comp_cone_functor IF.F.functor_axioms by fastforce
interpret \<chi>'o\<phi>: cone J C \<open>D' o \<phi>\<close> a' \<open>\<chi>' o \<phi>\<close> using \<chi>'o\<phi> by auto
have "cone J C (D' o \<phi>) a' (\<chi>' o \<phi>)" ..
hence 1: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"
using \<chi>.is_universal by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f (\<chi> o \<psi>) = \<chi>'"
proof
let ?f = "THE f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"
have f: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map ?f \<chi> = \<chi>' o \<phi>"
using 1 theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = \<chi>' o \<phi>"] by blast
have f_in_hom: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using f by blast
have "D'.cones_map ?f (\<chi> o \<psi>) = \<chi>'"
proof
fix j'
have "\<not>J'.arr j' \<Longrightarrow> D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'"
proof -
assume j': "\<not>J'.arr j'"
have "D'.cones_map ?f (\<chi> o \<psi>) j' = null"
using j' f_in_hom \<chi>o\<psi> by fastforce
thus ?thesis
using j' \<chi>'.is_extensional by simp
qed
moreover have "J'.arr j' \<Longrightarrow> D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'"
proof -
assume j': "J'.arr j'"
have "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi> (\<psi> j') \<cdot> ?f"
using j' f \<chi>o\<psi> by fastforce
also have "... = D.cones_map ?f \<chi> (\<psi> j')"
using j' f_in_hom \<chi> \<chi>.cone_\<chi> by fastforce
also have "... = \<chi>' j'"
using j' f \<chi> \<phi>\<psi> Fun.comp_def J'.map_simp by metis
finally show "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'" by auto
qed
ultimately show "D'.cones_map ?f (\<chi> o \<psi>) j' = \<chi>' j'" by blast
qed
thus "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map ?f (\<chi> o \<psi>) = \<chi>'" using f by auto
fix f'
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f' (\<chi> o \<psi>) = \<chi>'"
have "D.cones_map f' \<chi> = \<chi>' o \<phi>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j"
using f' \<chi> \<chi>'o\<phi>.is_extensional \<chi>.cone_\<chi> mem_Collect_eq restrict_apply by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j"
proof -
assume j: "J.arr j"
have "D.cones_map f' \<chi> j = C (\<chi> j) f'"
using j f' \<chi>.cone_\<chi> by auto
also have "... = C ((\<chi> o \<psi>) (\<phi> j)) f'"
using j f' \<psi>\<phi> by (metis comp_apply J.map_simp)
also have "... = D'.cones_map f' (\<chi> o \<psi>) (\<phi> j)"
using j f' \<chi>o\<psi> by fastforce
also have "... = (\<chi>' o \<phi>) j"
using j f' by auto
finally show "D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j" by auto
qed
ultimately show "D.cones_map f' \<chi> j = (\<chi>' o \<phi>) j" by blast
qed
hence "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f' \<chi> = \<chi>' o \<phi>"
using f' by auto
moreover have "\<And>P x x'. (\<exists>!x. P x) \<and> P x \<and> P x' \<Longrightarrow> x = x'"
by auto
ultimately show "f' = ?f" using 1 f by blast
qed
qed
have "limit_cone J' C D' a (\<chi> o \<psi>)" ..
thus "\<exists>a \<chi>. limit_cone J' C D' a \<chi>" by blast
qed
thus ?thesis using has_limits_of_shape_def by auto
qed
end
subsection "Diagonal Functors"
text\<open>
The existence of limits can also be expressed in terms of adjunctions: a category @{term C}
has limits of shape @{term J} if the diagonal functor taking each object @{term a}
in @{term C} to the constant-@{term a} diagram and each arrow \<open>f \<in> C.hom a a'\<close>
to the constant-@{term f} natural transformation between diagrams is a left adjoint functor.
\<close>
locale diagonal_functor =
C: category C +
J: category J +
J_C: functor_category J C
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
begin
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
notation J_C.comp (infixr "\<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>]" 55)
notation J_C.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] _\<guillemotright>")
definition map :: "'c \<Rightarrow> ('j, 'c) J_C.arr"
where "map f = (if C.arr f then J_C.MkArr (constant_functor.map J C (C.dom f))
(constant_functor.map J C (C.cod f))
(constant_transformation.map J C f)
else J_C.null)"
lemma is_functor:
shows "functor C J_C.comp map"
proof
fix f
show "\<not> C.arr f \<Longrightarrow> local.map f = J_C.null"
using map_def by simp
assume f: "C.arr f"
interpret Dom_f: constant_functor J C \<open>C.dom f\<close>
using f by (unfold_locales, auto)
interpret Cod_f: constant_functor J C \<open>C.cod f\<close>
using f by (unfold_locales, auto)
interpret Fun_f: constant_transformation J C f
using f by (unfold_locales, auto)
show 1: "J_C.arr (map f)"
using f map_def by (simp add: Fun_f.natural_transformation_axioms)
show "J_C.dom (map f) = map (C.dom f)"
proof -
have "constant_transformation J C (C.dom f)"
apply unfold_locales using f by auto
hence "constant_transformation.map J C (C.dom f) = Dom_f.map"
using Dom_f.map_def constant_transformation.map_def [of J C "C.dom f"] by auto
thus ?thesis using f 1 by (simp add: map_def J_C.dom_char)
qed
show "J_C.cod (map f) = map (C.cod f)"
proof -
have "constant_transformation J C (C.cod f)"
apply unfold_locales using f by auto
hence "constant_transformation.map J C (C.cod f) = Cod_f.map"
using Cod_f.map_def constant_transformation.map_def [of J C "C.cod f"] by auto
thus ?thesis using f 1 by (simp add: map_def J_C.cod_char)
qed
next
fix f g
assume g: "C.seq g f"
have f: "C.arr f" using g by auto
interpret Dom_f: constant_functor J C \<open>C.dom f\<close>
using f by (unfold_locales, auto)
interpret Cod_f: constant_functor J C \<open>C.cod f\<close>
using f by (unfold_locales, auto)
interpret Fun_f: constant_transformation J C f
using f by (unfold_locales, auto)
interpret Cod_g: constant_functor J C \<open>C.cod g\<close>
using g by (unfold_locales, auto)
interpret Fun_g: constant_transformation J C g
using g by (unfold_locales, auto)
interpret Fun_g: natural_transformation J C Cod_f.map Cod_g.map Fun_g.map
apply unfold_locales
using f g C.seqE [of g f] C.comp_arr_dom C.comp_cod_arr Fun_g.is_extensional by auto
interpret Fun_fg: vertical_composite
J C Dom_f.map Cod_f.map Cod_g.map Fun_f.map Fun_g.map ..
have 1: "J_C.arr (map f)"
using f map_def by (simp add: Fun_f.natural_transformation_axioms)
show "map (g \<cdot> f) = map g \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f"
proof -
have "map (C g f) = J_C.MkArr Dom_f.map Cod_g.map
(constant_transformation.map J C (C g f))"
using f g map_def by simp
also have "... = J_C.MkArr Dom_f.map Cod_g.map (\<lambda>j. if J.arr j then C g f else C.null)"
proof -
have "constant_transformation J C (g \<cdot> f)"
apply unfold_locales using g by auto
thus ?thesis using constant_transformation.map_def by metis
qed
also have "... = J_C.comp (J_C.MkArr Cod_f.map Cod_g.map Fun_g.map)
(J_C.MkArr Dom_f.map Cod_f.map Fun_f.map)"
proof -
have "J_C.MkArr Cod_f.map Cod_g.map Fun_g.map \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>]
J_C.MkArr Dom_f.map Cod_f.map Fun_f.map
= J_C.MkArr Dom_f.map Cod_g.map Fun_fg.map"
using J_C.comp_char J_C.comp_MkArr Fun_f.natural_transformation_axioms
Fun_g.natural_transformation_axioms
by blast
also have "... = J_C.MkArr Dom_f.map Cod_g.map
(\<lambda>j. if J.arr j then g \<cdot> f else C.null)"
proof -
have "Fun_fg.map = (\<lambda>j. if J.arr j then g \<cdot> f else C.null)"
using 1 f g Fun_fg.map_def by auto
thus ?thesis by auto
qed
finally show ?thesis by auto
qed
also have "... = map g \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map f"
using f g map_def by fastforce
finally show ?thesis by auto
qed
qed
sublocale "functor" C J_C.comp map
using is_functor by auto
text\<open>
The objects of \<open>[J, C]\<close> correspond bijectively to diagrams of shape @{term J}
in @{term C}.
\<close>
lemma ide_determines_diagram:
assumes "J_C.ide d"
shows "diagram J C (J_C.Map d)" and "J_C.MkIde (J_C.Map d) = d"
proof -
interpret \<delta>: natural_transformation J C \<open>J_C.Map d\<close> \<open>J_C.Map d\<close> \<open>J_C.Map d\<close>
using assms J_C.ide_char J_C.arr_MkArr by fastforce
interpret D: "functor" J C \<open>J_C.Map d\<close> ..
show "diagram J C (J_C.Map d)" ..
show "J_C.MkIde (J_C.Map d) = d"
using assms J_C.ide_char by (metis J_C.ideD(1) J_C.MkArr_Map)
qed
lemma diagram_determines_ide:
assumes "diagram J C D"
shows "J_C.ide (J_C.MkIde D)" and "J_C.Map (J_C.MkIde D) = D"
proof -
interpret D: diagram J C D using assms by auto
show "J_C.ide (J_C.MkIde D)" using J_C.ide_char
using D.functor_axioms J_C.ide_MkIde by auto
thus "J_C.Map (J_C.MkIde D) = D"
using J_C.in_homE by simp
qed
lemma bij_betw_ide_diagram:
shows "bij_betw J_C.Map (Collect J_C.ide) (Collect (diagram J C))"
proof (intro bij_betwI)
show "J_C.Map \<in> Collect J_C.ide \<rightarrow> Collect (diagram J C)"
using ide_determines_diagram by blast
show "J_C.MkIde \<in> Collect (diagram J C) \<rightarrow> Collect J_C.ide"
using diagram_determines_ide by blast
show "\<And>d. d \<in> Collect J_C.ide \<Longrightarrow> J_C.MkIde (J_C.Map d) = d"
using ide_determines_diagram by blast
show "\<And>D. D \<in> Collect (diagram J C) \<Longrightarrow> J_C.Map (J_C.MkIde D) = D"
using diagram_determines_ide by blast
qed
text\<open>
Arrows from from the diagonal functor correspond bijectively to cones.
\<close>
lemma arrow_determines_cone:
assumes "J_C.ide d" and "arrow_from_functor C J_C.comp map a d x"
shows "cone J C (J_C.Map d) a (J_C.Map x)"
and "J_C.MkArr (constant_functor.map J C a) (J_C.Map d) (J_C.Map x) = x"
proof -
interpret D: diagram J C \<open>J_C.Map d\<close>
using assms ide_determines_diagram by auto
interpret x: arrow_from_functor C J_C.comp map a d x
using assms by auto
interpret A: constant_functor J C a
using x.arrow by (unfold_locales, auto)
interpret \<alpha>: constant_transformation J C a
using x.arrow by (unfold_locales, auto)
have Dom_x: "J_C.Dom x = A.map"
proof -
have "J_C.dom x = map a" using x.arrow by blast
hence "J_C.Map (J_C.dom x) = J_C.Map (map a)" by simp
hence "J_C.Dom x = J_C.Map (map a)"
using A.value_is_ide x.arrow J_C.in_homE by (metis J_C.Map_dom)
moreover have "J_C.Map (map a) = \<alpha>.map"
using A.value_is_ide preserves_ide map_def by simp
ultimately show ?thesis using \<alpha>.map_def A.map_def by auto
qed
have Cod_x: "J_C.Cod x = J_C.Map d"
using x.arrow by auto
interpret \<chi>: natural_transformation J C A.map \<open>J_C.Map d\<close> \<open>J_C.Map x\<close>
using x.arrow J_C.arr_char [of x] Dom_x Cod_x by force
show "D.cone a (J_C.Map x)" ..
show "J_C.MkArr A.map (J_C.Map d) (J_C.Map x) = x"
using x.arrow Dom_x Cod_x \<chi>.natural_transformation_axioms
by (intro J_C.arr_eqI, auto)
qed
lemma cone_determines_arrow:
assumes "J_C.ide d" and "cone J C (J_C.Map d) a \<chi>"
shows "arrow_from_functor C J_C.comp map a d
(J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>)"
and "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>) = \<chi>"
proof -
interpret \<chi>: cone J C \<open>J_C.Map d\<close> a \<chi> using assms(2) by auto
let ?x = "J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi>"
interpret x: arrow_from_functor C J_C.comp map a d ?x
proof
have "\<guillemotleft>J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi> :
J_C.MkIde \<chi>.A.map \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde (J_C.Map d)\<guillemotright>"
using \<chi>.natural_transformation_axioms by auto
moreover have "J_C.MkIde \<chi>.A.map = map a"
using \<chi>.A.value_is_ide map_def \<chi>.A.map_def C.ide_char
by (metis (no_types, lifting) J_C.dom_MkArr preserves_arr preserves_dom)
moreover have "J_C.MkIde (J_C.Map d) = d"
using assms ide_determines_diagram(2) by simp
ultimately show "C.ide a \<and> \<guillemotleft>J_C.MkArr \<chi>.A.map (J_C.Map d) \<chi> : map a \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\<guillemotright>"
using \<chi>.A.value_is_ide by simp
qed
show "arrow_from_functor C J_C.comp map a d ?x" ..
show "J_C.Map (J_C.MkArr (constant_functor.map J C a) (J_C.Map d) \<chi>) = \<chi>"
by (simp add: \<chi>.natural_transformation_axioms)
qed
text\<open>
Transforming a cone by composing at the apex with an arrow @{term g} corresponds,
via the preceding bijections, to composition in \<open>[J, C]\<close> with the image of @{term g}
under the diagonal functor.
\<close>
lemma cones_map_is_composition:
assumes "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>" and "cone J C D a \<chi>"
shows "J_C.MkArr (constant_functor.map J C a') D (diagram.cones_map J C D g \<chi>)
= J_C.MkArr (constant_functor.map J C a) D \<chi> \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
proof -
interpret A: constant_transformation J C a
using assms(1) by (unfold_locales, auto)
interpret \<chi>: cone J C D a \<chi> using assms(2) by auto
have cone_\<chi>: "cone J C D a \<chi>" ..
interpret A': constant_transformation J C a'
using assms(1) by (unfold_locales, auto)
let ?\<chi>' = "\<chi>.D.cones_map g \<chi>"
interpret \<chi>': cone J C D a' ?\<chi>'
using assms(1) cone_\<chi> \<chi>.D.cones_map_mapsto by blast
let ?x = "J_C.MkArr \<chi>.A.map D \<chi>"
let ?x' = "J_C.MkArr \<chi>'.A.map D ?\<chi>'"
show "?x' = J_C.comp ?x (map g)"
proof (intro J_C.arr_eqI)
have x: "J_C.arr ?x"
using \<chi>.natural_transformation_axioms J_C.arr_char [of ?x] by simp
show x': "J_C.arr ?x'"
using \<chi>'.natural_transformation_axioms J_C.arr_char [of ?x'] by simp
have 3: "\<guillemotleft>?x : map a \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\<guillemotright>"
proof -
have 1: "map a = J_C.MkIde A.map"
using \<chi>.ide_apex A.equals_dom_if_value_is_ide A.equals_cod_if_value_is_ide map_def
by auto
have "J_C.arr ?x" using x by blast
moreover have "J_C.dom ?x = map a"
using x J_C.dom_char 1 x \<chi>.ide_apex A.equals_dom_if_value_is_ide \<chi>.D.functor_axioms
J_C.ide_char
by auto
moreover have "J_C.cod ?x = J_C.MkIde D" using x J_C.cod_char by auto
ultimately show ?thesis by fast
qed
have 4: "\<guillemotleft>?x' : map a' \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] J_C.MkIde D\<guillemotright>"
proof -
have 1: "map a' = J_C.MkIde A'.map"
using \<chi>'.ide_apex A'.equals_dom_if_value_is_ide A'.equals_cod_if_value_is_ide map_def
by auto
have "J_C.arr ?x'" using x' by blast
moreover have "J_C.dom ?x' = map a'"
using x' J_C.dom_char 1 x' \<chi>'.ide_apex A'.equals_dom_if_value_is_ide \<chi>.D.functor_axioms
J_C.ide_char
by force
moreover have "J_C.cod ?x' = J_C.MkIde D" using x' J_C.cod_char by auto
ultimately show ?thesis by fast
qed
have seq_xg: "J_C.seq ?x (map g)"
using assms(1) 3 preserves_hom [of g] by (intro J_C.seqI', auto)
show 2: "J_C.seq ?x (map g)"
using seq_xg J_C.seqI' by blast
show "J_C.Dom ?x' = J_C.Dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
have "J_C.Dom ?x' = J_C.Dom (J_C.dom ?x')"
using x' J_C.Dom_dom by simp
also have "... = J_C.Dom (map a')"
using 4 by force
also have "... = J_C.Dom (J_C.dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))"
using assms(1) 2 by auto
also have "... = J_C.Dom (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
using seq_xg J_C.Dom_dom J_C.seqI' by blast
finally show ?thesis by auto
qed
show "J_C.Cod ?x' = J_C.Cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
have "J_C.Cod ?x' = J_C.Cod (J_C.cod ?x')"
using x' J_C.Cod_cod by simp
also have "... = J_C.Cod (J_C.MkIde D)"
using 4 by force
also have "... = J_C.Cod (J_C.cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g))"
using 2 3 J_C.cod_comp J_C.in_homE by metis
also have "... = J_C.Cod (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
using seq_xg J_C.Cod_cod J_C.seqI' by blast
finally show ?thesis by auto
qed
show "J_C.Map ?x' = J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g)"
proof -
interpret g: constant_transformation J C g
apply unfold_locales using assms(1) by auto
interpret \<chi>og: vertical_composite J C A'.map \<chi>.A.map D g.map \<chi>
using assms(1) C.comp_arr_dom C.comp_cod_arr A'.is_extensional g.is_extensional
apply (unfold_locales, auto)
by (elim J.seqE, auto)
have "J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g) = \<chi>og.map"
using assms(1) 2 J_C.comp_char map_def by auto
also have "... = J_C.Map ?x'"
using x' \<chi>og.map_def J_C.arr_char [of ?x'] natural_transformation.is_extensional
assms(1) cone_\<chi> \<chi>og.map_simp_2
by fastforce
finally show ?thesis by auto
qed
qed
qed
text\<open>
Coextension along an arrow from a functor is equivalent to a transformation of cones.
\<close>
lemma coextension_iff_cones_map:
assumes x: "arrow_from_functor C J_C.comp map a d x"
and g: "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>"
and x': "\<guillemotleft>x' : map a' \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] d\<guillemotright>"
shows "arrow_from_functor.is_coext C J_C.comp map a x a' x' g
\<longleftrightarrow> J_C.Map x' = diagram.cones_map J C (J_C.Map d) g (J_C.Map x)"
proof -
interpret x: arrow_from_functor C J_C.comp map a d x
using assms by auto
interpret A': constant_functor J C a'
using assms(2) by (unfold_locales, auto)
have x': "arrow_from_functor C J_C.comp map a' d x'"
using A'.value_is_ide assms(3) by (unfold_locales, blast)
have d: "J_C.ide d" using J_C.ide_cod x.arrow by blast
let ?D = "J_C.Map d"
let ?\<chi> = "J_C.Map x"
let ?\<chi>' = "J_C.Map x'"
interpret D: diagram J C ?D
using ide_determines_diagram J_C.ide_cod x.arrow by blast
interpret \<chi>: cone J C ?D a ?\<chi>
using assms(1) d arrow_determines_cone by simp
interpret \<gamma>: constant_transformation J C g
using g \<chi>.ide_apex by (unfold_locales, auto)
interpret \<chi>og: vertical_composite J C A'.map \<chi>.A.map ?D \<gamma>.map ?\<chi>
using g C.comp_arr_dom C.comp_cod_arr \<gamma>.is_extensional by (unfold_locales, auto)
show ?thesis
proof
assume 0: "x.is_coext a' x' g"
show "?\<chi>' = D.cones_map g ?\<chi>"
proof -
have 1: "x' = x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
using 0 x.is_coext_def by blast
hence "?\<chi>' = J_C.Map x'"
using 0 x.is_coext_def by fast
moreover have "... = D.cones_map g ?\<chi>"
proof -
have "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) =
x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g"
using d g cones_map_is_composition arrow_determines_cone(2) \<chi>.cone_axioms
x.arrow_from_functor_axioms
by auto
hence f1: "J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)) = x'"
by (metis 1)
have "J_C.arr (J_C.MkArr A'.map (J_C.Map d) (D.cones_map g (J_C.Map x)))"
using 1 d g cones_map_is_composition preserves_arr arrow_determines_cone(2)
\<chi>.cone_axioms x.arrow_from_functor_axioms assms(3)
by auto
thus ?thesis
using f1 by auto
qed
ultimately show ?thesis by blast
qed
next
assume X': "?\<chi>' = D.cones_map g ?\<chi>"
show "x.is_coext a' x' g"
proof -
have 4: "J_C.seq x (map g)"
using g x.arrow mem_Collect_eq preserves_arr preserves_cod
by (elim C.in_homE, auto)
hence 1: "x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] map g =
J_C.MkArr (J_C.Dom (map g)) (J_C.Cod x)
(vertical_composite.map J C (J_C.Map (map g)) ?\<chi>)"
using J_C.comp_char [of x "map g"] by simp
have 2: "vertical_composite.map J C (J_C.Map (map g)) ?\<chi> = \<chi>og.map"
by (simp add: map_def \<gamma>.value_is_arr \<gamma>.natural_transformation_axioms)
have 3: "... = D.cones_map g ?\<chi>"
using g \<chi>og.map_simp_2 \<chi>.cone_axioms \<chi>og.is_extensional by auto
have "J_C.MkArr A'.map ?D ?\<chi>' = J_C.comp x (map g)"
proof -
have f1: "A'.map = J_C.Dom (map g)"
using \<gamma>.natural_transformation_axioms map_def g by auto
have "J_C.Map d = J_C.Cod x"
using x.arrow by auto
thus ?thesis using f1 X' 1 2 3 by argo
qed
moreover have "J_C.MkArr A'.map ?D ?\<chi>' = x'"
using d x' arrow_determines_cone by blast
ultimately show ?thesis
using g x.is_coext_def by simp
qed
qed
qed
end
locale right_adjoint_to_diagonal_functor =
C: category C +
J: category J +
J_C: functor_category J C +
\<Delta>: diagonal_functor J C +
"functor" J_C.comp C G +
Adj: meta_adjunction J_C.comp C \<Delta>.map G \<phi> \<psi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and G :: "('j, 'c) functor_category.arr \<Rightarrow> 'c"
and \<phi> :: "'c \<Rightarrow> ('j, 'c) functor_category.arr \<Rightarrow> 'c"
and \<psi> :: "('j, 'c) functor_category.arr \<Rightarrow> 'c \<Rightarrow> ('j, 'c) functor_category.arr" +
assumes adjoint: "adjoint_functors J_C.comp C \<Delta>.map G"
begin
interpretation Adj: adjunction J_C.comp C replete_setcat.comp Adj.\<phi>C Adj.\<phi>D \<Delta>.map G
\<phi> \<psi> Adj.\<eta> Adj.\<epsilon> Adj.\<Phi> Adj.\<Psi>
using Adj.induces_adjunction by simp
text\<open>
A right adjoint @{term G} to a diagonal functor maps each object @{term d} of
\<open>[J, C]\<close> (corresponding to a diagram @{term D} of shape @{term J} in @{term C}
to an object of @{term C}. This object is the limit object, and the component at @{term d}
of the counit of the adjunction determines the limit cone.
\<close>
lemma gives_limit_cones:
assumes "diagram J C D"
shows "limit_cone J C D (G (J_C.MkIde D)) (J_C.Map (Adj.\<epsilon> (J_C.MkIde D)))"
proof -
interpret D: diagram J C D using assms by auto
let ?d = "J_C.MkIde D"
let ?a = "G ?d"
let ?x = "Adj.\<epsilon> ?d"
let ?\<chi> = "J_C.Map ?x"
have "diagram J C D" ..
hence 1: "J_C.ide ?d" using \<Delta>.diagram_determines_ide by auto
hence 2: "J_C.Map (J_C.MkIde D) = D"
using assms 1 J_C.in_homE \<Delta>.diagram_determines_ide(2) by simp
interpret x: terminal_arrow_from_functor C J_C.comp \<Delta>.map ?a ?d ?x
apply unfold_locales
apply (metis (no_types, lifting) "1" preserves_ide Adj.\<epsilon>_in_terms_of_\<psi>
Adj.\<epsilon>o_def Adj.\<epsilon>o_in_hom)
by (metis 1 Adj.has_terminal_arrows_from_functor(1)
terminal_arrow_from_functor.is_terminal)
have 3: "arrow_from_functor C J_C.comp \<Delta>.map ?a ?d ?x" ..
interpret \<chi>: cone J C D ?a ?\<chi>
using 1 2 3 \<Delta>.arrow_determines_cone [of ?d] by auto
have cone_\<chi>: "D.cone ?a ?\<chi>" ..
interpret \<chi>: limit_cone J C D ?a ?\<chi>
proof
fix a' \<chi>'
assume cone_\<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J C D a' \<chi>' using cone_\<chi>' by auto
let ?x' = "J_C.MkArr \<chi>'.A.map D \<chi>'"
interpret x': arrow_from_functor C J_C.comp \<Delta>.map a' ?d ?x'
using 1 2 by (metis \<Delta>.cone_determines_arrow(1) cone_\<chi>')
have "arrow_from_functor C J_C.comp \<Delta>.map a' ?d ?x'" ..
hence 4: "\<exists>!g. x.is_coext a' ?x' g"
using x.is_terminal by simp
have 5: "\<And>g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<Longrightarrow> x.is_coext a' ?x' g \<longleftrightarrow> D.cones_map g ?\<chi> = \<chi>'"
proof -
fix g
assume g: "\<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright>"
show "x.is_coext a' ?x' g \<longleftrightarrow> D.cones_map g ?\<chi> = \<chi>'"
proof -
have "\<guillemotleft>?x' : \<Delta>.map a' \<rightarrow>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] ?d\<guillemotright>"
using x'.arrow by simp
thus ?thesis
using 1 2 3 g \<Delta>.coextension_iff_cones_map \<Delta>.cone_determines_arrow(2)
cone_\<chi>'
by force
qed
qed
have 6: "\<And>g. x.is_coext a' ?x' g \<Longrightarrow> \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright>"
using x.is_coext_def by simp
show "\<exists>!g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<and> D.cones_map g ?\<chi> = \<chi>'"
proof -
have "\<exists>g. \<guillemotleft>g : a' \<rightarrow>\<^sub>C ?a\<guillemotright> \<and> D.cones_map g ?\<chi> = \<chi>'"
using 4 5 6 by meson
thus ?thesis
using 4 5 6 by blast
qed
qed
show "D.limit_cone ?a ?\<chi>" ..
qed
corollary gives_limits:
assumes "diagram J C D"
shows "diagram.has_as_limit J C D (G (J_C.MkIde D))"
using assms gives_limit_cones by fastforce
end
lemma (in category) has_limits_iff_left_adjoint_diagonal:
assumes "category J"
shows "has_limits_of_shape J \<longleftrightarrow>
left_adjoint_functor C (functor_category.comp J C) (diagonal_functor.map J C)"
proof -
interpret J: category J using assms by auto
interpret J_C: functor_category J C ..
interpret \<Delta>: diagonal_functor J C ..
show ?thesis
proof
assume A: "left_adjoint_functor C J_C.comp \<Delta>.map"
interpret \<Delta>: left_adjoint_functor C J_C.comp \<Delta>.map using A by auto
interpret Adj: meta_adjunction J_C.comp C \<Delta>.map \<Delta>.G \<Delta>.\<phi> \<Delta>.\<psi>
using \<Delta>.induces_meta_adjunction by auto
have "meta_adjunction J_C.comp C \<Delta>.map \<Delta>.G \<Delta>.\<phi> \<Delta>.\<psi>" ..
hence 1: "adjoint_functors J_C.comp C \<Delta>.map \<Delta>.G"
using adjoint_functors_def by blast
interpret G: right_adjoint_to_diagonal_functor J C \<Delta>.G \<Delta>.\<phi> \<Delta>.\<psi>
using 1 by (unfold_locales, auto)
have "\<And>D. diagram J C D \<Longrightarrow> \<exists>a. diagram.has_as_limit J C D a"
using A G.gives_limits by blast
hence "\<And>D. diagram J C D \<Longrightarrow> \<exists>a \<chi>. limit_cone J C D a \<chi>"
by metis
thus "has_limits_of_shape J" using has_limits_of_shape_def by blast
next
text\<open>
If @{term "has_limits J"}, then every diagram @{term D} from @{term J} to
@{term[source=true] C} has a limit cone.
This means that, for every object @{term d} of the functor category
\<open>[J, C]\<close>, there exists an object @{term a} of @{term C} and a terminal arrow from
\<open>\<Delta> a\<close> to @{term d} in \<open>[J, C]\<close>. The terminal arrow is given by the
limit cone.
\<close>
assume A: "has_limits_of_shape J"
show "left_adjoint_functor C J_C.comp \<Delta>.map"
proof
fix d
assume D: "J_C.ide d"
interpret D: diagram J C \<open>J_C.Map d\<close>
using D \<Delta>.ide_determines_diagram by auto
let ?D = "J_C.Map d"
have "diagram J C (J_C.Map d)" ..
from this obtain a \<chi> where limit: "limit_cone J C ?D a \<chi>"
using A has_limits_of_shape_def by blast
interpret A: constant_functor J C a
using limit by (simp add: Limit.cone_def limit_cone_def)
interpret \<chi>: limit_cone J C ?D a \<chi> using limit by simp
have cone_\<chi>: "cone J C ?D a \<chi>" ..
let ?x = "J_C.MkArr A.map ?D \<chi>"
interpret x: arrow_from_functor C J_C.comp \<Delta>.map a d ?x
using D cone_\<chi> \<Delta>.cone_determines_arrow by auto
have "terminal_arrow_from_functor C J_C.comp \<Delta>.map a d ?x"
proof
show "\<And>a' x'. arrow_from_functor C J_C.comp \<Delta>.map a' d x' \<Longrightarrow> \<exists>!g. x.is_coext a' x' g"
proof -
fix a' x'
assume x': "arrow_from_functor C J_C.comp \<Delta>.map a' d x'"
interpret x': arrow_from_functor C J_C.comp \<Delta>.map a' d x' using x' by auto
interpret A': constant_functor J C a'
by (unfold_locales, simp add: x'.arrow)
let ?\<chi>' = "J_C.Map x'"
interpret \<chi>': cone J C ?D a' ?\<chi>'
using D x' \<Delta>.arrow_determines_cone by auto
have cone_\<chi>': "cone J C ?D a' ?\<chi>'" ..
let ?g = "\<chi>.induced_arrow a' ?\<chi>'"
show "\<exists>!g. x.is_coext a' x' g"
proof
show "x.is_coext a' x' ?g"
proof (unfold x.is_coext_def)
have 1: "\<guillemotleft>?g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map ?g \<chi> = ?\<chi>'"
using \<chi>.induced_arrow_def \<chi>.is_universal cone_\<chi>'
theI' [of "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<chi> = ?\<chi>'"]
by presburger
hence 2: "x' = ?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map ?g"
proof -
have "x' = J_C.MkArr A'.map ?D ?\<chi>'"
using D \<Delta>.arrow_determines_cone(2) x'.arrow_from_functor_axioms by auto
thus ?thesis
using 1 cone_\<chi> \<Delta>.cones_map_is_composition [of ?g a' a ?D \<chi>] by simp
qed
show "\<guillemotleft>?g : a' \<rightarrow> a\<guillemotright> \<and> x' = ?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map ?g"
using 1 2 by auto
qed
next
fix g
assume X: "x.is_coext a' x' g"
show "g = ?g"
proof -
have "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map g \<chi> = ?\<chi>'"
proof
show G: "\<guillemotleft>g : a' \<rightarrow> a\<guillemotright>" using X x.is_coext_def by blast
show "D.cones_map g \<chi> = ?\<chi>'"
proof -
have "?\<chi>' = J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map g)"
using X x.is_coext_def [of a' x' g] by fast
also have "... = D.cones_map g \<chi>"
proof -
interpret map_g: constant_transformation J C g
using G by (unfold_locales, auto)
interpret \<chi>': vertical_composite J C
map_g.F.map A.map \<open>\<chi>.\<Phi>.Ya.Cop_S.Map d\<close>
map_g.map \<chi>
proof (intro_locales)
have "map_g.G.map = A.map"
using G by blast
thus "natural_transformation_axioms J (\<cdot>) map_g.F.map A.map map_g.map"
using map_g.natural_transformation_axioms
by (simp add: natural_transformation_def)
qed
have "J_C.Map (?x \<cdot>\<^sub>[\<^sub>J\<^sub>,\<^sub>C\<^sub>] \<Delta>.map g) = vertical_composite.map J C map_g.map \<chi>"
proof -
have "J_C.seq ?x (\<Delta>.map g)"
using G x.arrow by auto
thus ?thesis
using G \<Delta>.map_def J_C.Map_comp' [of ?x "\<Delta>.map g"] by auto
qed
also have "... = D.cones_map g \<chi>"
using G cone_\<chi> \<chi>'.map_def map_g.map_def \<chi>.is_natural_2 \<chi>'.map_simp_2
by auto
finally show ?thesis by blast
qed
finally show ?thesis by auto
qed
qed
thus ?thesis
using cone_\<chi>' \<chi>.is_universal \<chi>.induced_arrow_def
theI_unique [of "\<lambda>g. \<guillemotleft>g : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map g \<chi> = ?\<chi>'" g]
by presburger
qed
qed
qed
qed
thus "\<exists>a x. terminal_arrow_from_functor C J_C.comp \<Delta>.map a d x" by auto
qed
qed
qed
section "Right Adjoint Functors Preserve Limits"
context right_adjoint_functor
begin
lemma preserves_limits:
fixes J :: "'j comp"
assumes "diagram J C E" and "diagram.has_as_limit J C E a"
shows "diagram.has_as_limit J D (G o E) (G a)"
proof -
text\<open>
From the assumption that @{term E} has a limit, obtain a limit cone @{term \<chi>}.
\<close>
interpret J: category J using assms(1) diagram_def by auto
interpret E: diagram J C E using assms(1) by auto
from assms(2) obtain \<chi> where \<chi>: "limit_cone J C E a \<chi>" by auto
interpret \<chi>: limit_cone J C E a \<chi> using \<chi> by auto
have a: "C.ide a" using \<chi>.ide_apex by auto
text\<open>
Form the @{term E}-image \<open>GE\<close> of the diagram @{term E}.
\<close>
interpret GE: composite_functor J C D E G ..
interpret GE: diagram J D GE.map ..
text\<open>Let \<open>G\<chi>\<close> be the @{term G}-image of the cone @{term \<chi>},
and note that it is a cone over \<open>GE\<close>.\<close>
let ?G\<chi> = "G o \<chi>"
interpret G\<chi>: cone J D GE.map \<open>G a\<close> ?G\<chi>
using \<chi>.cone_axioms preserves_cones by blast
text\<open>
Claim that \<open>G\<chi>\<close> is a limit cone for diagram \<open>GE\<close>.
\<close>
interpret G\<chi>: limit_cone J D GE.map \<open>G a\<close> ?G\<chi>
proof
text \<open>
Let @{term \<kappa>} be an arbitrary cone over \<open>GE\<close>.
\<close>
fix b \<kappa>
assume \<kappa>: "GE.cone b \<kappa>"
interpret \<kappa>: cone J D GE.map b \<kappa> using \<kappa> by auto
interpret Fb: constant_functor J C \<open>F b\<close>
apply unfold_locales
by (meson F_is_functor \<kappa>.ide_apex functor.preserves_ide)
interpret Adj: meta_adjunction C D F G \<phi> \<psi>
using induces_meta_adjunction by auto
interpret Adj: adjunction C D replete_setcat.comp
Adj.\<phi>C Adj.\<phi>D F G \<phi> \<psi> Adj.\<eta> Adj.\<epsilon> Adj.\<Phi> Adj.\<Psi>
using Adj.induces_adjunction by simp
text\<open>
For each arrow @{term j} of @{term J}, let @{term "\<chi>' j"} be defined to be
the adjunct of @{term "\<chi> j"}. We claim that @{term \<chi>'} is a cone over @{term E}.
\<close>
let ?\<chi>' = "\<lambda>j. if J.arr j then Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j) else C.null"
have cone_\<chi>': "E.cone (F b) ?\<chi>'"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> ?\<chi>' j = C.null" by simp
fix j
assume j: "J.arr j"
show "C.dom (?\<chi>' j) = Fb.map (J.dom j)" using j \<psi>_in_hom by simp
show "C.cod (?\<chi>' j) = E (J.cod j)" using j \<psi>_in_hom by simp
show "E j \<cdot>\<^sub>C ?\<chi>' (J.dom j) = ?\<chi>' j"
proof -
have "E j \<cdot>\<^sub>C ?\<chi>' (J.dom j) = (E j \<cdot>\<^sub>C Adj.\<epsilon> (E (J.dom j))) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using j C.comp_assoc by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "(E j \<cdot>\<^sub>C Adj.\<epsilon> (E (J.dom j))) \<cdot>\<^sub>C F (\<kappa> (J.dom j))
= (Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C Adj.FG.map (E j)) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using j Adj.\<epsilon>.naturality [of "E j"] by fastforce
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j))"
using C.comp_assoc by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j)) = F (GE.map j \<cdot>\<^sub>D \<kappa> (J.dom j))"
using j by simp
hence "Adj.FG.map (E j) \<cdot>\<^sub>C F (\<kappa> (J.dom j)) = F (\<kappa> j)"
using j \<kappa>.is_natural_1 by metis
thus ?thesis using j by simp
qed
finally show ?thesis by auto
qed
also have "... = ?\<chi>' j"
using j by simp
finally show ?thesis by auto
qed
show "?\<chi>' (J.cod j) \<cdot>\<^sub>C Fb.map j = ?\<chi>' j"
proof -
have "?\<chi>' (J.cod j) \<cdot>\<^sub>C Fb.map j = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> (J.cod j))"
using j Fb.value_is_ide Adj.\<epsilon>.preserves_hom C.comp_arr_dom [of "F (\<kappa> (J.cod j))"]
C.comp_assoc
by simp
also have "... = Adj.\<epsilon> (E (J.cod j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j \<kappa>.is_natural_1 \<kappa>.is_natural_2 Adj.\<epsilon>.naturality J.arr_cod_iff_arr
by (metis J.cod_cod \<kappa>.A.map_simp)
also have "... = ?\<chi>' j" using j by simp
finally show ?thesis by auto
qed
qed
text\<open>
Using the universal property of the limit cone @{term \<chi>}, obtain the unique arrow
@{term f} that transforms @{term \<chi>} into @{term \<chi>'}.
\<close>
from this \<chi>.is_universal [of "F b" ?\<chi>'] obtain f
where f: "\<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'"
by auto
text\<open>
Let @{term g} be the adjunct of @{term f}, and show that @{term g} transforms
@{term G\<chi>} into @{term \<kappa>}.
\<close>
let ?g = "G f \<cdot>\<^sub>D Adj.\<eta> b"
have 1: "\<guillemotleft>?g : b \<rightarrow>\<^sub>D G a\<guillemotright>" using f \<kappa>.ide_apex by fastforce
moreover have "GE.cones_map ?g ?G\<chi> = \<kappa>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> GE.cones_map ?g ?G\<chi> j = \<kappa> j"
using 1 G\<chi>.cone_axioms \<kappa>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> GE.cones_map ?g ?G\<chi> j = \<kappa> j"
proof -
fix j
assume j: "J.arr j"
have "GE.cones_map ?g ?G\<chi> j = G (\<chi> j) \<cdot>\<^sub>D ?g"
using j 1 G\<chi>.cone_axioms mem_Collect_eq restrict_apply by auto
also have "... = G (\<chi> j \<cdot>\<^sub>C f) \<cdot>\<^sub>D Adj.\<eta> b"
using j f \<chi>.preserves_hom [of j "J.dom j" "J.cod j"] D.comp_assoc by fastforce
also have "... = G (E.cones_map f \<chi> j) \<cdot>\<^sub>D Adj.\<eta> b"
proof -
have "\<chi> j \<cdot>\<^sub>C f = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
proof -
have "E.cone (C.cod f) \<chi>"
using f \<chi>.cone_axioms by blast
hence "\<chi> j \<cdot>\<^sub>C f = E.cones_map f \<chi> j"
using \<chi>.is_extensional by simp
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j f by simp
finally show ?thesis by blast
qed
thus ?thesis
using f mem_Collect_eq restrict_apply Adj.F.is_extensional by simp
qed
also have "... = (G (Adj.\<epsilon> (C.cod (E j))) \<cdot>\<^sub>D Adj.\<eta> (D.cod (GE.map j))) \<cdot>\<^sub>D \<kappa> j"
using j f Adj.\<eta>.naturality [of "\<kappa> j"] D.comp_assoc by auto
also have "... = D.cod (\<kappa> j) \<cdot>\<^sub>D \<kappa> j"
using j Adj.\<eta>\<epsilon>.triangle_G Adj.\<epsilon>_in_terms_of_\<psi> Adj.\<epsilon>o_def
Adj.\<eta>_in_terms_of_\<phi> Adj.\<eta>o_def Adj.unit_counit_G
by fastforce
also have "... = \<kappa> j"
using j D.comp_cod_arr by simp
finally show "GE.cones_map ?g ?G\<chi> j = \<kappa> j" by metis
qed
ultimately show "GE.cones_map ?g ?G\<chi> j = \<kappa> j" by auto
qed
ultimately have "\<guillemotleft>?g : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map ?g ?G\<chi> = \<kappa>" by auto
text\<open>
It remains to be shown that @{term g} is the unique such arrow.
Given any @{term g'} that transforms @{term G\<chi>} into @{term \<kappa>},
its adjunct transforms @{term \<chi>} into @{term \<chi>'}.
The adjunct of @{term g'} is therefore equal to @{term f},
which implies @{term g'} = @{term g}.
\<close>
moreover have "\<And>g'. \<guillemotleft>g' : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g' ?G\<chi> = \<kappa> \<Longrightarrow> g' = ?g"
proof -
fix g'
assume g': "\<guillemotleft>g' : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g' ?G\<chi> = \<kappa>"
have 1: "\<guillemotleft>\<psi> a g' : F b \<rightarrow>\<^sub>C a\<guillemotright>"
using g' a \<psi>_in_hom by simp
have 2: "E.cones_map (\<psi> a g') \<chi> = ?\<chi>'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j"
using 1 \<chi>.cone_axioms by auto
moreover have "J.arr j \<Longrightarrow> E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j"
proof -
fix j
assume j: "J.arr j"
have "E.cones_map (\<psi> a g') \<chi> j = \<chi> j \<cdot>\<^sub>C \<psi> a g'"
using 1 \<chi>.cone_axioms \<chi>.is_extensional by auto
also have "... = (\<chi> j \<cdot>\<^sub>C Adj.\<epsilon> a) \<cdot>\<^sub>C F g'"
using j a g' Adj.\<psi>_in_terms_of_\<epsilon> C.comp_assoc Adj.\<epsilon>_def by auto
also have "... = (Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (G (\<chi> j))) \<cdot>\<^sub>C F g'"
using j a g' Adj.\<epsilon>.naturality [of "\<chi> j"] by simp
also have "... = Adj.\<epsilon> (C.cod (E j)) \<cdot>\<^sub>C F (\<kappa> j)"
using j a g' G\<chi>.cone_axioms C.comp_assoc by auto
finally show "E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j" by (simp add: j)
qed
ultimately show "E.cones_map (\<psi> a g') \<chi> j = ?\<chi>' j" by auto
qed
have "\<psi> a g' = f"
proof -
have "\<exists>!f. \<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'"
using cone_\<chi>' \<chi>.is_universal by simp
moreover have "\<guillemotleft>\<psi> a g' : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map (\<psi> a g') \<chi> = ?\<chi>'"
using 1 2 by simp
ultimately show ?thesis
using ex1E [of "\<lambda>f. \<guillemotleft>f : F b \<rightarrow>\<^sub>C a\<guillemotright> \<and> E.cones_map f \<chi> = ?\<chi>'" "\<psi> a g' = f"]
using 1 2 Adj.\<epsilon>.is_extensional C.comp_null(2) C.ex_un_null \<chi>.cone_axioms f
mem_Collect_eq restrict_apply
by blast
qed
hence "\<phi> b (\<psi> a g') = \<phi> b f" by auto
hence "g' = \<phi> b f" using \<chi>.ide_apex g' by (simp add: \<phi>_\<psi>)
moreover have "?g = \<phi> b f" using f Adj.\<phi>_in_terms_of_\<eta> \<kappa>.ide_apex Adj.\<eta>_def by auto
ultimately show "g' = ?g" by argo
qed
ultimately show "\<exists>!g. \<guillemotleft>g : b \<rightarrow>\<^sub>D G a\<guillemotright> \<and> GE.cones_map g ?G\<chi> = \<kappa>" by blast
qed
have "GE.limit_cone (G a) ?G\<chi>" ..
thus ?thesis by auto
qed
end
section "Special Kinds of Limits"
subsection "Terminal Objects"
text\<open>
An object of a category @{term C} is a terminal object if and only if it is a limit of the
empty diagram in @{term C}.
\<close>
locale empty_diagram =
diagram J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c" +
assumes is_empty: "\<not>J.arr j"
begin
lemma has_as_limit_iff_terminal:
shows "has_as_limit a \<longleftrightarrow> C.terminal a"
proof
assume a: "has_as_limit a"
show "C.terminal a"
proof
have "\<exists>\<chi>. limit_cone a \<chi>" using a by auto
from this obtain \<chi> where \<chi>: "limit_cone a \<chi>" by blast
interpret \<chi>: limit_cone J C D a \<chi> using \<chi> by auto
have cone_\<chi>: "cone a \<chi>" ..
show "C.ide a" using \<chi>.ide_apex by auto
have 1: "\<chi> = (\<lambda>j. C.null)" using is_empty \<chi>.is_extensional by auto
show "\<And>a'. C.ide a' \<Longrightarrow> \<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>"
proof -
fix a'
assume a': "C.ide a'"
interpret A': constant_functor J C a'
apply unfold_locales using a' by auto
let ?\<chi>' = "\<lambda>j. C.null"
have cone_\<chi>': "cone a' ?\<chi>'"
using a' is_empty apply unfold_locales by auto
hence "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f \<chi> = ?\<chi>'"
using \<chi>.is_universal by force
moreover have "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> cones_map f \<chi> = ?\<chi>'"
using 1 cone_\<chi> by auto
ultimately show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" by blast
qed
qed
next
assume a: "C.terminal a"
show "has_as_limit a"
proof -
let ?\<chi> = "\<lambda>j. C.null"
have "C.ide a" using a C.terminal_def by simp
interpret A: constant_functor J C a
apply unfold_locales using \<open>C.ide a\<close> by simp
interpret \<chi>: cone J C D a ?\<chi>
using \<open>C.ide a\<close> is_empty by (unfold_locales, auto)
have cone_\<chi>: "cone a ?\<chi>" ..
have 1: "\<And>a' \<chi>'. cone a' \<chi>' \<Longrightarrow> \<chi>' = (\<lambda>j. C.null)"
proof -
fix a' \<chi>'
assume \<chi>': "cone a' \<chi>'"
interpret \<chi>': cone J C D a' \<chi>' using \<chi>' by auto
show "\<chi>' = (\<lambda>j. C.null)"
using is_empty \<chi>'.is_extensional by metis
qed
have "limit_cone a ?\<chi>"
proof
fix a' \<chi>'
assume \<chi>': "cone a' \<chi>'"
have 2: "\<chi>' = (\<lambda>j. C.null)" using 1 \<chi>' by simp
interpret \<chi>': cone J C D a' \<chi>' using \<chi>' by auto
have "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" using a C.terminal_def \<chi>'.ide_apex by simp
moreover have "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<Longrightarrow> cones_map f ?\<chi> = \<chi>'"
using 1 2 cones_map_mapsto cone_\<chi> \<chi>'.cone_axioms mem_Collect_eq by blast
ultimately show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f (\<lambda>j. C.null) = \<chi>'"
by blast
qed
thus ?thesis by auto
qed
qed
end
subsection "Products"
text\<open>
A \emph{product} in a category @{term C} is a limit of a discrete diagram in @{term C}.
\<close>
locale discrete_diagram =
J: category J +
diagram J C D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c" +
assumes is_discrete: "J.arr = J.ide"
begin
abbreviation mkCone
where "mkCone F \<equiv> (\<lambda>j. if J.arr j then F j else C.null)"
lemma cone_mkCone:
assumes "C.ide a" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j : a \<rightarrow> D j\<guillemotright>"
shows "cone a (mkCone F)"
proof -
interpret A: constant_functor J C a
apply unfold_locales using assms(1) by auto
show "cone a (mkCone F)"
using assms(2) is_discrete
apply unfold_locales
apply auto
apply (metis C.in_homE C.comp_cod_arr)
using C.comp_arr_ide by fastforce
qed
lemma mkCone_cone:
assumes "cone a \<pi>"
shows "mkCone \<pi> = \<pi>"
proof -
interpret \<pi>: cone J C D a \<pi>
using assms by auto
show "mkCone \<pi> = \<pi>" using \<pi>.is_extensional by auto
qed
end
text\<open>
The following locale defines a discrete diagram in a category @{term C},
given an index set @{term I} and a function @{term D} mapping @{term I}
to objects of @{term C}. Here we obtain the diagram shape @{term J}
using a discrete category construction that allows us to directly identify
the objects of @{term J} with the elements of @{term I}, however this construction
can only be applied in case the set @{term I} is not the universe of its
element type.
\<close>
locale discrete_diagram_from_map =
J: discrete_category I null +
C: category C
for I :: "'i set"
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'i \<Rightarrow> 'c"
and null :: 'i +
assumes maps_to_ide: "i \<in> I \<Longrightarrow> C.ide (D i)"
begin
definition map
where "map j \<equiv> if J.arr j then D j else C.null"
end
sublocale discrete_diagram_from_map \<subseteq> discrete_diagram J.comp C map
using map_def maps_to_ide J.arr_char J.Null_not_in_Obj J.null_char
by (unfold_locales, auto)
locale product_cone =
J: category J +
C: category C +
D: discrete_diagram J C D +
limit_cone J C D a \<pi>
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and C :: "'c comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 'c"
and a :: 'c
and \<pi> :: "'j \<Rightarrow> 'c"
begin
lemma is_cone:
shows "D.cone a \<pi>" ..
text\<open>
The following versions of @{prop is_universal} and @{prop induced_arrowI}
from the \<open>limit_cone\<close> locale are specialized to the case in which the
underlying diagram is a product diagram.
\<close>
lemma is_universal':
assumes "C.ide b" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j: b \<rightarrow> D j\<guillemotright>"
shows "\<exists>!f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<and> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof -
let ?\<chi> = "D.mkCone F"
interpret B: constant_functor J C b
apply unfold_locales using assms(1) by auto
have cone_\<chi>: "D.cone b ?\<chi>"
using assms D.is_discrete
apply unfold_locales
apply auto
apply (meson C.comp_ide_arr C.ide_in_hom C.seqI' D.preserves_ide)
using C.comp_arr_dom by blast
interpret \<chi>: cone J C D b ?\<chi> using cone_\<chi> by auto
have "\<exists>!f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using cone_\<chi> is_universal by force
moreover have
"\<And>f. \<guillemotleft>f : b \<rightarrow> a\<guillemotright> \<Longrightarrow> D.cones_map f \<pi> = ?\<chi> \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof -
fix f
assume f: "\<guillemotleft>f : b \<rightarrow> a\<guillemotright>"
show "D.cones_map f \<pi> = ?\<chi> \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j)"
proof
assume 1: "D.cones_map f \<pi> = ?\<chi>"
show "\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j"
proof -
have "\<And>j. J.arr j \<Longrightarrow> \<pi> j \<cdot> f = F j"
proof -
fix j
assume j: "J.arr j"
have "\<pi> j \<cdot> f = D.cones_map f \<pi> j"
using j f cone_axioms by force
also have "... = F j" using j 1 by simp
finally show "\<pi> j \<cdot> f = F j" by auto
qed
thus ?thesis by auto
qed
next
assume 1: "\<forall>j. J.arr j \<longrightarrow> \<pi> j \<cdot> f = F j"
show "D.cones_map f \<pi> = ?\<chi>"
using 1 f is_cone \<chi>.is_extensional D.is_discrete is_cone cone_\<chi> by auto
qed
qed
ultimately show ?thesis by blast
qed
abbreviation induced_arrow' :: "'c \<Rightarrow> ('j \<Rightarrow> 'c) \<Rightarrow> 'c"
where "induced_arrow' b F \<equiv> induced_arrow b (D.mkCone F)"
lemma induced_arrowI':
assumes "C.ide b" and "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>F j : b \<rightarrow> D j\<guillemotright>"
shows "\<And>j. J.arr j \<Longrightarrow> \<pi> j \<cdot> induced_arrow' b F = F j"
proof -
interpret B: constant_functor J C b
apply unfold_locales using assms(1) by auto
interpret \<chi>: cone J C D b \<open>D.mkCone F\<close>
using assms D.cone_mkCone by blast
have cone_\<chi>: "D.cone b (D.mkCone F)" ..
hence 1: "D.cones_map (induced_arrow' b F) \<pi> = D.mkCone F"
using induced_arrowI by blast
fix j
assume j: "J.arr j"
have "\<pi> j \<cdot> induced_arrow' b F = D.cones_map (induced_arrow' b F) \<pi> j"
using induced_arrowI(1) cone_\<chi> is_cone is_extensional by force
also have "... = F j"
using j 1 by auto
finally show "\<pi> j \<cdot> induced_arrow' b F = F j"
by auto
qed
end
context discrete_diagram
begin
lemma product_coneI:
assumes "limit_cone a \<pi>"
shows "product_cone J C D a \<pi>"
proof -
interpret L: limit_cone J C D a \<pi>
using assms by auto
show "product_cone J C D a \<pi>" ..
qed
end
context category
begin
definition has_as_product
where "has_as_product J D a \<equiv> (\<exists>\<pi>. product_cone J C D a \<pi>)"
lemma product_is_ide:
assumes "has_as_product J D a"
shows "ide a"
proof -
obtain \<pi> where \<pi>: "product_cone J C D a \<pi>"
using assms has_as_product_def by blast
interpret \<pi>: product_cone J C D a \<pi>
using \<pi> by auto
show ?thesis using \<pi>.ide_apex by auto
qed
text\<open>
A category has @{term I}-indexed products for an @{typ 'i}-set @{term I}
if every @{term I}-indexed discrete diagram has a product.
In order to reap the benefits of being able to directly identify the elements
of a set I with the objects of discrete category it generates (thereby avoiding
the use of coercion maps), it is necessary to assume that @{term "I \<noteq> UNIV"}.
If we want to assert that a category has products indexed by the universe of
some type @{typ 'i}, we have to pass to a larger type, such as @{typ "'i option"}.
\<close>
definition has_products
where "has_products (I :: 'i set) \<equiv>
I \<noteq> UNIV \<and>
(\<forall>J D. discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I
\<longrightarrow> (\<exists>a. has_as_product J D a))"
lemma ex_productE:
assumes "\<exists>a. has_as_product J D a"
obtains a \<pi> where "product_cone J C D a \<pi>"
using assms has_as_product_def someI_ex [of "\<lambda>a. has_as_product J D a"] by metis
lemma has_products_if_has_limits:
assumes "has_limits (undefined :: 'j)" and "I \<noteq> (UNIV :: 'j set)"
shows "has_products I"
proof (unfold has_products_def, intro conjI allI impI)
show "I \<noteq> UNIV" by fact
fix J D
assume D: "discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I"
interpret D: discrete_diagram J C D
using D by simp
have 1: "\<exists>a. D.has_as_limit a"
using assms D D.diagram_axioms D.J.category_axioms
by (simp add: has_limits_of_shape_def has_limits_def)
show "\<exists>a. has_as_product J D a"
using 1 has_as_product_def D.product_coneI by blast
qed
lemma has_finite_products_if_has_finite_limits:
assumes "\<And>J :: 'j comp. (finite (Collect (partial_magma.arr J))) \<Longrightarrow> has_limits_of_shape J"
and "finite (I :: 'j set)" and "I \<noteq> UNIV"
shows "has_products I"
proof (unfold has_products_def, intro conjI allI impI)
show "I \<noteq> UNIV" by fact
fix J D
assume D: "discrete_diagram J C D \<and> Collect (partial_magma.arr J) = I"
interpret D: discrete_diagram J C D
using D by simp
have 1: "\<exists>a. D.has_as_limit a"
using assms D has_limits_of_shape_def D.diagram_axioms by auto
show "\<exists>a. has_as_product J D a"
using 1 has_as_product_def D.product_coneI by blast
qed
lemma has_products_preserved_by_bijection:
assumes "has_products I" and "bij_betw \<phi> I I'" and "I' \<noteq> UNIV"
shows "has_products I'"
proof (unfold has_products_def, intro conjI allI impI)
show "I' \<noteq> UNIV" by fact
show "\<And>J' D'. discrete_diagram J' C D' \<and> Collect (partial_magma.arr J') = I'
\<Longrightarrow> \<exists>a. has_as_product J' D' a"
proof -
fix J' D'
assume 1: "discrete_diagram J' C D' \<and> Collect (partial_magma.arr J') = I'"
interpret J': category J'
using 1 by (simp add: discrete_diagram_def)
interpret D': discrete_diagram J' C D'
using 1 by simp
interpret J: discrete_category I \<open>SOME x. x \<notin> I\<close>
using assms has_products_def [of I] someI_ex [of "\<lambda>x. x \<notin> I"]
by unfold_locales auto
have 2: "Collect J.arr = I \<and> Collect J'.arr = I'"
using 1 by auto
have \<phi>: "bij_betw \<phi> (Collect J.arr) (Collect J'.arr)"
using 2 assms(2) by simp
let ?\<phi> = "\<lambda>j. if J.arr j then \<phi> j else J'.null"
let ?\<phi>' = "\<lambda>j'. if J'.arr j' then the_inv_into I \<phi> j' else J.null"
interpret \<phi>: "functor" J.comp J' ?\<phi>
proof -
have "\<phi> ` I = I'"
using \<phi> 2 bij_betw_def [of \<phi> I I'] by simp
hence "\<And>j. J.arr j \<Longrightarrow> J'.arr (?\<phi> j)"
using 1 D'.is_discrete by auto
thus "functor J.comp J' ?\<phi>"
using D'.is_discrete J.is_discrete J.seqE
by unfold_locales auto
qed
interpret \<phi>': "functor" J' J.comp ?\<phi>'
proof -
have "the_inv_into I \<phi> ` I' = I"
using assms(2) \<phi> bij_betw_the_inv_into bij_betw_imp_surj_on by metis
hence "\<And>j'. J'.arr j' \<Longrightarrow> J.arr (?\<phi>' j')"
using 2 D'.is_discrete J.is_discrete by auto
thus "functor J' J.comp ?\<phi>'"
using D'.is_discrete J.is_discrete J'.seqE
by unfold_locales auto
qed
let ?D = "\<lambda>i. D' (\<phi> i)"
interpret D: discrete_diagram_from_map I C ?D \<open>SOME j. j \<notin> I\<close>
using assms 1 D'.is_discrete bij_betw_imp_surj_on \<phi>.preserves_ide
by unfold_locales auto
obtain a where a: "has_as_product J.comp D.map a"
using assms D.discrete_diagram_axioms has_products_def [of I] by auto
obtain \<pi> where \<pi>: "product_cone J.comp C D.map a \<pi>"
using a has_as_product_def by blast
interpret \<pi>: product_cone J.comp C D.map a \<pi>
using \<pi> by simp
let ?\<pi>' = "\<pi> o ?\<phi>'"
interpret A: constant_functor J' C a
using \<pi>.ide_apex by unfold_locales simp
interpret \<pi>': natural_transformation J' C A.map D' ?\<pi>'
proof -
have "\<pi>.A.map \<circ> ?\<phi>' = A.map"
using \<phi> A.map_def \<phi>'.preserves_arr \<pi>.A.is_extensional J.not_arr_null by auto
moreover have "D.map \<circ> ?\<phi>' = D'"
proof
fix j'
have "J'.arr j' \<Longrightarrow> (D.map \<circ> ?\<phi>') j' = D' j'"
proof -
assume 2: "J'.arr j'"
have 3: "inj_on \<phi> I"
using assms(2) bij_betw_imp_inj_on by auto
have "\<phi> ` I = I'"
by (metis (no_types) assms(2) bij_betw_imp_surj_on)
hence "\<phi> ` I = Collect J'.arr"
using 1 by force
thus ?thesis
using 2 3 D.map_def \<phi>'.preserves_arr f_the_inv_into_f by fastforce
qed
moreover have "\<not> J'.arr j' \<Longrightarrow> (D.map \<circ> ?\<phi>') j' = D' j'"
using D.is_extensional D'.is_extensional
by (simp add: J.Null_not_in_Obj J.null_char)
ultimately show "(D.map \<circ> ?\<phi>') j' = D' j'" by blast
qed
ultimately show "natural_transformation J' C A.map D' ?\<pi>'"
using \<pi>.natural_transformation_axioms \<phi>'.natural_transformation_axioms
horizontal_composite [of J' J.comp ?\<phi>' ?\<phi>' ?\<phi>' C \<pi>.A.map D.map \<pi>]
by simp
qed
interpret \<pi>': cone J' C D' a ?\<pi>' ..
interpret \<pi>': product_cone J' C D' a ?\<pi>'
proof
fix a' \<chi>'
assume \<chi>': "D'.cone a' \<chi>'"
interpret \<chi>': cone J' C D' a' \<chi>'
using \<chi>' by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f (\<pi> \<circ> ?\<phi>') = \<chi>'"
proof -
let ?\<chi> = "\<chi>' o ?\<phi>"
interpret A': constant_functor J.comp C a'
using \<chi>'.ide_apex by unfold_locales simp
interpret \<chi>: natural_transformation J.comp C A'.map D.map ?\<chi>
proof -
have "\<chi>'.A.map \<circ> ?\<phi> = A'.map"
using \<phi> \<phi>.preserves_arr A'.map_def \<chi>'.A.is_extensional by auto
moreover have "D' \<circ> ?\<phi> = D.map"
using \<phi> D.map_def D'.is_extensional by auto
ultimately show "natural_transformation J.comp C A'.map D.map ?\<chi>"
using \<chi>'.natural_transformation_axioms \<phi>.natural_transformation_axioms
horizontal_composite [of J.comp J' ?\<phi> ?\<phi> ?\<phi> C \<chi>'.A.map D' \<chi>']
by simp
qed
interpret \<chi>: cone J.comp C D.map a' ?\<chi> ..
have *: "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using \<pi>.is_universal \<chi>.cone_axioms by simp
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f ?\<pi>' = \<chi>'"
proof -
have "\<exists>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D'.cones_map f ?\<pi>' = \<chi>'"
proof -
obtain f where f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f \<pi> = ?\<chi>"
using * by blast
have "D'.cones_map f ?\<pi>' = \<chi>'"
proof
fix j'
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
proof (cases "J'.arr j'")
assume j': "\<not> J'.arr j'"
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
using f j' \<chi>'.is_extensional \<pi>'.cone_axioms by auto
next
assume j': "J'.arr j'"
show "D'.cones_map f ?\<pi>' j' = \<chi>' j'"
proof -
have "D'.cones_map f ?\<pi>' j' = \<pi> (the_inv_into I \<phi> j') \<cdot> f"
using f j' \<pi>'.cone_axioms by auto
also have "... = D.cones_map f \<pi> (the_inv_into I \<phi> j')"
proof -
have "arr f \<and> dom f = a' \<and> cod f = a"
using f by blast
thus ?thesis
using \<phi>'.preserves_arr \<pi>.cone_\<chi> j' by auto
qed
also have "... = (\<chi>' \<circ> ?\<phi>) (the_inv_into I \<phi> j')"
using f by simp
also have "... = \<chi>' j'"
using assms(2) j' 2 bij_betw_def [of \<phi> I I'] bij_betw_imp_inj_on
\<phi>'.preserves_arr f_the_inv_into_f
by fastforce
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using f by blast
qed
moreover have "\<And>f f'. \<lbrakk> \<guillemotleft>f : a' \<rightarrow> a\<guillemotright>; D'.cones_map f ?\<pi>' = \<chi>';
\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>; D'.cones_map f' ?\<pi>' = \<chi>' \<rbrakk>
\<Longrightarrow> f = f'"
proof -
fix f f'
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright>" and f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>"
and f\<chi>': "D'.cones_map f ?\<pi>' = \<chi>'" and f'\<chi>': "D'.cones_map f' ?\<pi>' = \<chi>'"
have "D.cones_map f \<pi> = \<chi>' \<circ> ?\<phi> \<and> D.cones_map f' \<pi> = \<chi>' o ?\<phi>"
proof (intro conjI)
show "D.cones_map f \<pi> = \<chi>' \<circ> ?\<phi>"
proof
fix j
have "\<not> J.arr j \<Longrightarrow> D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
using f f\<chi>' \<pi>.cone_axioms \<chi>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
proof -
assume j: "J.arr j"
have 1: "j = the_inv_into I \<phi> (\<phi> j)"
using assms(2) j \<phi> the_inv_into_f_f bij_betw_imp_inj_on J.arr_char
by metis
have "D.cones_map f \<pi> j = D.cones_map f \<pi> (the_inv_into I \<phi> (\<phi> j))"
using 1 by simp
also have "... = (\<chi>' \<circ> ?\<phi>) j"
using f j f\<chi>' 1 \<pi>.cone_axioms \<pi>'.cone_axioms \<phi>.preserves_arr by auto
finally show "D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
ultimately show "D.cones_map f \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
show "D.cones_map f' \<pi> = \<chi>' \<circ> ?\<phi>"
proof
fix j
have "\<not> J.arr j \<Longrightarrow> D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
using f' f\<chi>' \<pi>.cone_axioms \<chi>.is_extensional by auto
moreover have "J.arr j \<Longrightarrow> D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j"
proof -
assume j: "J.arr j"
have 1: "j = the_inv_into I \<phi> (\<phi> j)"
using assms(2) j \<phi> the_inv_into_f_f bij_betw_imp_inj_on J.arr_char
by metis
have "D.cones_map f' \<pi> j = D.cones_map f' \<pi> (the_inv_into I \<phi> (\<phi> j))"
using 1 by simp
also have "... = (\<chi>' \<circ> ?\<phi>) j"
using f' j f'\<chi>' 1 \<pi>.cone_axioms \<pi>'.cone_axioms \<phi>.preserves_arr by auto
finally show "D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
ultimately show "D.cones_map f' \<pi> j = (\<chi>' \<circ> ?\<phi>) j" by blast
qed
qed
thus "f = f'"
using f f' * by auto
qed
ultimately show ?thesis by blast
qed
qed
qed
have "has_as_product J' D' a"
using has_as_product_def \<pi>'.product_cone_axioms by auto
thus "\<exists>a. has_as_product J' D' a" by blast
qed
qed
lemma ide_is_unary_product:
assumes "ide a"
shows "\<And>m n :: nat. m \<noteq> n \<Longrightarrow> has_as_product (discrete_category.comp {m :: nat} (n :: nat))
(\<lambda>i. if i = m then a else null) a"
proof -
fix m n :: nat
assume neq: "m \<noteq> n"
have "{m :: nat} \<noteq> UNIV"
proof -
have "finite {m :: nat}" by simp
moreover have "\<not>finite (UNIV :: nat set)" by simp
ultimately show ?thesis by fastforce
qed
interpret J: discrete_category "{m :: nat}" \<open>n :: nat\<close>
- using neq `{m :: nat} \<noteq> UNIV` by unfold_locales auto
+ using neq \<open>{m :: nat} \<noteq> UNIV\<close> by unfold_locales auto
let ?D = "\<lambda>i. if i = m then a else null"
interpret D: discrete_diagram J.comp C ?D
apply unfold_locales
using assms J.null_char neq
apply auto
by metis
interpret A: constant_functor J.comp C a
using assms by unfold_locales auto
show "has_as_product J.comp ?D a"
proof (unfold has_as_product_def)
let ?\<pi> = "\<lambda>i :: nat. if i = m then a else null"
interpret \<pi>: natural_transformation J.comp C A.map ?D ?\<pi>
using assms J.arr_char J.dom_char J.cod_char
by unfold_locales auto
interpret \<pi>: cone J.comp C ?D a ?\<pi> ..
interpret \<pi>: product_cone J.comp C ?D a ?\<pi>
proof
fix a' \<chi>'
assume \<chi>': "D.cone a' \<chi>'"
interpret \<chi>': cone J.comp C ?D a' \<chi>' using \<chi>' by auto
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
proof
show "\<guillemotleft>\<chi>' m : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map (\<chi>' m) ?\<pi> = \<chi>'"
proof
show 1: "\<guillemotleft>\<chi>' m : a' \<rightarrow> a\<guillemotright>"
using \<chi>'.preserves_hom \<chi>'.A.map_def J.arr_char J.dom_char J.cod_char
by auto
show "D.cones_map (\<chi>' m) ?\<pi> = \<chi>'"
proof
fix j
show "D.cones_map (\<chi>' m) (\<lambda>i. if i = m then a else null) j = \<chi>' j"
using J.arr_char J.dom_char J.cod_char J.ide_char \<pi>.cone_axioms comp_cod_arr
apply (cases "j = m")
apply simp
using \<chi>'.is_extensional by simp
qed
qed
show "\<And>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>' \<Longrightarrow> f = \<chi>' m"
proof -
fix f
assume f: "\<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map f ?\<pi> = \<chi>'"
show "f = \<chi>' m"
using assms f \<chi>'.preserves_hom J.arr_char J.dom_char J.cod_char \<pi>.cone_axioms
comp_cod_arr
by auto
qed
qed
qed
show "\<exists>\<pi>. product_cone J.comp C (\<lambda>i. if i = m then a else null) a \<pi>"
using \<pi>.product_cone_axioms by blast
qed
qed
lemma has_unary_products:
assumes "card I = 1" and "I \<noteq> UNIV"
shows "has_products I"
proof -
obtain i where i: "I = {i}"
using assms card_1_singletonE by blast
obtain n where n: "n \<notin> I"
using assms by auto
have "has_products {1 :: nat}"
proof (unfold has_products_def, intro conjI)
show "{1 :: nat} \<noteq> UNIV" by auto
show "\<forall>(J :: nat comp) D.
discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}
\<longrightarrow> (\<exists>a. has_as_product J D a)"
proof
fix J :: "nat comp"
show "\<forall>D. discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}
\<longrightarrow> (\<exists>a. has_as_product J D a)"
proof (intro allI impI)
fix D
assume D: "discrete_diagram J (\<cdot>) D \<and> Collect (partial_magma.arr J) = {1}"
interpret D: discrete_diagram J C D
using D by auto
interpret J: discrete_category \<open>{1 :: nat}\<close> D.J.null
by (metis D D.J.not_arr_null discrete_category.intro mem_Collect_eq)
have 1: "has_as_product J.comp D (D 1)"
proof -
have "has_as_product J.comp (\<lambda>i. if i = 1 then D 1 else null) (D 1)"
using ide_is_unary_product D.preserves_ide D.is_discrete D J.null_char
by (metis J.Null_not_in_Obj insertCI mem_Collect_eq)
moreover have "D = (\<lambda>i. if i = 1 then D 1 else null)"
proof
fix j
show "D j = (if j = 1 then D 1 else null)"
using D D.is_extensional by auto
qed
ultimately show ?thesis by simp
qed
moreover have "J = J.comp"
proof -
have "\<And>j j'. J j j' = J.comp j j'"
proof -
fix j j'
show "J j j' = J.comp j j'"
using J.comp_char D.is_discrete D
by (metis D.J.category_axioms D.J.ext D.J.ide_def J.null_char
category.seqE mem_Collect_eq)
qed
thus ?thesis by auto
qed
ultimately show "\<exists>a. has_as_product J D a" by auto
qed
qed
qed
moreover have "bij_betw (\<lambda>k. if k = 1 then i else n) {1 :: nat} I"
using i by (intro bij_betwI') auto
ultimately show "has_products I"
using assms has_products_preserved_by_bijection by blast
qed
end
subsection "Equalizers"
text\<open>
An \emph{equalizer} in a category @{term C} is a limit of a parallel pair
of arrows in @{term C}.
\<close>
locale parallel_pair_diagram =
J: parallel_pair +
C: category C
for C :: "'c comp" (infixr "\<cdot>" 55)
and f0 :: 'c
and f1 :: 'c +
assumes is_parallel: "C.par f0 f1"
begin
no_notation J.comp (infixr "\<cdot>" 55)
notation J.comp (infixr "\<cdot>\<^sub>J" 55)
definition map
where "map \<equiv> (\<lambda>j. if j = J.Zero then C.dom f0
else if j = J.One then C.cod f0
else if j = J.j0 then f0
else if j = J.j1 then f1
else C.null)"
lemma map_simp:
shows "map J.Zero = C.dom f0"
and "map J.One = C.cod f0"
and "map J.j0 = f0"
and "map J.j1 = f1"
proof -
show "map J.Zero = C.dom f0"
using map_def by metis
show "map J.One = C.cod f0"
using map_def J.Zero_not_eq_One by metis
show "map J.j0 = f0"
using map_def J.Zero_not_eq_j0 J.One_not_eq_j0 by metis
show "map J.j1 = f1"
using map_def J.Zero_not_eq_j1 J.One_not_eq_j1 J.j0_not_eq_j1 by metis
qed
end
sublocale parallel_pair_diagram \<subseteq> diagram J.comp C map
apply unfold_locales
apply (simp add: J.arr_char map_def)
using map_def is_parallel J.arr_char J.cod_simp J.dom_simp
apply auto[2]
proof -
show 1: "\<And>j. J.arr j \<Longrightarrow> C.cod (map j) = map (J.cod j)"
proof -
fix j
assume j: "J.arr j"
show "C.cod (map j) = map (J.cod j)"
proof -
have "j = J.Zero \<or> j = J.One \<Longrightarrow> ?thesis" using is_parallel map_def by auto
moreover have "j = J.j0 \<or> j = J.j1 \<Longrightarrow> ?thesis"
using is_parallel map_def J.Zero_not_eq_j0 J.One_not_eq_j0 J.Zero_not_eq_One
J.Zero_not_eq_j1 J.One_not_eq_j1 J.Zero_not_eq_One J.cod_simp
by presburger
ultimately show ?thesis using j J.arr_char by fast
qed
qed
next
fix j j'
assume jj': "J.seq j' j"
show "map (j' \<cdot>\<^sub>J j) = map j' \<cdot> map j"
proof -
have 1: "(j = J.Zero \<and> j' \<noteq> J.One) \<or> (j \<noteq> J.Zero \<and> j' = J.One)"
using jj' J.seq_char by blast
moreover have "j = J.Zero \<and> j' \<noteq> J.One \<Longrightarrow> ?thesis"
using jj' map_def is_parallel J.arr_char J.cod_simp J.dom_simp J.seq_char
by (metis (no_types, lifting) C.arr_dom_iff_arr C.comp_arr_dom C.dom_dom
J.comp_arr_dom)
moreover have "j \<noteq> J.Zero \<and> j' = J.One \<Longrightarrow> ?thesis"
using jj' J.ide_char map_def J.Zero_not_eq_One is_parallel
by (metis (no_types, lifting) C.arr_cod_iff_arr C.comp_arr_dom C.comp_cod_arr
C.comp_ide_arr C.ext C.ide_cod J.comp_simp(2))
ultimately show ?thesis by blast
qed
qed
context parallel_pair_diagram
begin
definition mkCone
where "mkCone e \<equiv> \<lambda>j. if J.arr j then if j = J.Zero then e else f0 \<cdot> e else C.null"
abbreviation is_equalized_by
where "is_equalized_by e \<equiv> C.seq f0 e \<and> f0 \<cdot> e = f1 \<cdot> e"
abbreviation has_as_equalizer
where "has_as_equalizer e \<equiv> limit_cone (C.dom e) (mkCone e)"
lemma cone_mkCone:
assumes "is_equalized_by e"
shows "cone (C.dom e) (mkCone e)"
proof -
interpret E: constant_functor J.comp C \<open>C.dom e\<close>
apply unfold_locales using assms by auto
show "cone (C.dom e) (mkCone e)"
using assms mkCone_def apply unfold_locales
apply auto[2]
using C.dom_comp C.seqE C.cod_comp J.Zero_not_eq_One J.arr_char' J.cod_char map_def
apply (metis (no_types, lifting) C.not_arr_null parallel_pair.cod_simp(1) preserves_arr)
proof -
fix j
assume j: "J.arr j"
show "map j \<cdot> mkCone e (J.dom j) = mkCone e j"
proof -
have 1: "\<forall>a. if a = J.Zero then map a = C.dom f0
else if a = J.One then map a = C.cod f0
else if a = J.j0 then map a = f0
else if a = J.j1 then map a = f1
else map a = C.null"
using map_def by auto
hence 2: "map j = f1 \<or> j = J.One \<or> j = J.Zero \<or> j = J.j0"
using j parallel_pair.arr_char by meson
have "j = J.Zero \<or> map j \<cdot> mkCone e (J.dom j) = mkCone e j"
using assms j 1 2 mkCone_def C.cod_comp
by (metis (no_types, lifting) C.comp_cod_arr J.arr_char J.dom_simp(2-4) is_parallel)
thus ?thesis
using assms 1 j
by (metis (no_types, lifting) C.comp_cod_arr C.seqE mkCone_def J.dom_simp(1))
qed
next
show "\<And>j. J.arr j \<Longrightarrow> mkCone e (J.cod j) \<cdot> E.map j = mkCone e j"
proof -
fix j
assume j: "J.arr j"
have "J.cod j = J.Zero \<Longrightarrow> mkCone e (J.cod j) \<cdot> E.map j = mkCone e j"
unfolding mkCone_def
using assms j J.arr_char J.cod_char C.comp_arr_dom mkCone_def J.Zero_not_eq_One
by (metis (no_types, lifting) C.seqE E.map_simp)
moreover have "J.cod j \<noteq> J.Zero \<Longrightarrow> mkCone e (J.cod j) \<cdot> E.map j = mkCone e j"
unfolding mkCone_def
using assms j C.comp_arr_dom by auto
ultimately show "mkCone e (J.cod j) \<cdot> E.map j = mkCone e j" by blast
qed
qed
qed
lemma is_equalized_by_cone:
assumes "cone a \<chi>"
shows "is_equalized_by (\<chi> (J.Zero))"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
show ?thesis
using assms J.arr_char J.dom_char J.cod_char
J.One_not_eq_j0 J.One_not_eq_j1 J.Zero_not_eq_j0 J.Zero_not_eq_j1 J.j0_not_eq_j1
by (metis (no_types, lifting) Limit.cone_def \<chi>.is_natural_1 \<chi>.naturality
\<chi>.preserves_reflects_arr constant_functor.map_simp map_simp(3) map_simp(4))
qed
lemma mkCone_cone:
assumes "cone a \<chi>"
shows "mkCone (\<chi> J.Zero) = \<chi>"
proof -
interpret \<chi>: cone J.comp C map a \<chi>
using assms by auto
have 1: "is_equalized_by (\<chi> J.Zero)"
using assms is_equalized_by_cone by blast
show ?thesis
proof
fix j
have "j = J.Zero \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using mkCone_def \<chi>.is_extensional by simp
moreover have "j = J.One \<or> j = J.j0 \<or> j = J.j1 \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using J.arr_char J.cod_char J.dom_char J.seq_char mkCone_def
\<chi>.is_natural_1 \<chi>.is_natural_2 \<chi>.A.map_simp map_def
by (metis (no_types, lifting) J.Zero_not_eq_j0 J.dom_simp(2))
ultimately have "J.arr j \<Longrightarrow> mkCone (\<chi> J.Zero) j = \<chi> j"
using J.arr_char by auto
thus "mkCone (\<chi> J.Zero) j = \<chi> j"
using mkCone_def \<chi>.is_extensional by fastforce
qed
qed
end
locale equalizer_cone =
J: parallel_pair +
C: category C +
D: parallel_pair_diagram C f0 f1 +
limit_cone J.comp C D.map "C.dom e" "D.mkCone e"
for C :: "'c comp" (infixr "\<cdot>" 55)
and f0 :: 'c
and f1 :: 'c
and e :: 'c
begin
lemma equalizes:
shows "D.is_equalized_by e"
proof
show 1: "C.seq f0 e"
proof (intro C.seqI)
show "C.arr e" using ide_apex C.arr_dom_iff_arr by fastforce
show "C.arr f0"
using D.map_simp D.preserves_arr J.arr_char by metis
show "C.dom f0 = C.cod e"
using J.arr_char J.ide_char D.mkCone_def D.map_simp preserves_cod [of J.Zero]
by auto
qed
hence 2: "C.seq f1 e"
using D.is_parallel by fastforce
show "f0 \<cdot> e = f1 \<cdot> e"
using D.map_simp D.mkCone_def J.arr_char naturality [of J.j0] naturality [of J.j1]
by force
qed
lemma is_universal':
assumes "D.is_equalized_by e'"
shows "\<exists>!h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<and> e \<cdot> h = e'"
proof -
have "D.cone (C.dom e') (D.mkCone e')"
using assms D.cone_mkCone by blast
moreover have 0: "D.cone (C.dom e) (D.mkCone e)" ..
ultimately have 1: "\<exists>!h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<and>
D.cones_map h (D.mkCone e) = D.mkCone e'"
using is_universal [of "C.dom e'" "D.mkCone e'"] by auto
have 2: "\<And>h. \<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright> \<Longrightarrow>
D.cones_map h (D.mkCone e) = D.mkCone e' \<longleftrightarrow> e \<cdot> h = e'"
proof -
fix h
assume h: "\<guillemotleft>h : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
show "D.cones_map h (D.mkCone e) = D.mkCone e' \<longleftrightarrow> e \<cdot> h = e'"
proof
assume 3: "D.cones_map h (D.mkCone e) = D.mkCone e'"
show "e \<cdot> h = e'"
proof -
have "e' = D.mkCone e' J.Zero"
using D.mkCone_def J.arr_char by simp
also have "... = D.cones_map h (D.mkCone e) J.Zero"
using 3 by simp
also have "... = e \<cdot> h"
using 0 h D.mkCone_def J.arr_char by auto
finally show ?thesis by auto
qed
next
assume e': "e \<cdot> h = e'"
show "D.cones_map h (D.mkCone e) = D.mkCone e'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
using h cone_axioms D.mkCone_def by auto
moreover have "j = J.Zero \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
using h e' cone_\<chi> D.mkCone_def J.arr_char [of J.Zero] by force
moreover have
"J.arr j \<and> j \<noteq> J.Zero \<Longrightarrow> D.cones_map h (D.mkCone e) j = D.mkCone e' j"
proof -
assume j: "J.arr j \<and> j \<noteq> J.Zero"
have "D.cones_map h (D.mkCone e) j = C (D.mkCone e j) h"
using j h equalizes D.mkCone_def D.cone_mkCone J.arr_char
J.Zero_not_eq_One J.Zero_not_eq_j0 J.Zero_not_eq_j1
by auto
also have "... = (f0 \<cdot> e) \<cdot> h"
using j D.mkCone_def J.arr_char J.Zero_not_eq_One J.Zero_not_eq_j0
J.Zero_not_eq_j1
by auto
also have "... = f0 \<cdot> e \<cdot> h"
using h equalizes C.comp_assoc by blast
also have "... = D.mkCone e' j"
using j e' h equalizes D.mkCone_def J.arr_char [of J.One] J.Zero_not_eq_One
by auto
finally show ?thesis by auto
qed
ultimately show "D.cones_map h (D.mkCone e) j = D.mkCone e' j" by blast
qed
qed
qed
thus ?thesis using 1 by blast
qed
lemma induced_arrowI':
assumes "D.is_equalized_by e'"
shows "\<guillemotleft>induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
and "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
proof -
interpret A': constant_functor J.comp C \<open>C.dom e'\<close>
using assms by (unfold_locales, auto)
have cone: "D.cone (C.dom e') (D.mkCone e')"
using assms D.cone_mkCone [of e'] by blast
have "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') =
D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) J.Zero"
using cone induced_arrowI(1) D.mkCone_def J.arr_char cone_\<chi> by force
also have "... = e'"
proof -
have
"D.cones_map (induced_arrow (C.dom e') (D.mkCone e')) (D.mkCone e) =
D.mkCone e'"
using cone induced_arrowI by blast
thus ?thesis
using J.arr_char D.mkCone_def by simp
qed
finally have 1: "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
by auto
show "\<guillemotleft>induced_arrow (C.dom e') (D.mkCone e') : C.dom e' \<rightarrow> C.dom e\<guillemotright>"
using 1 cone induced_arrowI by simp
show "e \<cdot> induced_arrow (C.dom e') (D.mkCone e') = e'"
using 1 cone induced_arrowI by simp
qed
end
context category
begin
definition has_as_equalizer
where "has_as_equalizer f0 f1 e \<equiv> par f0 f1 \<and> parallel_pair_diagram.has_as_equalizer C f0 f1 e"
definition has_equalizers
where "has_equalizers = (\<forall>f0 f1. par f0 f1 \<longrightarrow> (\<exists>e. has_as_equalizer f0 f1 e))"
end
section "Limits by Products and Equalizers"
text\<open>
A category with equalizers has limits of shape @{term J} if it has products
indexed by the set of arrows of @{term J} and the set of objects of @{term J}.
The proof is patterned after \cite{MacLane}, Theorem 2, page 109:
\begin{quotation}
\noindent
``The limit of \<open>F: J \<rightarrow> C\<close> is the equalizer \<open>e\<close>
of \<open>f, g: \<Pi>\<^sub>i F\<^sub>i \<rightarrow> \<Pi>\<^sub>u F\<^sub>c\<^sub>o\<^sub>d \<^sub>u (u \<in> arr J, i \<in> J)\<close>
where \<open>p\<^sub>u f = p\<^sub>c\<^sub>o\<^sub>d \<^sub>u\<close>, \<open>p\<^sub>u g = F\<^sub>u o p\<^sub>d\<^sub>o\<^sub>m \<^sub>u\<close>;
the limiting cone \<open>\<mu>\<close> is \<open>\<mu>\<^sub>j = p\<^sub>j e\<close>, for \<open>j \<in> J\<close>.''
\end{quotation}
\<close>
locale category_with_equalizers =
category C
for C :: "'c comp" (infixr "\<cdot>" 55) +
assumes has_equalizers: "has_equalizers"
begin
lemma has_limits_if_has_products:
fixes J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
assumes "category J" and "has_products (Collect (partial_magma.ide J))"
and "has_products (Collect (partial_magma.arr J))"
shows "has_limits_of_shape J"
proof (unfold has_limits_of_shape_def)
interpret J: category J using assms(1) by auto
have "\<And>D. diagram J C D \<Longrightarrow> (\<exists>a \<chi>. limit_cone J C D a \<chi>)"
proof -
fix D
assume D: "diagram J C D"
interpret D: diagram J C D using D by auto
text\<open>
First, construct the two required products and their cones.
\<close>
interpret Obj: discrete_category \<open>Collect J.ide\<close> J.null
using J.not_arr_null J.ideD(1) mem_Collect_eq by (unfold_locales, blast)
interpret \<Delta>o: discrete_diagram_from_map \<open>Collect J.ide\<close> C D J.null
using D.preserves_ide by (unfold_locales, auto)
have "\<exists>p. has_as_product Obj.comp \<Delta>o.map p"
using assms(2) \<Delta>o.diagram_axioms has_products_def Obj.arr_char
by (metis (no_types, lifting) Collect_cong \<Delta>o.discrete_diagram_axioms mem_Collect_eq)
from this obtain \<Pi>o \<pi>o where \<pi>o: "product_cone Obj.comp C \<Delta>o.map \<Pi>o \<pi>o"
using ex_productE [of Obj.comp \<Delta>o.map] by auto
interpret \<pi>o: product_cone Obj.comp C \<Delta>o.map \<Pi>o \<pi>o using \<pi>o by auto
have \<pi>o_in_hom: "\<And>j. Obj.arr j \<Longrightarrow> \<guillemotleft>\<pi>o j : \<Pi>o \<rightarrow> D j\<guillemotright>"
using \<pi>o.preserves_dom \<pi>o.preserves_cod \<Delta>o.map_def by auto
interpret Arr: discrete_category \<open>Collect J.arr\<close> J.null
using J.not_arr_null by (unfold_locales, blast)
interpret \<Delta>a: discrete_diagram_from_map \<open>Collect J.arr\<close> C \<open>D o J.cod\<close> J.null
by (unfold_locales, auto)
have "\<exists>p. has_as_product Arr.comp \<Delta>a.map p"
using assms(3) has_products_def [of "Collect J.arr"] \<Delta>a.discrete_diagram_axioms
by blast
from this obtain \<Pi>a \<pi>a where \<pi>a: "product_cone Arr.comp C \<Delta>a.map \<Pi>a \<pi>a"
using ex_productE [of Arr.comp \<Delta>a.map] by auto
interpret \<pi>a: product_cone Arr.comp C \<Delta>a.map \<Pi>a \<pi>a using \<pi>a by auto
have \<pi>a_in_hom: "\<And>j. Arr.arr j \<Longrightarrow> \<guillemotleft>\<pi>a j : \<Pi>a \<rightarrow> D (J.cod j)\<guillemotright>"
using \<pi>a.preserves_cod \<pi>a.preserves_dom \<Delta>a.map_def by auto
text\<open>
Next, construct a parallel pair of arrows \<open>f, g: \<Pi>o \<rightarrow> \<Pi>a\<close>
that expresses the commutativity constraints imposed by the diagram.
\<close>
interpret \<Pi>o: constant_functor Arr.comp C \<Pi>o
using \<pi>o.ide_apex by (unfold_locales, auto)
let ?\<chi> = "\<lambda>j. if Arr.arr j then \<pi>o (J.cod j) else null"
interpret \<chi>: cone Arr.comp C \<Delta>a.map \<Pi>o ?\<chi>
using \<pi>o.ide_apex \<pi>o_in_hom \<Delta>a.map_def \<Delta>o.map_def \<Delta>o.is_discrete \<pi>o.is_natural_2
comp_cod_arr
by (unfold_locales, auto)
let ?f = "\<pi>a.induced_arrow \<Pi>o ?\<chi>"
have f_in_hom: "\<guillemotleft>?f : \<Pi>o \<rightarrow> \<Pi>a\<guillemotright>"
using \<chi>.cone_axioms \<pi>a.induced_arrowI by blast
have f_map: "\<Delta>a.cones_map ?f \<pi>a = ?\<chi>"
using \<chi>.cone_axioms \<pi>a.induced_arrowI by blast
have ff: "\<And>j. J.arr j \<Longrightarrow> \<pi>a j \<cdot> ?f = \<pi>o (J.cod j)"
proof -
fix j
assume j: "J.arr j"
have "\<pi>a j \<cdot> ?f = \<Delta>a.cones_map ?f \<pi>a j"
using f_in_hom \<pi>a.is_cone \<pi>a.is_extensional by auto
also have "... = \<pi>o (J.cod j)"
using j f_map by fastforce
finally show "\<pi>a j \<cdot> ?f = \<pi>o (J.cod j)" by auto
qed
let ?\<chi>' = "\<lambda>j. if Arr.arr j then D j \<cdot> \<pi>o (J.dom j) else null"
interpret \<chi>': cone Arr.comp C \<Delta>a.map \<Pi>o ?\<chi>'
using \<pi>o.ide_apex \<pi>o_in_hom \<Delta>o.map_def \<Delta>a.map_def comp_arr_dom comp_cod_arr
by (unfold_locales, auto)
let ?g = "\<pi>a.induced_arrow \<Pi>o ?\<chi>'"
have g_in_hom: "\<guillemotleft>?g : \<Pi>o \<rightarrow> \<Pi>a\<guillemotright>"
using \<chi>'.cone_axioms \<pi>a.induced_arrowI by blast
have g_map: "\<Delta>a.cones_map ?g \<pi>a = ?\<chi>'"
using \<chi>'.cone_axioms \<pi>a.induced_arrowI by blast
have gg: "\<And>j. J.arr j \<Longrightarrow> \<pi>a j \<cdot> ?g = D j \<cdot> \<pi>o (J.dom j)"
proof -
fix j
assume j: "J.arr j"
have "\<pi>a j \<cdot> ?g = \<Delta>a.cones_map ?g \<pi>a j"
using g_in_hom \<pi>a.is_cone \<pi>a.is_extensional by force
also have "... = D j \<cdot> \<pi>o (J.dom j)"
using j g_map by fastforce
finally show "\<pi>a j \<cdot> ?g = D j \<cdot> \<pi>o (J.dom j)" by auto
qed
interpret PP: parallel_pair_diagram C ?f ?g
using f_in_hom g_in_hom
by (elim in_homE, unfold_locales, auto)
from PP.is_parallel obtain e where equ: "PP.has_as_equalizer e"
using has_equalizers has_equalizers_def has_as_equalizer_def by blast
interpret EQU: limit_cone PP.J.comp C PP.map \<open>dom e\<close> \<open>PP.mkCone e\<close>
using equ by auto
interpret EQU: equalizer_cone C ?f ?g e ..
text\<open>
An arrow @{term h} with @{term "cod h = \<Pi>o"} equalizes @{term f} and @{term g}
if and only if it satisfies the commutativity condition required for a cone over
@{term D}.
\<close>
have E: "\<And>h. \<guillemotleft>h : dom h \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> h = ?g \<cdot> h \<longleftrightarrow> (\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h)"
proof
fix h
assume h: "\<guillemotleft>h : dom h \<rightarrow> \<Pi>o\<guillemotright>"
show "?f \<cdot> h = ?g \<cdot> h \<Longrightarrow> \<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
proof -
assume E: "?f \<cdot> h = ?g \<cdot> h"
have "\<And>j. J.arr j \<Longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
proof -
fix j
assume j: "J.arr j"
have "?\<chi> j \<cdot> h = \<Delta>a.cones_map ?f \<pi>a j \<cdot> h"
using j f_map by fastforce
also have "... = \<pi>a j \<cdot> ?f \<cdot> h"
using j f_in_hom \<Delta>a.map_def \<pi>a.cone_\<chi> comp_assoc by auto
also have "... = \<pi>a j \<cdot> ?g \<cdot> h"
using j E by simp
also have "... = \<Delta>a.cones_map ?g \<pi>a j \<cdot> h"
using j g_in_hom \<Delta>a.map_def \<pi>a.cone_\<chi> comp_assoc by auto
also have "... = ?\<chi>' j \<cdot> h"
using j g_map by force
finally show "?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h" by auto
qed
thus "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h" by blast
qed
show "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h \<Longrightarrow> ?f \<cdot> h = ?g \<cdot> h"
proof -
assume 1: "\<forall>j. J.arr j \<longrightarrow> ?\<chi> j \<cdot> h = ?\<chi>' j \<cdot> h"
have 2: "\<And>j. j \<in> Collect J.arr \<Longrightarrow> \<pi>a j \<cdot> ?f \<cdot> h = \<pi>a j \<cdot> ?g \<cdot> h"
proof -
fix j
assume j: "j \<in> Collect J.arr"
have "\<pi>a j \<cdot> ?f \<cdot> h = (\<pi>a j \<cdot> ?f) \<cdot> h"
using comp_assoc by simp
also have "... = ?\<chi> j \<cdot> h"
proof -
have "\<pi>a j \<cdot> ?f = \<Delta>a.cones_map ?f \<pi>a j"
using j f_in_hom \<pi>a.cone_axioms \<Delta>a.map_def \<pi>a.cone_\<chi> by auto
thus ?thesis using f_map by fastforce
qed
also have "... = ?\<chi>' j \<cdot> h"
using 1 j by auto
also have "... = (\<pi>a j \<cdot> ?g) \<cdot> h"
proof -
have "\<pi>a j \<cdot> ?g = \<Delta>a.cones_map ?g \<pi>a j"
using j g_in_hom \<pi>a.cone_axioms \<Delta>a.map_def \<pi>a.cone_\<chi> by auto
thus ?thesis using g_map by simp
qed
also have "... = \<pi>a j \<cdot> ?g \<cdot> h"
using comp_assoc by simp
finally show "\<pi>a j \<cdot> ?f \<cdot> h = \<pi>a j \<cdot> ?g \<cdot> h"
by auto
qed
show "C ?f h = C ?g h"
proof -
have "\<And>j. Arr.arr j \<Longrightarrow> \<guillemotleft>\<pi>a j \<cdot> ?f \<cdot> h : dom h \<rightarrow> \<Delta>a.map j\<guillemotright>"
using f_in_hom h \<pi>a_in_hom by (elim in_homE, auto)
hence 3: "\<exists>!k. \<guillemotleft>k : dom h \<rightarrow> \<Pi>a\<guillemotright> \<and> (\<forall>j. Arr.arr j \<longrightarrow> \<pi>a j \<cdot> k = \<pi>a j \<cdot> ?f \<cdot> h)"
using h \<pi>a \<pi>a.is_universal' [of "dom h" "\<lambda>j. \<pi>a j \<cdot> ?f \<cdot> h"] \<Delta>a.map_def
ide_dom [of h]
by blast
have 4: "\<And>P x x'. \<exists>!k. P k x \<Longrightarrow> P x x \<Longrightarrow> P x' x \<Longrightarrow> x' = x" by auto
let ?P = "\<lambda> k x. \<guillemotleft>k : dom h \<rightarrow> \<Pi>a\<guillemotright> \<and>
(\<forall>j. j \<in> Collect J.arr \<longrightarrow> \<pi>a j \<cdot> k = \<pi>a j \<cdot> x)"
have "?P (?g \<cdot> h) (?g \<cdot> h)"
using g_in_hom h by force
moreover have "?P (?f \<cdot> h) (?g \<cdot> h)"
using 2 f_in_hom g_in_hom h by force
ultimately show ?thesis
using 3 4 [of ?P "?f \<cdot> h" "?g \<cdot> h"] by auto
qed
qed
qed
have E': "\<And>e. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> e = ?g \<cdot> e \<longleftrightarrow>
(\<forall>j. J.arr j \<longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e)"
proof -
have 1: "\<And>e j. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow> J.arr j \<Longrightarrow>
?\<chi> j \<cdot> e = (D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e"
proof -
fix e j
assume e: "\<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright>"
assume j: "J.arr j"
have "\<guillemotleft>\<pi>o (J.cod j) \<cdot> e : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
using e j \<pi>o_in_hom by auto
thus "?\<chi> j \<cdot> e = (D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e"
using j comp_arr_dom comp_cod_arr by (elim in_homE, auto)
qed
have 2: "\<And>e j. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow> J.arr j \<Longrightarrow> ?\<chi>' j \<cdot> e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e"
proof -
fix e j
assume e: "\<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright>"
assume j: "J.arr j"
show "?\<chi>' j \<cdot> e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e"
using j comp_assoc by fastforce
qed
show "\<And>e. \<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright> \<Longrightarrow>
?f \<cdot> e = ?g \<cdot> e \<longleftrightarrow>
(\<forall>j. J.arr j \<longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> e) \<cdot> dom e = D j \<cdot> \<pi>o (J.dom j) \<cdot> e)"
using 1 2 E by presburger
qed
text\<open>
The composites of @{term e} with the projections from the product @{term \<Pi>o}
determine a limit cone @{term \<mu>} for @{term D}. The component of @{term \<mu>}
at an object @{term j} of @{term[source=true] J} is the composite @{term "C (\<pi>o j) e"}.
However, we need to extend @{term \<mu>} to all arrows @{term j} of @{term[source=true] J},
so the correct definition is @{term "\<mu> j = C (D j) (C (\<pi>o (J.dom j)) e)"}.
\<close>
have e_in_hom: "\<guillemotleft>e : dom e \<rightarrow> \<Pi>o\<guillemotright>"
using EQU.equalizes f_in_hom in_homI
by (metis (no_types, lifting) seqE in_homE)
have e_map: "C ?f e = C ?g e"
using EQU.equalizes f_in_hom in_homI by fastforce
interpret domE: constant_functor J C \<open>dom e\<close>
using e_in_hom by (unfold_locales, auto)
let ?\<mu> = "\<lambda>j. if J.arr j then D j \<cdot> \<pi>o (J.dom j) \<cdot> e else null"
have \<mu>: "\<And>j. J.arr j \<Longrightarrow> \<guillemotleft>?\<mu> j : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
proof -
fix j
assume j: "J.arr j"
show "\<guillemotleft>?\<mu> j : dom e \<rightarrow> D (J.cod j)\<guillemotright>"
using j e_in_hom \<pi>o_in_hom [of "J.dom j"] by auto
qed
interpret \<mu>: cone J C D \<open>dom e\<close> ?\<mu>
using \<mu> comp_cod_arr e_in_hom e_map E'
apply unfold_locales
apply auto
by (metis D.is_natural_1 comp_assoc)
text\<open>
If @{term \<tau>} is any cone over @{term D} then @{term \<tau>} restricts to a cone over
@{term \<Delta>o} for which the induced arrow to @{term \<Pi>o} equalizes @{term f} and @{term g}.
\<close>
have R: "\<And>a \<tau>. cone J C D a \<tau> \<Longrightarrow>
cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>) \<and>
?f \<cdot> \<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)
= ?g \<cdot> \<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
proof -
fix a \<tau>
assume cone_\<tau>: "cone J C D a \<tau>"
interpret \<tau>: cone J C D a \<tau> using cone_\<tau> by auto
interpret A: constant_functor Obj.comp C a
using \<tau>.ide_apex by (unfold_locales, auto)
interpret \<tau>o: cone Obj.comp C \<Delta>o.map a \<open>\<Delta>o.mkCone \<tau>\<close>
using A.value_is_ide \<Delta>o.map_def comp_cod_arr comp_arr_dom
by (unfold_locales, auto)
let ?e = "\<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
have mkCone_\<tau>: "\<Delta>o.mkCone \<tau> \<in> \<Delta>o.cones a"
using \<tau>o.cone_axioms by auto
have e: "\<guillemotleft>?e : a \<rightarrow> \<Pi>o\<guillemotright>"
using mkCone_\<tau> \<pi>o.induced_arrowI by simp
have ee: "\<And>j. J.ide j \<Longrightarrow> \<pi>o j \<cdot> ?e = \<tau> j"
proof -
fix j
assume j: "J.ide j"
have "\<pi>o j \<cdot> ?e = \<Delta>o.cones_map ?e \<pi>o j"
using j e \<pi>o.cone_axioms by force
also have "... = \<Delta>o.mkCone \<tau> j"
using j mkCone_\<tau> \<pi>o.induced_arrowI [of "\<Delta>o.mkCone \<tau>" a] by fastforce
also have "... = \<tau> j"
using j by simp
finally show "\<pi>o j \<cdot> ?e = \<tau> j" by auto
qed
have "\<And>j. J.arr j \<Longrightarrow>
(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
proof -
fix j
assume j: "J.arr j"
have 1: "\<guillemotleft>\<pi>o (J.cod j) : \<Pi>o \<rightarrow> D (J.cod j)\<guillemotright>" using j \<pi>o_in_hom by simp
have 2: "(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e
= D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e"
proof -
have "seq (D (J.cod j)) (\<pi>o (J.cod j))"
using j 1 by auto
moreover have "seq (\<pi>o (J.cod j)) ?e"
using j e by fastforce
ultimately show ?thesis using comp_arr_dom by auto
qed
also have 3: "... = \<pi>o (J.cod j) \<cdot> ?e"
using j e 1 comp_cod_arr by (elim in_homE, auto)
also have "... = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
using j e ee 2 3 \<tau>.naturality \<tau>.A.map_simp \<tau>.ide_apex comp_cod_arr by auto
finally show "(D (J.cod j) \<cdot> \<pi>o (J.cod j) \<cdot> ?e) \<cdot> dom ?e = D j \<cdot> \<pi>o (J.dom j) \<cdot> ?e"
by auto
qed
hence "C ?f ?e = C ?g ?e"
using E' \<pi>o.induced_arrowI \<tau>o.cone_axioms mem_Collect_eq by blast
thus "cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>) \<and> C ?f ?e = C ?g ?e"
using \<tau>o.cone_axioms by auto
qed
text\<open>
Finally, show that @{term \<mu>} is a limit cone.
\<close>
interpret \<mu>: limit_cone J C D \<open>dom e\<close> ?\<mu>
proof
fix a \<tau>
assume cone_\<tau>: "cone J C D a \<tau>"
interpret \<tau>: cone J C D a \<tau> using cone_\<tau> by auto
interpret A: constant_functor Obj.comp C a
using \<tau>.ide_apex by unfold_locales auto
have cone_\<tau>o: "cone Obj.comp C \<Delta>o.map a (\<Delta>o.mkCone \<tau>)"
using A.value_is_ide \<Delta>o.map_def D.preserves_ide comp_cod_arr comp_arr_dom
\<tau>.preserves_hom
by unfold_locales auto
show "\<exists>!h. \<guillemotleft>h : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h ?\<mu> = \<tau>"
proof
let ?e' = "\<pi>o.induced_arrow a (\<Delta>o.mkCone \<tau>)"
have e'_in_hom: "\<guillemotleft>?e' : a \<rightarrow> \<Pi>o\<guillemotright>"
using cone_\<tau> R \<pi>o.induced_arrowI by auto
have e'_map: "?f \<cdot> ?e' = ?g \<cdot> ?e' \<and> \<Delta>o.cones_map ?e' \<pi>o = \<Delta>o.mkCone \<tau>"
using cone_\<tau> R \<pi>o.induced_arrowI [of "\<Delta>o.mkCone \<tau>" a] by auto
have equ: "PP.is_equalized_by ?e'"
using e'_map e'_in_hom f_in_hom seqI' by blast
let ?h = "EQU.induced_arrow a (PP.mkCone ?e')"
have h_in_hom: "\<guillemotleft>?h : a \<rightarrow> dom e\<guillemotright>"
using EQU.induced_arrowI PP.cone_mkCone [of ?e'] e'_in_hom equ by fastforce
have h_map: "PP.cones_map ?h (PP.mkCone e) = PP.mkCone ?e'"
using EQU.induced_arrowI [of "PP.mkCone ?e'" a] PP.cone_mkCone [of ?e']
e'_in_hom equ
by fastforce
have 3: "D.cones_map ?h ?\<mu> = \<tau>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map ?h ?\<mu> j = \<tau> j"
using h_in_hom \<mu>.cone_axioms cone_\<tau> \<tau>.is_extensional by force
moreover have "J.arr j \<Longrightarrow> D.cones_map ?h ?\<mu> j = \<tau> j"
proof -
fix j
assume j: "J.arr j"
have 1: "\<guillemotleft>\<pi>o (J.dom j) \<cdot> e : dom e \<rightarrow> D (J.dom j)\<guillemotright>"
using j e_in_hom \<pi>o_in_hom [of "J.dom j"] by auto
have "D.cones_map ?h ?\<mu> j = ?\<mu> j \<cdot> ?h"
using h_in_hom j \<mu>.cone_axioms by auto
also have "... = D j \<cdot> (\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h"
using j comp_assoc by simp
also have "... = D j \<cdot> \<tau> (J.dom j)"
proof -
have "(\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h = \<tau> (J.dom j)"
proof -
have "(\<pi>o (J.dom j) \<cdot> e) \<cdot> ?h = \<pi>o (J.dom j) \<cdot> e \<cdot> ?h"
using j 1 e_in_hom h_in_hom \<pi>o arrI comp_assoc by auto
also have "... = \<pi>o (J.dom j) \<cdot> ?e'"
using equ e'_in_hom EQU.induced_arrowI' [of ?e'] by auto
also have "... = \<Delta>o.cones_map ?e' \<pi>o (J.dom j)"
using j e'_in_hom \<pi>o.cone_axioms by auto
also have "... = \<tau> (J.dom j)"
using j e'_map by simp
finally show ?thesis by auto
qed
thus ?thesis by simp
qed
also have "... = \<tau> j"
using j \<tau>.is_natural_1 by simp
finally show "D.cones_map ?h ?\<mu> j = \<tau> j" by auto
qed
ultimately show "D.cones_map ?h ?\<mu> j = \<tau> j" by auto
qed
show "\<guillemotleft>?h : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map ?h ?\<mu> = \<tau>"
using h_in_hom 3 by simp
show "\<And>h'. \<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h' ?\<mu> = \<tau> \<Longrightarrow> h' = ?h"
proof -
fix h'
assume h': "\<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright> \<and> D.cones_map h' ?\<mu> = \<tau>"
have h'_in_hom: "\<guillemotleft>h' : a \<rightarrow> dom e\<guillemotright>" using h' by simp
have h'_map: "D.cones_map h' ?\<mu> = \<tau>" using h' by simp
show "h' = ?h"
proof -
have 1: "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright> \<and> ?f \<cdot> e \<cdot> h' = ?g \<cdot> e \<cdot> h' \<and>
\<Delta>o.cones_map (C e h') \<pi>o = \<Delta>o.mkCone \<tau>"
proof -
have 2: "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright>" using h'_in_hom e_in_hom by auto
moreover have "?f \<cdot> e \<cdot> h' = ?g \<cdot> e \<cdot> h'"
proof -
have "?f \<cdot> e \<cdot> h' = (?f \<cdot> e) \<cdot> h'"
using comp_assoc by auto
also have "... = ?g \<cdot> e \<cdot> h'"
using EQU.equalizes comp_assoc by auto
finally show ?thesis by auto
qed
moreover have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o = \<Delta>o.mkCone \<tau>"
proof
have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o = \<Delta>o.cones_map h' (\<Delta>o.cones_map e \<pi>o)"
using \<pi>o.cone_axioms e_in_hom h'_in_hom \<Delta>o.cones_map_comp [of e h']
by fastforce
fix j
have "\<not>Obj.arr j \<Longrightarrow> \<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j"
using 2 e_in_hom h'_in_hom \<pi>o.cone_axioms by auto
moreover have "Obj.arr j \<Longrightarrow> \<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j"
proof -
assume j: "Obj.arr j"
have "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<pi>o j \<cdot> e \<cdot> h'"
using 2 j \<pi>o.cone_axioms by auto
also have "... = (\<pi>o j \<cdot> e) \<cdot> h'"
using comp_assoc by auto
also have "... = \<Delta>o.mkCone ?\<mu> j \<cdot> h'"
using j e_in_hom \<pi>o_in_hom comp_ide_arr [of "D j" "\<pi>o j \<cdot> e"]
by fastforce
also have "... = \<Delta>o.mkCone \<tau> j"
using j h' \<mu>.cone_axioms mem_Collect_eq by auto
finally show "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j" by auto
qed
ultimately show "\<Delta>o.cones_map (e \<cdot> h') \<pi>o j = \<Delta>o.mkCone \<tau> j" by auto
qed
ultimately show ?thesis by auto
qed
have "\<guillemotleft>e \<cdot> h' : a \<rightarrow> \<Pi>o\<guillemotright>" using 1 by simp
moreover have "e \<cdot> h' = ?e'"
using 1 cone_\<tau>o e'_in_hom e'_map \<pi>o.is_universal \<pi>o by blast
ultimately show "h' = ?h"
using 1 h'_in_hom h'_map EQU.is_universal' [of "e \<cdot> h'"]
EQU.induced_arrowI' [of ?e'] equ
by (elim in_homE) auto
qed
qed
qed
qed
have "limit_cone J C D (dom e) ?\<mu>" ..
thus "\<exists>a \<mu>. limit_cone J C D a \<mu>" by auto
qed
thus "\<forall>D. diagram J C D \<longrightarrow> (\<exists>a \<mu>. limit_cone J C D a \<mu>)" by blast
qed
end
section "Limits in a Set Category"
text\<open>
In this section, we consider the special case of limits in a set category.
\<close>
locale diagram_in_set_category =
J: category J +
S: set_category S \<AA> +
diagram J S D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and \<AA> :: "'a rel"
and D :: "'j \<Rightarrow> 's"
begin
notation S.in_hom ("\<guillemotleft>_ : _ \<rightarrow> _\<guillemotright>")
text\<open>
An object @{term a} of a set category @{term[source=true] S} is a limit of a diagram in
@{term[source=true] S} if and only if there is a bijection between the set
@{term "S.hom S.unity a"} of points of @{term a} and the set of cones over the diagram
that have apex @{term S.unity}.
\<close>
lemma limits_are_sets_of_cones:
shows "has_as_limit a \<longleftrightarrow> S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
proof
text\<open>
If \<open>has_limit a\<close>, then by the universal property of the limit cone,
composition in @{term[source=true] S} yields a bijection between @{term "S.hom S.unity a"}
and @{term "cones S.unity"}.
\<close>
assume a: "has_as_limit a"
hence "S.ide a"
using limit_cone_def cone.ide_apex by metis
from a obtain \<chi> where \<chi>: "limit_cone a \<chi>" by auto
interpret \<chi>: limit_cone J S D a \<chi> using \<chi> by auto
have "bij_betw (\<lambda>f. cones_map f \<chi>) (S.hom S.unity a) (cones S.unity)"
using \<chi>.bij_betw_hom_and_cones S.ide_unity by simp
thus "S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
using \<open>S.ide a\<close> by blast
next
text\<open>
Conversely, an arbitrary bijection @{term \<phi>} between @{term "S.hom S.unity a"}
and cones unity extends pointwise to a natural bijection @{term "\<Phi> a'"} between
@{term "S.hom a' a"} and @{term "cones a'"}, showing that @{term a} is a limit.
In more detail, the hypotheses give us a correspondence between points of @{term a}
and cones with apex @{term "S.unity"}. We extend this to a correspondence between
functions to @{term a} and general cones, with each arrow from @{term a'} to @{term a}
determining a cone with apex @{term a'}. If @{term "f \<in> hom a' a"} then composition
with @{term f} takes each point @{term y} of @{term a'} to the point @{term "S f y"}
of @{term a}. To this we may apply the given bijection @{term \<phi>} to obtain
@{term "\<phi> (S f y) \<in> cones S.unity"}. The component @{term "\<phi> (S f y) j"} at @{term j}
of this cone is a point of @{term "S.cod (D j)"}. Thus, @{term "f \<in> hom a' a"} determines
a cone @{term \<chi>f} with apex @{term a'} whose component at @{term j} is the
unique arrow @{term "\<chi>f j"} of @{term[source=true] S} such that
@{term "\<chi>f j \<in> hom a' (cod (D j))"} and @{term "S (\<chi>f j) y = \<phi> (S f y) j"}
for all points @{term y} of @{term a'}.
The cone @{term \<chi>a} corresponding to @{term "a \<in> S.hom a a"} is then a limit cone.
\<close>
assume a: "S.ide a \<and> (\<exists>\<phi>. bij_betw \<phi> (S.hom S.unity a) (cones S.unity))"
hence ide_a: "S.ide a" by auto
show "has_as_limit a"
proof -
from a obtain \<phi> where \<phi>: "bij_betw \<phi> (S.hom S.unity a) (cones S.unity)" by blast
have X: "\<And>f j y. \<lbrakk> \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>; J.arr j; \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<rbrakk>
\<Longrightarrow> \<guillemotleft>\<phi> (f \<cdot> y) j : S.unity \<rightarrow> S.cod (D j)\<guillemotright>"
proof -
fix f j y
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>" and j: "J.arr j" and y: "\<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright>"
interpret \<chi>: cone J S D S.unity \<open>\<phi> (S f y)\<close>
using f y \<phi> bij_betw_imp_funcset funcset_mem by blast
show "\<guillemotleft>\<phi> (f \<cdot> y) j : S.unity \<rightarrow> S.cod (D j)\<guillemotright>" using j by auto
qed
text\<open>
We want to define the component @{term "\<chi>j \<in> S.hom (S.dom f) (S.cod (D j))"}
at @{term j} of a cone by specifying how it acts by composition on points
@{term "y \<in> S.hom S.unity (S.dom f)"}. We can do this because @{term[source=true] S}
is a set category.
\<close>
let ?P = "\<lambda>f j \<chi>j. \<guillemotleft>\<chi>j : S.dom f \<rightarrow> S.cod (D j)\<guillemotright> \<and>
(\<forall>y. \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<longrightarrow> \<chi>j \<cdot> y = \<phi> (f \<cdot> y) j)"
let ?\<chi> = "\<lambda>f j. if J.arr j then (THE \<chi>j. ?P f j \<chi>j) else S.null"
have \<chi>: "\<And>f j. \<lbrakk> \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>; J.arr j \<rbrakk> \<Longrightarrow> ?P f j (?\<chi> f j)"
proof -
fix b f j
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>" and j: "J.arr j"
interpret B: constant_functor J S \<open>S.dom f\<close>
using f by (unfold_locales) auto
have "(\<lambda>y. \<phi> (f \<cdot> y) j) \<in> S.hom S.unity (S.dom f) \<rightarrow> S.hom S.unity (S.cod (D j))"
using f j X Pi_I' by simp
hence "\<exists>!\<chi>j. ?P f j \<chi>j"
using f j S.fun_complete' by (elim S.in_homE) auto
thus "?P f j (?\<chi> f j)" using j theI' [of "?P f j"] by simp
qed
text\<open>
The arrows @{term "\<chi> f j"} are in fact the components of a cone with apex
@{term "S.dom f"}.
\<close>
have cone: "\<And>f. \<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright> \<Longrightarrow> cone (S.dom f) (?\<chi> f)"
proof -
fix f
assume f: "\<guillemotleft>f : S.dom f \<rightarrow> a\<guillemotright>"
interpret B: constant_functor J S \<open>S.dom f\<close>
using f by unfold_locales auto
show "cone (S.dom f) (?\<chi> f)"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> ?\<chi> f j = S.null" by simp
fix j
assume j: "J.arr j"
have 0: "\<guillemotleft>?\<chi> f j : S.dom f \<rightarrow> S.cod (D j)\<guillemotright>" using f j \<chi> by simp
show "S.dom (?\<chi> f j) = B.map (J.dom j)" using f j \<chi> by auto
show "S.cod (?\<chi> f j) = D (J.cod j)" using f j \<chi> by auto
have par2: "S.par (?\<chi> f (J.cod j) \<cdot> B.map j) (?\<chi> f j)"
using f j 0 \<chi> [of f "J.cod j"] by (elim S.in_homE, auto)
have nat: "\<And>y. \<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright> \<Longrightarrow>
(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y \<and>
(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y"
proof -
fix y
assume y: "\<guillemotleft>y : S.unity \<rightarrow> S.dom f\<guillemotright>"
show "(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y \<and>
(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y"
proof
have 1: "\<phi> (f \<cdot> y) \<in> cones S.unity"
using f y \<phi> bij_betw_imp_funcset PiE
S.seqI S.cod_comp S.dom_comp mem_Collect_eq
by fastforce
interpret \<chi>: cone J S D S.unity \<open>\<phi> (f \<cdot> y)\<close>
using 1 by simp
have "(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = D j \<cdot> ?\<chi> f (J.dom j) \<cdot> y"
using S.comp_assoc by simp
also have "... = D j \<cdot> \<phi> (f \<cdot> y) (J.dom j)"
using f y \<chi> \<chi>.is_extensional by simp
also have "... = \<phi> (f \<cdot> y) j" using j by auto
also have "... = ?\<chi> f j \<cdot> y"
using f j y \<chi> by force
finally show "(D j \<cdot> ?\<chi> f (J.dom j)) \<cdot> y = ?\<chi> f j \<cdot> y" by auto
have "(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f (J.cod j) \<cdot> y"
using j B.map_simp par2 B.value_is_ide S.comp_arr_ide
by (metis (no_types, lifting))
also have "... = \<phi> (f \<cdot> y) (J.cod j)"
using f y \<chi> \<chi>.is_extensional by simp
also have "... = \<phi> (f \<cdot> y) j"
using j \<chi>.is_natural_2
by (metis J.arr_cod \<chi>.A.map_simp J.cod_cod)
also have "... = ?\<chi> f j \<cdot> y"
using f y \<chi> \<chi>.is_extensional by simp
finally show "(?\<chi> f (J.cod j) \<cdot> B.map j) \<cdot> y = ?\<chi> f j \<cdot> y" by auto
qed
qed
show "D j \<cdot> ?\<chi> f (J.dom j) = ?\<chi> f j"
proof -
have "S.par (D j \<cdot> ?\<chi> f (J.dom j)) (?\<chi> f j)"
using f j 0 \<chi> [of f "J.dom j"] by (elim S.in_homE, auto)
thus ?thesis
using nat 0
apply (intro S.arr_eqI' [of "D j \<cdot> ?\<chi> f (J.dom j)" "?\<chi> f j"])
apply force
by auto
qed
show "?\<chi> f (J.cod j) \<cdot> B.map j = ?\<chi> f j"
using par2 nat 0 f j \<chi>
apply (intro S.arr_eqI' [of "?\<chi> f (J.cod j) \<cdot> B.map j" "?\<chi> f j"])
apply force
by (metis (no_types, lifting) S.in_homE)
qed
qed
interpret \<chi>a: cone J S D a \<open>?\<chi> a\<close> using a cone [of a] by fastforce
text\<open>
Finally, show that \<open>\<chi> a\<close> is a limit cone.
\<close>
interpret \<chi>a: limit_cone J S D a \<open>?\<chi> a\<close>
proof
fix a' \<chi>'
assume cone_\<chi>': "cone a' \<chi>'"
interpret \<chi>': cone J S D a' \<chi>' using cone_\<chi>' by auto
show "\<exists>!f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and> cones_map f (?\<chi> a) = \<chi>'"
proof
let ?\<psi> = "inv_into (S.hom S.unity a) \<phi>"
have \<psi>: "?\<psi> \<in> cones S.unity \<rightarrow> S.hom S.unity a"
using \<phi> bij_betw_inv_into bij_betwE by blast
let ?P = "\<lambda>f. \<guillemotleft>f : a' \<rightarrow> a\<guillemotright> \<and>
(\<forall>y. y \<in> S.hom S.unity a' \<longrightarrow> f \<cdot> y = ?\<psi> (cones_map y \<chi>'))"
have 1: "\<exists>!f. ?P f"
proof -
have "(\<lambda>y. ?\<psi> (cones_map y \<chi>')) \<in> S.hom S.unity a' \<rightarrow> S.hom S.unity a"
proof
fix x
assume "x \<in> S.hom S.unity a'"
hence "\<guillemotleft>x : S.unity \<rightarrow> a'\<guillemotright>" by simp
hence "cones_map x \<in> cones a' \<rightarrow> cones S.unity"
using cones_map_mapsto [of x] by (elim S.in_homE) auto
hence "cones_map x \<chi>' \<in> cones S.unity"
using cone_\<chi>' by blast
thus "?\<psi> (cones_map x \<chi>') \<in> S.hom S.unity a"
using \<psi> by auto
qed
thus ?thesis
using S.fun_complete' a \<chi>'.ide_apex by simp
qed
let ?f = "THE f. ?P f"
have f: "?P ?f" using 1 theI' [of ?P] by simp
have f_in_hom: "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright>" using f by simp
have f_map: "cones_map ?f (?\<chi> a) = \<chi>'"
proof -
have 1: "cone a' (cones_map ?f (?\<chi> a))"
proof -
have "cones_map ?f \<in> cones a \<rightarrow> cones a'"
using f_in_hom cones_map_mapsto [of ?f] by (elim S.in_homE) auto
hence "cones_map ?f (?\<chi> a) \<in> cones a'"
using \<chi>a.cone_axioms by blast
thus ?thesis by simp
qed
interpret f\<chi>a: cone J S D a' \<open>cones_map ?f (?\<chi> a)\<close>
using 1 by simp
show ?thesis
proof
fix j
have "\<not>J.arr j \<Longrightarrow> cones_map ?f (?\<chi> a) j = \<chi>' j"
using 1 \<chi>'.is_extensional f\<chi>a.is_extensional by presburger
moreover have "J.arr j \<Longrightarrow> cones_map ?f (?\<chi> a) j = \<chi>' j"
proof -
assume j: "J.arr j"
show "cones_map ?f (?\<chi> a) j = \<chi>' j"
proof (intro S.arr_eqI' [of "cones_map ?f (?\<chi> a) j" "\<chi>' j"])
show par: "S.par (cones_map ?f (?\<chi> a) j) (\<chi>' j)"
using j \<chi>'.preserves_cod \<chi>'.preserves_dom \<chi>'.preserves_reflects_arr
f\<chi>a.preserves_cod f\<chi>a.preserves_dom f\<chi>a.preserves_reflects_arr
by presburger
fix y
assume "\<guillemotleft>y : S.unity \<rightarrow> S.dom (cones_map ?f (?\<chi> a) j)\<guillemotright>"
hence y: "\<guillemotleft>y : S.unity \<rightarrow> a'\<guillemotright>"
using j f\<chi>a.preserves_dom by simp
have 1: "\<guillemotleft>?\<chi> a j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using j \<chi>a.preserves_hom by force
have 2: "\<guillemotleft>?f \<cdot> y : S.unity \<rightarrow> a\<guillemotright>"
using f_in_hom y by blast
have "cones_map ?f (?\<chi> a) j \<cdot> y = (?\<chi> a j \<cdot> ?f) \<cdot> y"
proof -
have "S.cod ?f = a" using f_in_hom by blast
thus ?thesis using j \<chi>a.cone_axioms by simp
qed
also have "... = ?\<chi> a j \<cdot> ?f \<cdot> y"
using 1 j y f_in_hom S.comp_assoc S.seqI' by blast
also have "... = \<phi> (a \<cdot> ?f \<cdot> y) j"
using 1 2 ide_a f j y \<chi> [of a] by (simp add: S.ide_in_hom)
also have "... = \<phi> (?f \<cdot> y) j"
using a 2 y S.comp_cod_arr by (elim S.in_homE, auto)
also have "... = \<phi> (?\<psi> (cones_map y \<chi>')) j"
using j y f by simp
also have "... = cones_map y \<chi>' j"
proof -
have "cones_map y \<chi>' \<in> cones S.unity"
using cone_\<chi>' y cones_map_mapsto by force
hence "\<phi> (?\<psi> (cones_map y \<chi>')) = cones_map y \<chi>'"
using \<phi> bij_betw_inv_into_right [of \<phi>] by simp
thus ?thesis by auto
qed
also have "... = \<chi>' j \<cdot> y"
using cone_\<chi>' j y by auto
finally show "cones_map ?f (?\<chi> a) j \<cdot> y = \<chi>' j \<cdot> y"
by auto
qed
qed
ultimately show "cones_map ?f (?\<chi> a) j = \<chi>' j" by blast
qed
qed
show "\<guillemotleft>?f : a' \<rightarrow> a\<guillemotright> \<and> cones_map ?f (?\<chi> a) = \<chi>'"
using f_in_hom f_map by simp
show "\<And>f'. \<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> cones_map f' (?\<chi> a) = \<chi>' \<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright> \<and> cones_map f' (?\<chi> a) = \<chi>'"
have f'_in_hom: "\<guillemotleft>f' : a' \<rightarrow> a\<guillemotright>" using f' by simp
have f'_map: "cones_map f' (?\<chi> a) = \<chi>'" using f' by simp
show "f' = ?f"
proof (intro S.arr_eqI' [of f' ?f])
show "S.par f' ?f"
using f_in_hom f'_in_hom by (elim S.in_homE, auto)
show "\<And>y'. \<guillemotleft>y' : S.unity \<rightarrow> S.dom f'\<guillemotright> \<Longrightarrow> f' \<cdot> y' = ?f \<cdot> y'"
proof -
fix y'
assume y': "\<guillemotleft>y' : S.unity \<rightarrow> S.dom f'\<guillemotright>"
have 0: "\<phi> (f' \<cdot> y') = cones_map y' \<chi>'"
proof
fix j
have 1: "\<guillemotleft>f' \<cdot> y' : S.unity \<rightarrow> a\<guillemotright>" using f'_in_hom y' by auto
hence 2: "\<phi> (f' \<cdot> y') \<in> cones S.unity"
using \<phi> bij_betw_imp_funcset [of \<phi> "S.hom S.unity a" "cones S.unity"]
by auto
interpret \<chi>'': cone J S D S.unity \<open>\<phi> (f' \<cdot> y')\<close> using 2 by auto
have "\<not>J.arr j \<Longrightarrow> \<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j"
using f' y' cone_\<chi>' \<chi>''.is_extensional mem_Collect_eq restrict_apply
by (elim S.in_homE, auto)
moreover have "J.arr j \<Longrightarrow> \<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j"
proof -
assume j: "J.arr j"
have 3: "\<guillemotleft>?\<chi> a j : a \<rightarrow> D (J.cod j)\<guillemotright>"
using j \<chi>a.preserves_hom by force
have "\<phi> (f' \<cdot> y') j = \<phi> (a \<cdot> f' \<cdot> y') j"
using a f' y' j S.comp_cod_arr by (elim S.in_homE, auto)
also have "... = ?\<chi> a j \<cdot> f' \<cdot> y'"
using 1 3 \<chi> [of a] a f' y' j by fastforce
also have "... = (?\<chi> a j \<cdot> f') \<cdot> y'"
using S.comp_assoc by simp
also have "... = cones_map f' (?\<chi> a) j \<cdot> y'"
using f' y' j \<chi>a.cone_axioms by auto
also have "... = \<chi>' j \<cdot> y'"
using f' by blast
also have "... = cones_map y' \<chi>' j"
using y' j cone_\<chi>' f' mem_Collect_eq restrict_apply by force
finally show "\<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j" by auto
qed
ultimately show "\<phi> (f' \<cdot> y') j = cones_map y' \<chi>' j" by auto
qed
hence "f' \<cdot> y' = ?\<psi> (cones_map y' \<chi>')"
using \<phi> f'_in_hom y' S.comp_in_homI
bij_betw_inv_into_left [of \<phi> "S.hom S.unity a" "cones S.unity" "f' \<cdot> y'"]
by (elim S.in_homE, auto)
moreover have "?f \<cdot> y' = ?\<psi> (cones_map y' \<chi>')"
using \<phi> 0 1 f f_in_hom f'_in_hom y' S.comp_in_homI
bij_betw_inv_into_left [of \<phi> "S.hom S.unity a" "cones S.unity" "?f \<cdot> y'"]
by (elim S.in_homE, auto)
ultimately show "f' \<cdot> y' = ?f \<cdot> y'" by auto
qed
qed
qed
qed
qed
have "limit_cone a (?\<chi> a)" ..
thus ?thesis by auto
qed
qed
end
locale diagram_in_replete_set_category =
J: category J +
S: replete_set_category S +
diagram J S D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 's"
begin
sublocale diagram_in_set_category J S \<open>cardSuc (cmax (card_of (UNIV :: 's set)) natLeq)\<close> D
..
end
context set_category
begin
text\<open>
A set category has an equalizer for any parallel pair of arrows.
\<close>
lemma has_equalizers:
shows "has_equalizers"
proof (unfold has_equalizers_def)
have "\<And>f0 f1. par f0 f1 \<Longrightarrow> \<exists>e. has_as_equalizer f0 f1 e"
proof -
fix f0 f1
assume par: "par f0 f1"
interpret J: parallel_pair .
interpret PP: parallel_pair_diagram S f0 f1
apply unfold_locales using par by auto
interpret PP: diagram_in_set_category J.comp S \<AA> PP.map ..
text\<open>
Let @{term a} be the object corresponding to the set of all images of equalizing points
of @{term "dom f0"}, and let @{term e} be the inclusion of @{term a} in @{term "dom f0"}.
\<close>
let ?a = "mkIde (img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e})"
have 0: "{e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e} \<subseteq> hom unity (dom f0)"
by auto
hence 1: "img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e} \<subseteq> Univ"
using img_point_in_Univ by auto
have 2: "|img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}| <o \<AA>"
proof -
have "|hom unity (dom f0)| =o |Dom f0|"
using par bij_betw_points_and_set [of "dom f0"]
by (simp add: card_of_ordIsoI)
moreover have "|Dom f0| <o \<AA>"
using par set_card by simp
ultimately have "|hom unity (dom f0)| <o \<AA>"
using ordIso_ordLess_trans by blast
moreover have
"|{e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}| \<le>o |hom unity (dom f0)|"
using 0 by simp
ultimately have "|{e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}| <o \<AA>"
using ordLeq_ordLess_trans by blast
thus ?thesis
using card_of_image ordLeq_ordLess_trans by blast
qed
have ide_a: "ide ?a" using 1 2 ide_mkIde by auto
have set_a: "set ?a = img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using 1 2 set_mkIde by simp
have incl_in_a: "incl_in ?a (dom f0)"
proof -
have "ide (dom f0)"
using PP.is_parallel by simp
moreover have "set ?a \<subseteq> set (dom f0)"
proof -
have "set ?a = img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using img_point_in_Univ set_a by blast
thus ?thesis
using imageE img_point_elem_set mem_Collect_eq subsetI by auto
qed
ultimately show ?thesis
using incl_in_def \<open>ide ?a\<close> by simp
qed
text\<open>
Then @{term "set a"} is in bijective correspondence with @{term "PP.cones unity"}.
\<close>
let ?\<phi> = "\<lambda>t. PP.mkCone (mkPoint (dom f0) t)"
let ?\<psi> = "\<lambda>\<chi>. img (\<chi> (J.Zero))"
have bij: "bij_betw ?\<phi> (set ?a) (PP.cones unity)"
proof (intro bij_betwI)
show "?\<phi> \<in> set ?a \<rightarrow> PP.cones unity"
proof
fix t
assume t: "t \<in> set ?a"
hence 1: "t \<in> img ` {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using set_a by blast
then have 2: "mkPoint (dom f0) t \<in> hom unity (dom f0)"
using mkPoint_in_hom imageE mem_Collect_eq mkPoint_img(2) by auto
with 1 have 3: "mkPoint (dom f0) t \<in> {e. e \<in> hom unity (dom f0) \<and> f0 \<cdot> e = f1 \<cdot> e}"
using mkPoint_img(2) by auto
then have "PP.is_equalized_by (mkPoint (dom f0) t)"
using CollectD par by fastforce
thus "PP.mkCone (mkPoint (dom f0) t) \<in> PP.cones unity"
using 2 PP.cone_mkCone [of "mkPoint (dom f0) t"] by auto
qed
show "?\<psi> \<in> PP.cones unity \<rightarrow> set ?a"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> PP.cones unity"
interpret \<chi>: cone J.comp S PP.map unity \<chi> using \<chi> by auto
have "\<chi> (J.Zero) \<in> hom unity (dom f0) \<and> f0 \<cdot> \<chi> (J.Zero) = f1 \<cdot> \<chi> (J.Zero)"
using \<chi> PP.map_def PP.is_equalized_by_cone J.arr_char by auto
hence "img (\<chi> (J.Zero)) \<in> set ?a"
using set_a by simp
thus "?\<psi> \<chi> \<in> set ?a" by blast
qed
show "\<And>t. t \<in> set ?a \<Longrightarrow> ?\<psi> (?\<phi> t) = t"
using set_a J.arr_char PP.mkCone_def imageE mem_Collect_eq mkPoint_img(2)
by auto
show "\<And>\<chi>. \<chi> \<in> PP.cones unity \<Longrightarrow> ?\<phi> (?\<psi> \<chi>) = \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> PP.cones unity"
interpret \<chi>: cone J.comp S PP.map unity \<chi> using \<chi> by auto
have 1: "\<chi> (J.Zero) \<in> hom unity (dom f0) \<and> f0 \<cdot> \<chi> (J.Zero) = f1 \<cdot> \<chi> (J.Zero)"
using \<chi> PP.map_def PP.is_equalized_by_cone J.arr_char by auto
hence "img (\<chi> (J.Zero)) \<in> set ?a"
using set_a by simp
hence "img (\<chi> (J.Zero)) \<in> set (dom f0)"
using incl_in_a incl_in_def by auto
hence "mkPoint (dom f0) (img (\<chi> J.Zero)) = \<chi> J.Zero"
using 1 mkPoint_img(2) by blast
hence "?\<phi> (?\<psi> \<chi>) = PP.mkCone (\<chi> J.Zero)" by simp
also have "... = \<chi>"
using \<chi> PP.mkCone_cone by simp
finally show "?\<phi> (?\<psi> \<chi>) = \<chi>" by auto
qed
qed
text\<open>
It follows that @{term a} is a limit of \<open>PP\<close>, and that the limit cone gives an
equalizer of @{term f0} and @{term f1}.
\<close>
have "\<exists>\<mu>. bij_betw \<mu> (hom unity ?a) (set ?a)"
using bij_betw_points_and_set ide_a by auto
from this obtain \<mu> where \<mu>: "bij_betw \<mu> (hom unity ?a) (set ?a)" by blast
have "bij_betw (?\<phi> o \<mu>) (hom unity ?a) (PP.cones unity)"
using bij \<mu> bij_betw_comp_iff by blast
hence "\<exists>\<phi>. bij_betw \<phi> (hom unity ?a) (PP.cones unity)" by auto
hence "PP.has_as_limit ?a"
using ide_a PP.limits_are_sets_of_cones by simp
from this obtain \<epsilon> where \<epsilon>: "limit_cone J.comp S PP.map ?a \<epsilon>" by auto
interpret \<epsilon>: limit_cone J.comp S PP.map ?a \<epsilon> using \<epsilon> by auto
have "PP.mkCone (\<epsilon> (J.Zero)) = \<epsilon>"
using \<epsilon> PP.mkCone_cone \<epsilon>.cone_axioms by simp
moreover have "dom (\<epsilon> (J.Zero)) = ?a"
using J.ide_char \<epsilon>.preserves_hom \<epsilon>.A.map_def by simp
ultimately have "PP.has_as_equalizer (\<epsilon> J.Zero)"
using \<epsilon> by simp
thus "\<exists>e. has_as_equalizer f0 f1 e"
using par has_as_equalizer_def by auto
qed
thus "\<forall>f0 f1. par f0 f1 \<longrightarrow> (\<exists>e. has_as_equalizer f0 f1 e)" by auto
qed
end
sublocale set_category \<subseteq> category_with_equalizers S
apply unfold_locales using has_equalizers by auto
context set_category
begin
text\<open>
The aim of the next results is to characterize the conditions under which a set
category has products. In a traditional development of category theory,
one shows that the category \textbf{Set} of \emph{all} sets has all small
(\emph{i.e.}~set-indexed) products. In the present context we do not have a
category of \emph{all} sets, but rather only a category of all sets with
elements at a particular type. Clearly, we cannot expect such a category
to have products indexed by arbitrarily large sets. The existence of
@{term I}-indexed products in a set category @{term[source=true] S} implies that the universe
\<open>S.Univ\<close> of @{term[source=true] S} must be large enough to admit the formation of
@{term I}-tuples of its elements. Conversely, for a set category @{term[source=true] S}
the ability to form @{term I}-tuples in @{term Univ} implies that
@{term[source=true] S} has @{term I}-indexed products. Below we make this precise by
defining the notion of when a set category @{term[source=true] S}
``admits @{term I}-indexed tupling'' and we show that @{term[source=true] S}
has @{term I}-indexed products if and only if it admits @{term I}-indexed tupling.
The definition of ``@{term[source=true] S} admits @{term I}-indexed tupling'' says that
there is an injective map, from the space of extensional functions from
@{term I} to @{term Univ}, to @{term Univ}. However for a convenient
statement and proof of the desired result, the definition of extensional
function from theory @{theory "HOL-Library.FuncSet"} needs to be modified.
The theory @{theory "HOL-Library.FuncSet"} uses the definite, but arbitrarily chosen value
@{term undefined} as the value to be assumed by an extensional function outside
of its domain. In the context of the \<open>set_category\<close>, though, it is
more natural to use \<open>S.unity\<close>, which is guaranteed to be an element of the
universe of @{term[source=true] S}, for this purpose. Doing things that way makes it
simpler to establish a bijective correspondence between cones over @{term D} with apex
@{term unity} and the set of extensional functions @{term d} that map
each arrow @{term j} of @{term J} to an element @{term "d j"} of @{term "set (D j)"}.
Possibly it makes sense to go back and make this change in \<open>set_category\<close>,
but that would mean completely abandoning @{theory "HOL-Library.FuncSet"} and essentially
introducing a duplicate version for use with \<open>set_category\<close>.
As a compromise, what I have done here is to locally redefine the few notions from
@{theory "HOL-Library.FuncSet"} that I need in order to prove the next set of results.
The redefined notions are primed to avoid confusion with the original versions.
\<close>
definition extensional'
where "extensional' A \<equiv> {f. \<forall>x. x \<notin> A \<longrightarrow> f x = unity}"
abbreviation PiE'
where "PiE' A B \<equiv> Pi A B \<inter> extensional' A"
abbreviation restrict'
where "restrict' f A \<equiv> \<lambda>x. if x \<in> A then f x else unity"
lemma extensional'I [intro]:
assumes "\<And>x. x \<notin> A \<Longrightarrow> f x = unity"
shows "f \<in> extensional' A"
using assms extensional'_def by auto
lemma extensional'_arb:
assumes "f \<in> extensional' A" and "x \<notin> A"
shows "f x = unity"
using assms extensional'_def by fast
lemma extensional'_monotone:
assumes "A \<subseteq> B"
shows "extensional' A \<subseteq> extensional' B"
proof
fix f
assume f: "f \<in> extensional' A"
have 1: "\<forall>x. x \<notin> A \<longrightarrow> f x = unity" using f extensional'_def by fast
hence "\<forall>x. x \<notin> B \<longrightarrow> f x = unity" using assms by auto
thus "f \<in> extensional' B" using extensional'_def by blast
qed
lemma PiE'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> PiE' A B \<subseteq> PiE' A C"
by auto
end
locale discrete_diagram_in_replete_set_category =
S: replete_set_category S +
discrete_diagram J S D +
diagram_in_replete_set_category J S D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and S :: "'s comp" (infixr "\<cdot>" 55)
and D :: "'j \<Rightarrow> 's"
begin
text\<open>
For @{term D} a discrete diagram in a set category, there is a bijective correspondence
between cones over @{term D} with apex unity and the set of extensional functions @{term d}
that map each arrow @{term j} of @{term[source=true] J} to an element of
@{term "S.set (D j)"}.
\<close>
abbreviation I
where "I \<equiv> Collect J.arr"
definition funToCone
where "funToCone F \<equiv> \<lambda>j. if J.arr j then S.mkPoint (D j) (F j) else S.null"
definition coneToFun
where "coneToFun \<chi> \<equiv> \<lambda>j. if J.arr j then S.img (\<chi> j) else S.unity"
lemma funToCone_mapsto:
shows "funToCone \<in> S.PiE' I (S.set o D) \<rightarrow> cones S.unity"
proof
fix F
assume F: "F \<in> S.PiE' I (S.set o D)"
interpret U: constant_functor J S S.unity
apply unfold_locales using S.ide_unity by auto
have 1: "S.ide (S.mkIde S.Univ)"
using S.ide_mkIde by simp
have "cone S.unity (funToCone F)"
proof
show "\<And>j. \<not>J.arr j \<Longrightarrow> funToCone F j = S.null"
using funToCone_def by simp
fix j
assume j: "J.arr j"
have "funToCone F j = S.mkPoint (D j) (F j)"
using j funToCone_def by simp
moreover have "... \<in> S.hom S.unity (D j)"
using F j is_discrete S.img_mkPoint(1) [of "D j"] by force
ultimately have 2: "funToCone F j \<in> S.hom S.unity (D j)" by auto
show 3: "S.dom (funToCone F j) = U.map (J.dom j)"
using 2 j U.map_simp by auto
show 4: "S.cod (funToCone F j) = D (J.cod j)"
using 2 j is_discrete by auto
show "D j \<cdot> funToCone F (J.dom j) = funToCone F j"
using 2 j is_discrete S.comp_cod_arr by auto
show "funToCone F (J.cod j) \<cdot> (U.map j) = funToCone F j"
using 3 j is_discrete U.map_simp S.arr_dom_iff_arr S.comp_arr_dom U.preserves_arr
by (metis J.ide_char)
qed
thus "funToCone F \<in> cones S.unity" by auto
qed
lemma coneToFun_mapsto:
shows "coneToFun \<in> cones S.unity \<rightarrow> S.PiE' I (S.set o D)"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> cones S.unity"
interpret \<chi>: cone J S D S.unity \<chi> using \<chi> by auto
show "coneToFun \<chi> \<in> S.PiE' I (S.set o D)"
proof
show "coneToFun \<chi> \<in> Pi I (S.set o D)"
using S.mkPoint_img(1) coneToFun_def is_discrete \<chi>.component_in_hom
by (simp add: S.img_point_elem_set restrict_apply')
show "coneToFun \<chi> \<in> S.extensional' I"
proof
fix x
show "x \<notin> I \<Longrightarrow> coneToFun \<chi> x = S.unity"
using coneToFun_def by simp
qed
qed
qed
lemma funToCone_coneToFun:
assumes "\<chi> \<in> cones S.unity"
shows "funToCone (coneToFun \<chi>) = \<chi>"
proof
interpret \<chi>: cone J S D S.unity \<chi> using assms by auto
fix j
have "\<not>J.arr j \<Longrightarrow> funToCone (coneToFun \<chi>) j = \<chi> j"
using funToCone_def \<chi>.is_extensional by simp
moreover have "J.arr j \<Longrightarrow> funToCone (coneToFun \<chi>) j = \<chi> j"
using funToCone_def coneToFun_def S.mkPoint_img(2) is_discrete \<chi>.component_in_hom
by auto
ultimately show "funToCone (coneToFun \<chi>) j = \<chi> j" by blast
qed
lemma coneToFun_funToCone:
assumes "F \<in> S.PiE' I (S.set o D)"
shows "coneToFun (funToCone F) = F"
proof
fix i
have "i \<notin> I \<Longrightarrow> coneToFun (funToCone F) i = F i"
using assms coneToFun_def S.extensional'_arb [of F I i] by auto
moreover have "i \<in> I \<Longrightarrow> coneToFun (funToCone F) i = F i"
proof -
assume i: "i \<in> I"
have "coneToFun (funToCone F) i = S.img (funToCone F i)"
using i coneToFun_def by simp
also have "... = S.img (S.mkPoint (D i) (F i))"
using i funToCone_def by auto
also have "... = F i"
using assms i is_discrete S.img_mkPoint(2) by force
finally show "coneToFun (funToCone F) i = F i" by auto
qed
ultimately show "coneToFun (funToCone F) i = F i" by auto
qed
lemma bij_coneToFun:
shows "bij_betw coneToFun (cones S.unity) (S.PiE' I (S.set o D))"
using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone
bij_betwI
by blast
lemma bij_funToCone:
shows "bij_betw funToCone (S.PiE' I (S.set o D)) (cones S.unity)"
using coneToFun_mapsto funToCone_mapsto funToCone_coneToFun coneToFun_funToCone
bij_betwI
by blast
end
context replete_set_category
begin
text\<open>
A set category admits @{term I}-indexed tupling if there is an injective map that takes
each extensional function from @{term I} to @{term Univ} to an element of @{term Univ}.
\<close>
definition admits_tupling
where "admits_tupling I \<equiv> \<exists>\<pi>. \<pi> \<in> PiE' I (\<lambda>_. Univ) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>_. Univ))"
lemma admits_tupling_monotone:
assumes "admits_tupling I" and "I' \<subseteq> I"
shows "admits_tupling I'"
proof -
from assms(1) obtain \<pi>
where \<pi>: "\<pi> \<in> PiE' I (\<lambda>_. Univ) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>_. Univ))"
using admits_tupling_def by metis
have "\<pi> \<in> PiE' I' (\<lambda>_. Univ) \<rightarrow> Univ"
proof
fix f
assume f: "f \<in> PiE' I' (\<lambda>_. Univ)"
have "f \<in> PiE' I (\<lambda>_. Univ)"
using assms(2) f extensional'_def [of I'] terminal_unity extensional'_monotone by auto
thus "\<pi> f \<in> Univ" using \<pi> by auto
qed
moreover have "inj_on \<pi> (PiE' I' (\<lambda>_. Univ))"
proof -
have 1: "\<And>F A A'. inj_on F A \<and> A' \<subseteq> A \<Longrightarrow> inj_on F A'"
using subset_inj_on by blast
moreover have "PiE' I' (\<lambda>_. Univ) \<subseteq> PiE' I (\<lambda>_. Univ)"
using assms(2) extensional'_def [of I'] terminal_unity by auto
ultimately show ?thesis using \<pi> assms(2) by blast
qed
ultimately show ?thesis using admits_tupling_def by metis
qed
lemma has_products_iff_admits_tupling:
fixes I :: "'i set"
shows "has_products I \<longleftrightarrow> I \<noteq> UNIV \<and> admits_tupling I"
proof
text\<open>
If @{term[source=true] S} has @{term I}-indexed products, then for every @{term I}-indexed
discrete diagram @{term D} in @{term[source=true] S} there is an object @{term \<Pi>D}
of @{term[source=true] S} whose points are in bijective correspondence with the set of
cones over @{term D} with apex @{term unity}. In particular this is true for
the diagram @{term D} that assigns to each element of @{term I} the
``universal object'' @{term "mkIde Univ"}.
\<close>
assume has_products: "has_products I"
have I: "I \<noteq> UNIV" using has_products has_products_def by auto
interpret J: discrete_category I \<open>SOME x. x \<notin> I\<close>
using I someI_ex [of "\<lambda>x. x \<notin> I"] by (unfold_locales, auto)
let ?D = "\<lambda>i. mkIde Univ"
interpret D: discrete_diagram_from_map I S ?D \<open>SOME j. j \<notin> I\<close>
using J.not_arr_null J.arr_char ide_mkIde
by (unfold_locales, auto)
interpret D: discrete_diagram_in_replete_set_category J.comp S D.map ..
have "discrete_diagram J.comp S D.map" ..
from this obtain \<Pi>D \<chi> where \<chi>: "product_cone J.comp S D.map \<Pi>D \<chi>"
using has_products has_products_def [of I] ex_productE [of "J.comp" D.map]
D.diagram_axioms
by blast
interpret \<chi>: product_cone J.comp S D.map \<Pi>D \<chi>
using \<chi> by auto
have "D.has_as_limit \<Pi>D"
using \<chi>.limit_cone_axioms by auto
hence \<Pi>D: "ide \<Pi>D \<and> (\<exists>\<phi>. bij_betw \<phi> (hom unity \<Pi>D) (D.cones unity))"
using D.limits_are_sets_of_cones by simp
from this obtain \<phi> where \<phi>: "bij_betw \<phi> (hom unity \<Pi>D) (D.cones unity)"
by blast
have \<phi>': "inv_into (hom unity \<Pi>D) \<phi> \<in> D.cones unity \<rightarrow> hom unity \<Pi>D \<and>
inj_on (inv_into (hom unity \<Pi>D) \<phi>) (D.cones unity)"
using \<phi> bij_betw_inv_into bij_betw_imp_inj_on bij_betw_imp_funcset by blast
let ?\<pi> = "img o (inv_into (hom unity \<Pi>D) \<phi>) o D.funToCone"
have 1: "D.funToCone \<in> PiE' I (set o D.map) \<rightarrow> D.cones unity"
using D.funToCone_mapsto extensional'_def [of I] by auto
have 2: "inv_into (hom unity \<Pi>D) \<phi> \<in> D.cones unity \<rightarrow> hom unity \<Pi>D"
using \<phi>' by auto
have 3: "img \<in> hom unity \<Pi>D \<rightarrow> Univ"
using img_point_in_Univ by blast
have 4: "inj_on D.funToCone (PiE' I (set o D.map))"
proof -
have "D.I = I" by auto
thus ?thesis
using D.bij_funToCone bij_betw_imp_inj_on by auto
qed
have 5: "inj_on (inv_into (hom unity \<Pi>D) \<phi>) (D.cones unity)"
using \<phi>' by auto
have 6: "inj_on img (hom unity \<Pi>D)"
using \<Pi>D bij_betw_points_and_set bij_betw_imp_inj_on [of img "hom unity \<Pi>D" "set \<Pi>D"]
by simp
have "?\<pi> \<in> PiE' I (set o D.map) \<rightarrow> Univ"
using 1 2 3 by force
moreover have "inj_on ?\<pi> (PiE' I (set o D.map))"
proof -
have 7: "\<And>A B C D F G H. F \<in> A \<rightarrow> B \<and> G \<in> B \<rightarrow> C \<and> H \<in> C \<rightarrow> D
\<and> inj_on F A \<and> inj_on G B \<and> inj_on H C
\<Longrightarrow> inj_on (H o G o F) A"
proof (intro inj_onI)
fix A :: "'a set" and B :: "'b set" and C :: "'c set" and D :: "'d set"
and F :: "'a \<Rightarrow> 'b" and G :: "'b \<Rightarrow> 'c" and H :: "'c \<Rightarrow> 'd"
assume a1: "F \<in> A \<rightarrow> B \<and> G \<in> B \<rightarrow> C \<and> H \<in> C \<rightarrow> D \<and>
inj_on F A \<and> inj_on G B \<and> inj_on H C"
fix a a'
assume a: "a \<in> A" and a': "a' \<in> A" and eq: "(H o G o F) a = (H o G o F) a'"
have "H (G (F a)) = H (G (F a'))" using eq by simp
moreover have "G (F a) \<in> C \<and> G (F a') \<in> C" using a a' a1 by auto
ultimately have "G (F a) = G (F a')" using a1 inj_onD by metis
moreover have "F a \<in> B \<and> F a' \<in> B" using a a' a1 by auto
ultimately have "F a = F a'" using a1 inj_onD by metis
thus "a = a'" using a a' a1 inj_onD by metis
qed
show ?thesis
using 1 2 3 4 5 6 7 [of D.funToCone "PiE' I (set o D.map)" "D.cones unity"
"inv_into (hom unity \<Pi>D) \<phi>" "hom unity \<Pi>D"
img Univ]
by fastforce
qed
moreover have "PiE' I (set o D.map) = PiE' I (\<lambda>x. Univ)"
proof -
have "\<And>i. i \<in> I \<Longrightarrow> (set o D.map) i = Univ"
using J.arr_char D.map_def set_mkIde by simp
thus ?thesis by blast
qed
ultimately have "?\<pi> \<in> (PiE' I (\<lambda>x. Univ)) \<rightarrow> Univ \<and> inj_on ?\<pi> (PiE' I (\<lambda>x. Univ))"
by auto
thus "I \<noteq> UNIV \<and> admits_tupling I"
using I admits_tupling_def by auto
next
assume ex_\<pi>: "I \<noteq> UNIV \<and> admits_tupling I"
show "has_products I"
proof (unfold has_products_def)
from ex_\<pi> obtain \<pi>
where \<pi>: "\<pi> \<in> (PiE' I (\<lambda>x. Univ)) \<rightarrow> Univ \<and> inj_on \<pi> (PiE' I (\<lambda>x. Univ))"
using admits_tupling_def by metis
text\<open>
Given an @{term I}-indexed discrete diagram @{term D}, obtain the object @{term \<Pi>D}
of @{term[source=true] S} corresponding to the set @{term "\<pi> ` PiE I D"} of all
@{term "\<pi> d"} where \<open>d \<in> d \<in> J \<rightarrow>\<^sub>E Univ\<close> and @{term "d i \<in> D i"}
for all @{term "i \<in> I"}.
The elements of @{term \<Pi>D} are in bijective correspondence with the set of cones
over @{term D}, hence @{term \<Pi>D} is a limit of @{term D}.
\<close>
have "\<And>J D. discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I
\<Longrightarrow> \<exists>\<Pi>D. has_as_product J D \<Pi>D"
proof
fix J :: "'i comp" and D
assume D: "discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I"
interpret J: category J
using D discrete_diagram.axioms(1) by blast
interpret D: discrete_diagram J S D
using D by simp
interpret D: discrete_diagram_in_replete_set_category J S D ..
let ?\<Pi>D = "mkIde (\<pi> ` PiE' I (set o D))"
have 0: "ide ?\<Pi>D"
proof -
have "set o D \<in> I \<rightarrow> Pow Univ"
using Pow_iff incl_in_def o_apply elem_set_implies_incl_in
set_subset_Univ subsetI
by (metis (mono_tags, lifting) Pi_I')
hence "\<pi> ` PiE' I (set o D) \<subseteq> Univ"
using \<pi> by blast
thus ?thesis using \<pi> ide_mkIde by simp
qed
hence set_\<Pi>D: "\<pi> ` PiE' I (set o D) = set ?\<Pi>D"
using 0 ide_in_hom arr_mkIde set_mkIde by auto
text\<open>
The elements of @{term \<Pi>D} are all values of the form @{term "\<pi> d"},
where @{term d} satisfies @{term "d i \<in> set (D i)"} for all @{term "i \<in> I"}.
Such @{term d} correspond bijectively to cones.
Since @{term \<pi>} is injective, the values @{term "\<pi> d"} correspond bijectively to cones.
\<close>
let ?\<phi> = "mkPoint ?\<Pi>D o \<pi> o D.coneToFun"
let ?\<phi>' = "D.funToCone o inv_into (PiE' I (set o D)) \<pi> o img"
have 1: "\<pi> \<in> PiE' I (set o D) \<rightarrow> set ?\<Pi>D \<and> inj_on \<pi> (PiE' I (set o D))"
proof -
have "PiE' I (set o D) \<subseteq> PiE' I (\<lambda>x. Univ)"
using set_subset_Univ elem_set_implies_incl_in elem_set_implies_set_eq_singleton
incl_in_def PiE'_mono
by (metis comp_apply subsetI)
thus ?thesis using \<pi> subset_inj_on set_\<Pi>D Pi_I' imageI by fastforce
qed
have 2: "inv_into (PiE' I (set o D)) \<pi> \<in> set ?\<Pi>D \<rightarrow> PiE' I (set o D)"
proof
fix y
assume y: "y \<in> set ?\<Pi>D"
have "y \<in> \<pi> ` (PiE' I (set o D))" using y set_\<Pi>D by auto
thus "inv_into (PiE' I (set o D)) \<pi> y \<in> PiE' I (set o D)"
using inv_into_into [of y \<pi> "PiE' I (set o D)"] by simp
qed
have 3: "\<And>x. x \<in> set ?\<Pi>D \<Longrightarrow> \<pi> (inv_into (PiE' I (set o D)) \<pi> x) = x"
using set_\<Pi>D by (simp add: f_inv_into_f)
have 4: "\<And>d. d \<in> PiE' I (set o D) \<Longrightarrow> inv_into (PiE' I (set o D)) \<pi> (\<pi> d) = d"
using 1 by auto
have 5: "D.I = I"
using D by auto
have "bij_betw ?\<phi> (D.cones unity) (hom unity ?\<Pi>D)"
proof (intro bij_betwI)
show "?\<phi> \<in> D.cones unity \<rightarrow> hom unity ?\<Pi>D"
proof
fix \<chi>
assume \<chi>: "\<chi> \<in> D.cones unity"
show "?\<phi> \<chi> \<in> hom unity ?\<Pi>D"
using \<chi> 0 1 5 D.coneToFun_mapsto mkPoint_in_hom [of ?\<Pi>D]
by (simp, blast)
qed
show "?\<phi>' \<in> hom unity ?\<Pi>D \<rightarrow> D.cones unity"
proof
fix x
assume x: "x \<in> hom unity ?\<Pi>D"
hence "img x \<in> set ?\<Pi>D"
using img_point_elem_set by blast
hence "inv_into (PiE' I (set o D)) \<pi> (img x) \<in> Pi I (set \<circ> D) \<inter> extensional' I"
using 2 by blast
thus "?\<phi>' x \<in> D.cones unity"
using 5 D.funToCone_mapsto by auto
qed
show "\<And>x. x \<in> hom unity ?\<Pi>D \<Longrightarrow> ?\<phi> (?\<phi>' x) = x"
proof -
fix x
assume x: "x \<in> hom unity ?\<Pi>D"
show "?\<phi> (?\<phi>' x) = x"
proof -
have "D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \<pi> (img x)))
= inv_into (PiE' I (set o D)) \<pi> (img x)"
using x 1 5 img_point_elem_set set_\<Pi>D D.coneToFun_funToCone by force
hence "\<pi> (D.coneToFun (D.funToCone (inv_into (PiE' I (set o D)) \<pi> (img x))))
= img x"
using x 3 img_point_elem_set set_\<Pi>D by force
thus ?thesis using x 0 mkPoint_img by auto
qed
qed
show "\<And>\<chi>. \<chi> \<in> D.cones unity \<Longrightarrow> ?\<phi>' (?\<phi> \<chi>) = \<chi>"
proof -
fix \<chi>
assume \<chi>: "\<chi> \<in> D.cones unity"
show "?\<phi>' (?\<phi> \<chi>) = \<chi>"
proof -
have "img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>))) = \<pi> (D.coneToFun \<chi>)"
using \<chi> 0 1 5 D.coneToFun_mapsto img_mkPoint(2) by blast
hence "inv_into (PiE' I (set o D)) \<pi> (img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>))))
= D.coneToFun \<chi>"
using \<chi> D.coneToFun_mapsto 4 5 by (metis PiE)
hence "D.funToCone (inv_into (PiE' I (set o D)) \<pi>
(img (mkPoint ?\<Pi>D (\<pi> (D.coneToFun \<chi>)))))
= \<chi>"
using \<chi> D.funToCone_coneToFun by auto
thus ?thesis by auto
qed
qed
qed
hence "bij_betw (inv_into (D.cones unity) ?\<phi>) (hom unity ?\<Pi>D) (D.cones unity)"
using bij_betw_inv_into by blast
hence "\<exists>\<phi>. bij_betw \<phi> (hom unity ?\<Pi>D) (D.cones unity)" by blast
hence "D.has_as_limit ?\<Pi>D"
using \<open>ide ?\<Pi>D\<close> D.limits_are_sets_of_cones by simp
from this obtain \<chi> where \<chi>: "limit_cone J S D ?\<Pi>D \<chi>" by blast
interpret \<chi>: limit_cone J S D ?\<Pi>D \<chi> using \<chi> by auto
interpret P: product_cone J S D ?\<Pi>D \<chi>
using \<chi> D.product_coneI by blast
have "product_cone J S D ?\<Pi>D \<chi>" ..
thus "has_as_product J D ?\<Pi>D"
using has_as_product_def by auto
qed
thus "I \<noteq> UNIV \<and>
(\<forall>J D. discrete_diagram J S D \<and> Collect (partial_magma.arr J) = I
\<longrightarrow> (\<exists>\<Pi>D. has_as_product J D \<Pi>D))"
using ex_\<pi> by blast
qed
qed
text\<open>
Characterization of the completeness properties enjoyed by a set category:
A set category @{term[source=true] S} has all limits at a type @{typ 'j},
if and only if @{term[source=true] S} admits @{term I}-indexed tupling
for all @{typ 'j}-sets @{term I} such that @{term "I \<noteq> UNIV"}.
\<close>
theorem has_limits_iff_admits_tupling:
shows "has_limits (undefined :: 'j) \<longleftrightarrow> (\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I)"
proof
assume has_limits: "has_limits (undefined :: 'j)"
show "\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I"
using has_limits has_products_if_has_limits has_products_iff_admits_tupling by blast
next
assume admits_tupling: "\<forall>I :: 'j set. I \<noteq> UNIV \<longrightarrow> admits_tupling I"
show "has_limits (undefined :: 'j)"
proof -
have 1: "\<And>I :: 'j set. I \<noteq> UNIV \<Longrightarrow> has_products I"
using admits_tupling has_products_iff_admits_tupling by auto
have "\<And>J :: 'j comp. category J \<Longrightarrow> has_products (Collect (partial_magma.arr J))"
proof -
fix J :: "'j comp"
assume J: "category J"
interpret J: category J using J by auto
have "Collect J.arr \<noteq> UNIV" using J.not_arr_null by blast
thus "has_products (Collect J.arr)"
using 1 by simp
qed
hence "\<And>J :: 'j comp. category J \<Longrightarrow> has_limits_of_shape J"
proof -
fix J :: "'j comp"
assume J: "category J"
interpret J: category J using J by auto
show "has_limits_of_shape J"
proof -
have "Collect J.arr \<noteq> UNIV" using J.not_arr_null by fast
moreover have "Collect J.ide \<noteq> UNIV" using J.not_arr_null by blast
ultimately show ?thesis
using 1 has_limits_if_has_products J.category_axioms by metis
qed
qed
thus "has_limits (undefined :: 'j)"
using has_limits_def by metis
qed
qed
end
section "Limits in Functor Categories"
text\<open>
In this section, we consider the special case of limits in functor categories,
with the objective of showing that limits in a functor category \<open>[A, B]\<close>
are given pointwise, and that \<open>[A, B]\<close> has all limits that @{term B} has.
\<close>
locale parametrized_diagram =
J: category J +
A: category A +
B: category B +
JxA: product_category J A +
binary_functor J A B D
for J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and A :: "'a comp" (infixr "\<cdot>\<^sub>A" 55)
and B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and D :: "'j * 'a \<Rightarrow> 'b"
begin
(* Notation for A.in_hom and B.in_hom is being inherited, but from where? *)
notation J.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J _\<guillemotright>")
notation JxA.comp (infixr "\<cdot>\<^sub>J\<^sub>x\<^sub>A" 55)
notation JxA.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J\<^sub>x\<^sub>A _\<guillemotright>")
text\<open>
A choice of limit cone for each diagram \<open>D (-, a)\<close>, where @{term a}
is an object of @{term[source=true] A}, extends to a functor \<open>L: A \<rightarrow> B\<close>,
where the action of @{term L} on arrows of @{term[source=true] A} is determined by
universality.
\<close>
abbreviation L
where "L \<equiv> \<lambda>l \<chi>. \<lambda>a. if A.arr a then
limit_cone.induced_arrow J B (\<lambda>j. D (j, A.cod a))
(l (A.cod a)) (\<chi> (A.cod a))
(l (A.dom a)) (vertical_composite.map J B
(\<chi> (A.dom a)) (\<lambda>j. D (j, a)))
else B.null"
abbreviation P
where "P \<equiv> \<lambda>l \<chi>. \<lambda>a f. \<guillemotleft>f : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright> \<and>
diagram.cones_map J B (\<lambda>j. D (j, A.cod a)) f (\<chi> (A.cod a)) =
vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a))"
lemma L_arr:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "\<And>a. A.arr a \<Longrightarrow> (\<exists>!f. P l \<chi> a f) \<and> P l \<chi> a (L l \<chi> a)"
proof
fix a
assume a: "A.arr a"
interpret \<chi>_dom_a: limit_cone J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>l (A.dom a)\<close> \<open>\<chi> (A.dom a)\<close>
using a assms by auto
interpret \<chi>_cod_a: limit_cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.cod a)\<close> \<open>\<chi> (A.cod a)\<close>
using a assms by auto
interpret Da: natural_transformation J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close>
\<open>\<lambda>j. D (j, a)\<close>
using a fixing_arr_gives_natural_transformation_2 by simp
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close>
\<open>\<chi> (A.dom a)\<close> \<open>\<lambda>j. D (j, a)\<close> ..
interpret Dao\<chi>_dom_a: cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.dom a)\<close> Dao\<chi>_dom_a.map ..
show "P l \<chi> a (L l \<chi> a)"
using a Dao\<chi>_dom_a.cone_axioms \<chi>_cod_a.induced_arrowI [of Dao\<chi>_dom_a.map "l (A.dom a)"]
by auto
show "\<exists>!f. P l \<chi> a f"
using \<chi>_cod_a.is_universal Dao\<chi>_dom_a.cone_axioms by blast
qed
lemma L_ide:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "\<And>a. A.ide a \<Longrightarrow> L l \<chi> a = l a"
proof -
let ?L = "L l \<chi>"
let ?P = "P l \<chi>"
fix a
assume a: "A.ide a"
interpret \<chi>a: limit_cone J B \<open>\<lambda>j. D (j, a)\<close> \<open>l a\<close> \<open>\<chi> a\<close> using a assms by auto
have Pa: "?P a = (\<lambda>f. f \<in> B.hom (l a) (l a) \<and>
diagram.cones_map J B (\<lambda>j. D (j, a)) f (\<chi> a) = \<chi> a)"
using a vcomp_ide_dom \<chi>a.natural_transformation_axioms by simp
have "?P a (?L a)" using assms a L_arr [of l \<chi> a] by fastforce
moreover have "?P a (l a)"
proof -
have "?P a (l a) \<longleftrightarrow> l a \<in> B.hom (l a) (l a) \<and> \<chi>a.D.cones_map (l a) (\<chi> a) = \<chi> a"
using Pa by meson
thus ?thesis
using a \<chi>a.ide_apex \<chi>a.cone_axioms \<chi>a.D.cones_map_ide [of "\<chi> a" "l a"] by force
qed
moreover have "\<exists>!f. ?P a f"
using a Pa \<chi>a.is_universal \<chi>a.cone_axioms by force
ultimately show "?L a = l a" by blast
qed
lemma chosen_limits_induce_functor:
assumes "\<forall>a. A.ide a \<longrightarrow> limit_cone J B (\<lambda>j. D (j, a)) (l a) (\<chi> a)"
shows "functor A B (L l \<chi>)"
proof -
let ?L = "L l \<chi>"
let ?P = "\<lambda>a. \<lambda>f. \<guillemotleft>f : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright> \<and>
diagram.cones_map J B (\<lambda>j. D (j, A.cod a)) f (\<chi> (A.cod a))
= vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a))"
interpret L: "functor" A B ?L
apply unfold_locales
using assms L_arr [of l] L_ide
apply auto[4]
proof -
fix a' a
assume 1: "A.arr (A a' a)"
have a: "A.arr a" using 1 by auto
have a': "\<guillemotleft>a' : A.cod a \<rightarrow>\<^sub>A A.cod a'\<guillemotright>" using 1 by auto
have a'a: "A.seq a' a" using 1 by auto
interpret \<chi>_dom_a: limit_cone J B \<open>\<lambda>j. D (j, A.dom a)\<close> \<open>l (A.dom a)\<close> \<open>\<chi> (A.dom a)\<close>
using a assms by auto
interpret \<chi>_cod_a: limit_cone J B \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>l (A.cod a)\<close> \<open>\<chi> (A.cod a)\<close>
using a'a assms by auto
interpret \<chi>_dom_a'a: limit_cone J B \<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.dom (a' \<cdot>\<^sub>A a))\<close>
using a'a assms by auto
interpret \<chi>_cod_a'a: limit_cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.cod (a' \<cdot>\<^sub>A a))\<close>
using a'a assms by auto
interpret Da: natural_transformation J B
\<open>\<lambda>j. D (j, A.dom a)\<close> \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, a)\<close>
using a fixing_arr_gives_natural_transformation_2 by simp
interpret Da': natural_transformation J B
\<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>\<lambda>j. D (j, a')\<close>
using a a'a fixing_arr_gives_natural_transformation_2 by fastforce
interpret Da'o\<chi>_cod_a: vertical_composite J B
\<chi>_cod_a.A.map \<open>\<lambda>j. D (j, A.cod a)\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<chi> (A.cod a)\<close> \<open>\<lambda>j. D (j, a')\<close>..
interpret Da'o\<chi>_cod_a: cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>l (A.cod a)\<close> Da'o\<chi>_cod_a.map
..
interpret Da'a: natural_transformation J B
\<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close> \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, a' \<cdot>\<^sub>A a)\<close>
using a'a fixing_arr_gives_natural_transformation_2 [of "a' \<cdot>\<^sub>A a"] by auto
interpret Da'ao\<chi>_dom_a'a:
vertical_composite J B \<chi>_dom_a'a.A.map \<open>\<lambda>j. D (j, A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close> \<open>\<chi> (A.dom (a' \<cdot>\<^sub>A a))\<close>
\<open>\<lambda>j. D (j, a' \<cdot>\<^sub>A a)\<close> ..
interpret Da'ao\<chi>_dom_a'a: cone J B \<open>\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))\<close>
\<open>l (A.dom (a' \<cdot>\<^sub>A a))\<close> Da'ao\<chi>_dom_a'a.map ..
show "?L (a' \<cdot>\<^sub>A a) = ?L a' \<cdot>\<^sub>B ?L a"
proof -
have "?P (a' \<cdot>\<^sub>A a) (?L (a' \<cdot>\<^sub>A a))" using assms a'a L_arr [of l \<chi> "a' \<cdot>\<^sub>A a"] by fastforce
moreover have "?P (a' \<cdot>\<^sub>A a) (?L a' \<cdot>\<^sub>B ?L a)"
proof
have La: "\<guillemotleft>?L a : l (A.dom a) \<rightarrow>\<^sub>B l (A.cod a)\<guillemotright>"
using assms a L_arr by fast
moreover have La': "\<guillemotleft>?L a' : l (A.cod a) \<rightarrow>\<^sub>B l (A.cod a')\<guillemotright>"
using assms a a' L_arr [of l \<chi> a'] by auto
ultimately have seq: "B.seq (?L a') (?L a)" by (elim B.in_homE, auto)
thus La'_La: "\<guillemotleft>?L a' \<cdot>\<^sub>B ?L a : l (A.dom (a' \<cdot>\<^sub>A a)) \<rightarrow>\<^sub>B l (A.cod (a' \<cdot>\<^sub>A a))\<guillemotright>"
using a a' 1 La La' by (intro B.comp_in_homI, auto)
show "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a)))
= Da'ao\<chi>_dom_a'a.map"
proof -
have "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a)))
= (\<chi>_cod_a'a.D.cones_map (?L a) o \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi> (A.cod a'))"
proof -
have "\<chi>_cod_a'a.D.cones_map (?L a' \<cdot>\<^sub>B ?L a) (\<chi> (A.cod (a' \<cdot>\<^sub>A a))) =
restrict (\<chi>_cod_a'a.D.cones_map (?L a) \<circ> \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi>_cod_a'a.D.cones (B.cod (?L a')))
(\<chi> (A.cod (a' \<cdot>\<^sub>A a)))"
using seq \<chi>_cod_a'a.cone_axioms \<chi>_cod_a'a.D.cones_map_comp [of "?L a'" "?L a"]
by argo
also have "... = (\<chi>_cod_a'a.D.cones_map (?L a) o \<chi>_cod_a'a.D.cones_map (?L a'))
(\<chi> (A.cod a'))"
proof -
have "\<chi> (A.cod a') \<in> \<chi>_cod_a'a.D.cones (l (A.cod a'))"
using \<chi>_cod_a'a.cone_axioms a'a by simp
moreover have "B.cod (?L a') = l (A.cod a')"
using assms a' L_arr [of l] by auto
ultimately show ?thesis
using a' a'a by simp
qed
finally show ?thesis by blast
qed
also have "... = \<chi>_cod_a'a.D.cones_map (?L a)
(\<chi>_cod_a'a.D.cones_map (?L a') (\<chi> (A.cod a')))"
by simp
also have "... = \<chi>_cod_a'a.D.cones_map (?L a) Da'o\<chi>_cod_a.map"
proof -
have "?P a' (?L a')" using assms a' L_arr [of l \<chi> a'] by fast
moreover have
"?P a' = (\<lambda>f. f \<in> B.hom (l (A.cod a)) (l (A.cod a')) \<and>
\<chi>_cod_a'a.D.cones_map f (\<chi> (A.cod a')) = Da'o\<chi>_cod_a.map)"
using a'a by force
ultimately show ?thesis using a'a by force
qed
also have "... = vertical_composite.map J B
(\<chi>_cod_a.D.cones_map (?L a) (\<chi> (A.cod a)))
(\<lambda>j. D (j, a'))"
using assms \<chi>_cod_a.D.diagram_axioms \<chi>_cod_a'a.D.diagram_axioms
Da'.natural_transformation_axioms \<chi>_cod_a.cone_axioms La
cones_map_vcomp [of J B "\<lambda>j. D (j, A.cod a)" "\<lambda>j. D (j, A.cod (a' \<cdot>\<^sub>A a))"
"\<lambda>j. D (j, a')" "l (A.cod a)" "\<chi> (A.cod a)"
"?L a" "l (A.dom a)"]
by blast
also have "... = vertical_composite.map J B
(vertical_composite.map J B (\<chi> (A.dom a)) (\<lambda>j. D (j, a)))
(\<lambda>j. D (j, a'))"
using assms a L_arr by presburger
also have "... = vertical_composite.map J B (\<chi> (A.dom a))
(vertical_composite.map J B (\<lambda>j. D (j, a)) (\<lambda>j. D (j, a')))"
using a'a Da.natural_transformation_axioms Da'.natural_transformation_axioms
\<chi>_dom_a.natural_transformation_axioms vcomp_assoc
by auto
also have
"... = vertical_composite.map J B (\<chi> (A.dom (a' \<cdot>\<^sub>A a))) (\<lambda>j. D (j, a' \<cdot>\<^sub>A a))"
using a'a preserves_comp_2 by simp
finally show ?thesis by auto
qed
qed
moreover have "\<exists>!f. ?P (a' \<cdot>\<^sub>A a) f"
using \<chi>_cod_a'a.is_universal
[of "l (A.dom (a' \<cdot>\<^sub>A a))"
"vertical_composite.map J B (\<chi> (A.dom (a' \<cdot>\<^sub>A a))) (\<lambda>j. D (j, a' \<cdot>\<^sub>A a))"]
Da'ao\<chi>_dom_a'a.cone_axioms
by fast
ultimately show ?thesis by blast
qed
qed
show ?thesis ..
qed
end
locale diagram_in_functor_category =
A: category A +
B: category B +
A_B: functor_category A B +
diagram J A_B.comp D
for A :: "'a comp" (infixr "\<cdot>\<^sub>A" 55)
and B :: "'b comp" (infixr "\<cdot>\<^sub>B" 55)
and J :: "'j comp" (infixr "\<cdot>\<^sub>J" 55)
and D :: "'j \<Rightarrow> ('a, 'b) functor_category.arr"
begin
interpretation JxA: product_category J A ..
interpretation A_BxA: product_category A_B.comp A ..
interpretation E: evaluation_functor A B ..
interpretation Curry: currying J A B ..
notation JxA.comp (infixr "\<cdot>\<^sub>J\<^sub>x\<^sub>A" 55)
notation JxA.in_hom ("\<guillemotleft>_ : _ \<rightarrow>\<^sub>J\<^sub>x\<^sub>A _\<guillemotright>")
text\<open>
Evaluation of a functor or natural transformation from @{term[source=true] J}
to \<open>[A, B]\<close> at an arrow @{term a} of @{term[source=true] A}.
\<close>
abbreviation at
where "at a \<tau> \<equiv> \<lambda>j. Curry.uncurry \<tau> (j, a)"
lemma at_simp:
assumes "A.arr a" and "J.arr j" and "A_B.arr (\<tau> j)"
shows "at a \<tau> j = A_B.Map (\<tau> j) a"
using assms Curry.uncurry_def E.map_simp by simp
lemma functor_at_ide_is_functor:
assumes "functor J A_B.comp F" and "A.ide a"
shows "functor J B (at a F)"
proof -
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using assms(1) Curry.uncurry_preserves_functors by simp
interpret uncurry_F: binary_functor J A B \<open>Curry.uncurry F\<close> ..
show ?thesis using assms(2) uncurry_F.fixing_ide_gives_functor_2 by simp
qed
lemma functor_at_arr_is_transformation:
assumes "functor J A_B.comp F" and "A.arr a"
shows "natural_transformation J B (at (A.dom a) F) (at (A.cod a) F) (at a F)"
proof -
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using assms(1) Curry.uncurry_preserves_functors by simp
interpret uncurry_F: binary_functor J A B \<open>Curry.uncurry F\<close> ..
show ?thesis
using assms(2) uncurry_F.fixing_arr_gives_natural_transformation_2 by simp
qed
lemma transformation_at_ide_is_transformation:
assumes "natural_transformation J A_B.comp F G \<tau>" and "A.ide a"
shows "natural_transformation J B (at a F) (at a G) (at a \<tau>)"
proof -
interpret \<tau>: natural_transformation J A_B.comp F G \<tau> using assms(1) by auto
interpret uncurry_F: "functor" JxA.comp B \<open>Curry.uncurry F\<close>
using Curry.uncurry_preserves_functors \<tau>.F.functor_axioms by simp
interpret uncurry_f: binary_functor J A B \<open>Curry.uncurry F\<close> ..
interpret uncurry_G: "functor" JxA.comp B \<open>Curry.uncurry G\<close>
using Curry.uncurry_preserves_functors \<tau>.G.functor_axioms by simp
interpret uncurry_G: binary_functor J A B \<open>Curry.uncurry G\<close> ..
interpret uncurry_\<tau>: natural_transformation
JxA.comp B \<open>Curry.uncurry F\<close> \<open>Curry.uncurry G\<close> \<open>Curry.uncurry \<tau>\<close>
using Curry.uncurry_preserves_transformations \<tau>.natural_transformation_axioms
by simp
interpret uncurry_\<tau>: binary_functor_transformation J A B
\<open>Curry.uncurry F\<close> \<open>Curry.uncurry G\<close> \<open>Curry.uncurry \<tau>\<close> ..
show ?thesis
using assms(2) uncurry_\<tau>.fixing_ide_gives_natural_transformation_2 by simp
qed
lemma constant_at_ide_is_constant:
assumes "cone x \<chi>" and a: "A.ide a"
shows "at a (constant_functor.map J A_B.comp x) =
constant_functor.map J B (A_B.Map x a)"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close>
using x A_B.ide_char by simp
interpret Da: "functor" J B \<open>at a D\<close>
using a functor_at_ide_is_functor functor_axioms by blast
interpret Da: diagram J B \<open>at a D\<close> ..
interpret Xa: constant_functor J B \<open>A_B.Map x a\<close>
using a Fun_x.preserves_ide by unfold_locales simp
show "at a \<chi>.A.map = Xa.map"
using a x Curry.uncurry_def E.map_def Xa.is_extensional by auto
qed
lemma at_ide_is_diagram:
assumes a: "A.ide a"
shows "diagram J B (at a D)"
proof -
interpret Da: "functor" J B "at a D"
using a functor_at_ide_is_functor functor_axioms by simp
show ?thesis ..
qed
lemma cone_at_ide_is_cone:
assumes "cone x \<chi>" and a: "A.ide a"
shows "diagram.cone J B (at a D) (A_B.Map x a) (at a \<chi>)"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close>
using x A_B.ide_char by simp
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret Xa: constant_functor J B \<open>A_B.Map x a\<close>
using a by (unfold_locales, simp)
interpret \<chi>a: natural_transformation J B Xa.map \<open>at a D\<close> \<open>at a \<chi>\<close>
using assms(1) x a transformation_at_ide_is_transformation \<chi>.natural_transformation_axioms
constant_at_ide_is_constant
by fastforce
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close> ..
show cone_\<chi>a: "Da.cone (A_B.Map x a) (at a \<chi>)" ..
qed
lemma at_preserves_comp:
assumes "A.seq a' a"
shows "at (A a' a) D = vertical_composite.map J B (at a D) (at a' D)"
proof -
interpret Da: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at a D\<close>
using assms functor_at_arr_is_transformation functor_axioms by blast
interpret Da': natural_transformation J B \<open>at (A.cod a) D\<close> \<open>at (A.cod a') D\<close> \<open>at a' D\<close>
using assms functor_at_arr_is_transformation [of D a'] functor_axioms by fastforce
interpret Da'oDa: vertical_composite J B
\<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at (A.cod a') D\<close>
\<open>at a D\<close> \<open>at a' D\<close> ..
interpret Da'a: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a') D\<close>
\<open>at (a' \<cdot>\<^sub>A a) D\<close>
using assms functor_at_arr_is_transformation [of D "a' \<cdot>\<^sub>A a"] functor_axioms by simp
show "at (a' \<cdot>\<^sub>A a) D = Da'oDa.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) Da'oDa.map" ..
show "natural_transformation J B (at (A.dom a) D) (at (A.cod a') D) (at (a' \<cdot>\<^sub>A a) D)" ..
show "\<And>j. J.ide j \<Longrightarrow> at (a' \<cdot>\<^sub>A a) D j = Da'oDa.map j"
proof -
fix j
assume j: "J.ide j"
interpret Dj: "functor" A B \<open>A_B.Map (D j)\<close>
using j preserves_ide A_B.ide_char by simp
show "at (a' \<cdot>\<^sub>A a) D j = Da'oDa.map j"
using assms j Dj.preserves_comp at_simp Da'oDa.map_simp_ide by auto
qed
qed
qed
lemma cones_map_pointwise:
assumes "cone x \<chi>" and "cone x' \<chi>'"
and f: "f \<in> A_B.hom x' x"
shows "cones_map f \<chi> = \<chi>' \<longleftrightarrow>
(\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>')"
proof
interpret \<chi>: cone J A_B.comp D x \<chi> using assms(1) by auto
interpret \<chi>': cone J A_B.comp D x' \<chi>' using assms(2) by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
have x': "A_B.ide x'" using \<chi>'.ide_apex by auto
interpret \<chi>f: cone J A_B.comp D x' \<open>cones_map f \<chi>\<close>
using x' f assms(1) cones_map_mapsto by blast
interpret Fun_x: "functor" A B \<open>A_B.Map x\<close> using x A_B.ide_char by simp
interpret Fun_x': "functor" A B \<open>A_B.Map x'\<close> using x' A_B.ide_char by simp
show "cones_map f \<chi> = \<chi>' \<Longrightarrow>
(\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>')"
proof -
assume \<chi>': "cones_map f \<chi> = \<chi>'"
have "\<And>a. A.ide a \<Longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using a assms(1) cone_at_ide_is_cone by simp
interpret \<chi>'a: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close> \<open>at a \<chi>'\<close>
using a assms(2) cone_at_ide_is_cone by simp
have 1: "\<guillemotleft>A_B.Map f a : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>"
using f a A_B.arr_char A_B.Map_cod A_B.Map_dom mem_Collect_eq
natural_transformation.preserves_hom A.ide_in_hom
by (metis (no_types, lifting) A_B.in_homE)
interpret \<chi>fa: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close>
\<open>Da.cones_map (A_B.Map f a) (at a \<chi>)\<close>
using 1 \<chi>a.cone_axioms Da.cones_map_mapsto by force
show "Da.cones_map (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j"
using \<chi>'a.is_extensional \<chi>fa.is_extensional [of j] by simp
moreover have "J.arr j \<Longrightarrow> Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j"
using a f 1 \<chi>.cone_axioms \<chi>a.cone_axioms at_simp
apply simp
apply (elim A_B.in_homE B.in_homE, auto)
using \<chi>' \<chi>.A.map_simp A_B.Map_comp [of "\<chi> j" f a a] by auto
ultimately show "Da.cones_map (A_B.Map f a) (at a \<chi>) j = at a \<chi>' j" by blast
qed
qed
thus "\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
by simp
qed
show "\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'
\<Longrightarrow> cones_map f \<chi> = \<chi>'"
proof -
assume A:
"\<forall>a. A.ide a \<longrightarrow> diagram.cones_map J B (at a D) (A_B.Map f a) (at a \<chi>) = at a \<chi>'"
show "cones_map f \<chi> = \<chi>'"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J A_B.comp \<chi>'.A.map D (cones_map f \<chi>)" ..
show "natural_transformation J A_B.comp \<chi>'.A.map D \<chi>'" ..
show "\<And>j. J.ide j \<Longrightarrow> cones_map f \<chi> j = \<chi>' j"
proof (intro A_B.arr_eqI)
fix j
assume j: "J.ide j"
show 1: "A_B.arr (cones_map f \<chi> j)"
using j \<chi>f.preserves_reflects_arr by simp
show "A_B.arr (\<chi>' j)" using j by auto
have Dom_\<chi>f_j: "A_B.Dom (cones_map f \<chi> j) = A_B.Map x'"
using x' j 1 A_B.Map_dom \<chi>'.A.map_simp \<chi>f.preserves_dom J.ide_in_hom
by (metis (no_types, lifting) J.ideD(2) \<chi>f.preserves_reflects_arr)
also have Dom_\<chi>'_j: "... = A_B.Dom (\<chi>' j)"
using x' j A_B.Map_dom [of "\<chi>' j"] \<chi>'.preserves_hom \<chi>'.A.map_simp by simp
finally show "A_B.Dom (cones_map f \<chi> j) = A_B.Dom (\<chi>' j)" by auto
have Cod_\<chi>f_j: "A_B.Cod (cones_map f \<chi> j) = A_B.Map (D (J.cod j))"
using j A_B.Map_cod A_B.cod_char J.ide_in_hom \<chi>f.preserves_hom
by (metis (no_types, lifting) "1" J.ideD(1) \<chi>f.preserves_cod)
also have Cod_\<chi>'_j: "... = A_B.Cod (\<chi>' j)"
using j A_B.Map_cod [of "\<chi>' j"] \<chi>'.preserves_hom by simp
finally show "A_B.Cod (cones_map f \<chi> j) = A_B.Cod (\<chi>' j)" by auto
show "A_B.Map (cones_map f \<chi> j) = A_B.Map (\<chi>' j)"
proof (intro NaturalTransformation.eqI)
interpret \<chi>fj: natural_transformation A B \<open>A_B.Map x'\<close> \<open>A_B.Map (D (J.cod j))\<close>
\<open>A_B.Map (cones_map f \<chi> j)\<close>
using j \<chi>f.preserves_reflects_arr A_B.arr_char [of "cones_map f \<chi> j"]
Dom_\<chi>f_j Cod_\<chi>f_j
by simp
show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j)))
(A_B.Map (cones_map f \<chi> j))" ..
interpret \<chi>'j: natural_transformation A B \<open>A_B.Map x'\<close> \<open>A_B.Map (D (J.cod j))\<close>
\<open>A_B.Map (\<chi>' j)\<close>
using j A_B.arr_char [of "\<chi>' j"] Dom_\<chi>'_j Cod_\<chi>'_j by simp
show "natural_transformation A B (A_B.Map x') (A_B.Map (D (J.cod j)))
(A_B.Map (\<chi>' j))" ..
show "\<And>a. A.ide a \<Longrightarrow> A_B.Map (cones_map f \<chi> j) a = A_B.Map (\<chi>' j) a"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
have cone_\<chi>a: "Da.cone (A_B.Map x a) (at a \<chi>)"
using a assms(1) cone_at_ide_is_cone by simp
interpret \<chi>a: cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using cone_\<chi>a by auto
interpret Fun_f: natural_transformation A B \<open>A_B.Dom f\<close> \<open>A_B.Cod f\<close>
\<open>A_B.Map f\<close>
using f A_B.arr_char by fast
have fa: "A_B.Map f a \<in> B.hom (A_B.Map x' a) (A_B.Map x a)"
using a f Fun_f.preserves_hom A.ide_in_hom by auto
have "A_B.Map (cones_map f \<chi> j) a = Da.cones_map (A_B.Map f a) (at a \<chi>) j"
proof -
have "A_B.Map (cones_map f \<chi> j) a = A_B.Map (A_B.comp (\<chi> j) f) a"
using assms(1) f \<chi>.is_extensional by auto
also have "... = B (A_B.Map (\<chi> j) a) (A_B.Map f a)"
using f j a \<chi>.preserves_hom A.ide_in_hom J.ide_in_hom A_B.Map_comp
\<chi>.A.map_simp
by (metis (no_types, lifting) A.comp_ide_self A.ideD(1) A_B.seqI'
J.ideD(1) mem_Collect_eq)
also have "... = Da.cones_map (A_B.Map f a) (at a \<chi>) j"
using j a cone_\<chi>a fa Curry.uncurry_def E.map_simp by auto
finally show ?thesis by auto
qed
also have "... = at a \<chi>' j" using j a A by simp
also have "... = A_B.Map (\<chi>' j) a"
using j Curry.uncurry_def E.map_simp \<chi>'j.is_extensional by simp
finally show "A_B.Map (cones_map f \<chi> j) a = A_B.Map (\<chi>' j) a" by auto
qed
qed
qed
qed
qed
qed
text\<open>
If @{term \<chi>} is a cone with apex @{term a} over @{term D}, then @{term \<chi>}
is a limit cone if, for each object @{term x} of @{term X}, the cone obtained
by evaluating @{term \<chi>} at @{term x} is a limit cone with apex @{term "A_B.Map a x"}
for the diagram in @{term C} obtained by evaluating @{term D} at @{term x}.
\<close>
lemma cone_is_limit_if_pointwise_limit:
assumes cone_\<chi>: "cone x \<chi>"
and "\<forall>a. A.ide a \<longrightarrow> diagram.limit_cone J B (at a D) (A_B.Map x a) (at a \<chi>)"
shows "limit_cone x \<chi>"
proof -
interpret \<chi>: cone J A_B.comp D x \<chi> using assms by auto
have x: "A_B.ide x" using \<chi>.ide_apex by auto
show "limit_cone x \<chi>"
proof
fix x' \<chi>'
assume cone_\<chi>': "cone x' \<chi>'"
interpret \<chi>': cone J A_B.comp D x' \<chi>' using cone_\<chi>' by auto
have x': "A_B.ide x'" using \<chi>'.ide_apex by auto
text\<open>
The universality of the limit cone \<open>at a \<chi>\<close> yields, for each object
\<open>a\<close> of \<open>A\<close>, a unique arrow \<open>fa\<close> that transforms
\<open>at a \<chi>\<close> to \<open>at a \<chi>'\<close>.
\<close>
have EU: "\<And>a. A.ide a \<Longrightarrow>
\<exists>!fa. fa \<in> B.hom (A_B.Map x' a) (A_B.Map x a) \<and>
diagram.cones_map J B (at a D) fa (at a \<chi>) = at a \<chi>'"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret \<chi>a: limit_cone J B \<open>at a D\<close> \<open>A_B.Map x a\<close> \<open>at a \<chi>\<close>
using assms(2) a by auto
interpret \<chi>'a: cone J B \<open>at a D\<close> \<open>A_B.Map x' a\<close> \<open>at a \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
have "Da.cone (A_B.Map x' a) (at a \<chi>')" ..
thus "\<exists>!fa. fa \<in> B.hom (A_B.Map x' a) (A_B.Map x a) \<and>
Da.cones_map fa (at a \<chi>) = at a \<chi>'"
using \<chi>a.is_universal by simp
qed
text\<open>
Our objective is to show the existence of a unique arrow \<open>f\<close> that transforms
\<open>\<chi>\<close> into \<open>\<chi>'\<close>. We obtain \<open>f\<close> by bundling the arrows \<open>fa\<close>
of \<open>C\<close> and proving that this yields a natural transformation from \<open>X\<close>
to \<open>C\<close>, hence an arrow of \<open>[X, C]\<close>.
\<close>
show "\<exists>!f. \<guillemotleft>f : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f \<chi> = \<chi>'"
proof
let ?P = "\<lambda>a fa. \<guillemotleft>fa : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright> \<and>
diagram.cones_map J B (at a D) fa (at a \<chi>) = at a \<chi>'"
have AaPa: "\<And>a. A.ide a \<Longrightarrow> ?P a (THE fa. ?P a fa)"
proof -
fix a
assume a: "A.ide a"
have "\<exists>!fa. ?P a fa" using a EU by simp
thus "?P a (THE fa. ?P a fa)" using a theI' [of "?P a"] by fastforce
qed
have AaPa_in_hom:
"\<And>a. A.ide a \<Longrightarrow> \<guillemotleft>THE fa. ?P a fa : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>"
using AaPa by blast
have AaPa_map:
"\<And>a. A.ide a \<Longrightarrow>
diagram.cones_map J B (at a D) (THE fa. ?P a fa) (at a \<chi>) = at a \<chi>'"
using AaPa by blast
let ?Fun_f = "\<lambda>a. if A.ide a then (THE fa. ?P a fa) else B.null"
interpret Fun_x: "functor" A B \<open>\<lambda>a. A_B.Map x a\<close>
using x A_B.ide_char by simp
interpret Fun_x': "functor" A B \<open>\<lambda>a. A_B.Map x' a\<close>
using x' A_B.ide_char by simp
text\<open>
The arrows \<open>Fun_f a\<close> are the components of a natural transformation.
It is more work to verify the naturality than it seems like it ought to be.
\<close>
interpret \<phi>: transformation_by_components A B
\<open>\<lambda>a. A_B.Map x' a\<close> \<open>\<lambda>a. A_B.Map x a\<close> ?Fun_f
proof
fix a
assume a: "A.ide a"
show "\<guillemotleft>?Fun_f a : A_B.Map x' a \<rightarrow>\<^sub>B A_B.Map x a\<guillemotright>" using a AaPa by simp
next
fix a
assume a: "A.arr a"
text\<open>
\newcommand\xdom{\mathop{\rm dom}}
\newcommand\xcod{\mathop{\rm cod}}
$$\xymatrix{
{x_{\xdom a}} \drtwocell\omit{\omit(A)} \ar[d]_{\chi_{\xdom a}} \ar[r]^{x_a} & {x_{\xcod a}}
\ar[d]^{\chi_{\xcod a}} \\
{D_{\xdom a}} \ar[r]^{D_a} & {D_{\xcod a}} \\
{x'_{\xdom a}} \urtwocell\omit{\omit(B)} \ar@/^5em/[uu]^{f_{\xdom a}}_{\hspace{1em}(C)} \ar[u]^{\chi'_{\xdom a}}
\ar[r]_{x'_a} & {x'_{\xcod a}} \ar[u]_{x'_{\xcod a}} \ar@/_5em/[uu]_{f_{\xcod a}}
}$$
\<close>
let ?x_dom_a = "A_B.Map x (A.dom a)"
let ?x_cod_a = "A_B.Map x (A.cod a)"
let ?x_a = "A_B.Map x a"
have x_a: "\<guillemotleft>?x_a : ?x_dom_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright>"
using a x A_B.ide_char by auto
let ?x'_dom_a = "A_B.Map x' (A.dom a)"
let ?x'_cod_a = "A_B.Map x' (A.cod a)"
let ?x'_a = "A_B.Map x' a"
have x'_a: "\<guillemotleft>?x'_a : ?x'_dom_a \<rightarrow>\<^sub>B ?x'_cod_a\<guillemotright>"
using a x' A_B.ide_char by auto
let ?f_dom_a = "?Fun_f (A.dom a)"
let ?f_cod_a = "?Fun_f (A.cod a)"
have f_dom_a: "\<guillemotleft>?f_dom_a : ?x'_dom_a \<rightarrow>\<^sub>B ?x_dom_a\<guillemotright>" using a AaPa by simp
have f_cod_a: "\<guillemotleft>?f_cod_a : ?x'_cod_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright>" using a AaPa by simp
interpret D_dom_a: diagram J B \<open>at (A.dom a) D\<close> using a at_ide_is_diagram by simp
interpret D_cod_a: diagram J B \<open>at (A.cod a) D\<close> using a at_ide_is_diagram by simp
interpret Da: natural_transformation J B \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close> \<open>at a D\<close>
using a functor_axioms functor_at_arr_is_transformation by simp
interpret \<chi>_dom_a: limit_cone J B \<open>at (A.dom a) D\<close> \<open>A_B.Map x (A.dom a)\<close>
\<open>at (A.dom a) \<chi>\<close>
using assms(2) a by auto
interpret \<chi>_cod_a: limit_cone J B \<open>at (A.cod a) D\<close> \<open>A_B.Map x (A.cod a)\<close>
\<open>at (A.cod a) \<chi>\<close>
using assms(2) a by auto
interpret \<chi>'_dom_a: cone J B \<open>at (A.dom a) D\<close> \<open>A_B.Map x' (A.dom a)\<close>
\<open>at (A.dom a) \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
interpret \<chi>'_cod_a: cone J B \<open>at (A.cod a) D\<close> \<open>A_B.Map x' (A.cod a)\<close>
\<open>at (A.cod a) \<chi>'\<close>
using a cone_\<chi>' cone_at_ide_is_cone by auto
text\<open>
Now construct cones with apexes \<open>x_dom_a\<close> and \<open>x'_dom_a\<close>
over @{term "at (A.cod a) D"} by forming the vertical composites of
@{term "at (A.dom a) \<chi>"} and @{term "at (A.cod a) \<chi>'"} with the natural
transformation @{term "at a D"}.
\<close>
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close>
\<open>at (A.dom a) \<chi>\<close> \<open>at a D\<close> ..
interpret Dao\<chi>_dom_a: cone J B \<open>at (A.cod a) D\<close> ?x_dom_a Dao\<chi>_dom_a.map
using \<chi>_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone
by metis
interpret Dao\<chi>'_dom_a: vertical_composite J B
\<chi>'_dom_a.A.map \<open>at (A.dom a) D\<close> \<open>at (A.cod a) D\<close>
\<open>at (A.dom a) \<chi>'\<close> \<open>at a D\<close> ..
interpret Dao\<chi>'_dom_a: cone J B \<open>at (A.cod a) D\<close> ?x'_dom_a Dao\<chi>'_dom_a.map
using \<chi>'_dom_a.cone_axioms Da.natural_transformation_axioms vcomp_transformation_cone
by metis
have Dao\<chi>_dom_a: "D_cod_a.cone ?x_dom_a Dao\<chi>_dom_a.map" ..
have Dao\<chi>'_dom_a: "D_cod_a.cone ?x'_dom_a Dao\<chi>'_dom_a.map" ..
text\<open>
These cones are also obtained by transforming the cones @{term "at (A.cod a) \<chi>"}
and @{term "at (A.cod a) \<chi>'"} by \<open>x_a\<close> and \<open>x'_a\<close>, respectively.
\<close>
have A: "Dao\<chi>_dom_a.map = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>)"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
using Dao\<chi>_dom_a.is_extensional \<chi>_cod_a.cone_axioms x_a by force
moreover have
"J.arr j \<Longrightarrow> Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
proof -
assume j: "J.arr j"
have "Dao\<chi>_dom_a.map j = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> (J.dom j)"
using j Dao\<chi>_dom_a.map_simp_2 by simp
also have "... = A_B.Map (D j) a \<cdot>\<^sub>B A_B.Map (\<chi> (J.dom j)) (A.dom a)"
using a j at_simp by simp
also have "... = A_B.Map (A_B.comp (D j) (\<chi> (J.dom j))) a"
using a j A_B.Map_comp
by (metis (no_types, lifting) A.comp_arr_dom \<chi>.is_natural_1
\<chi>.preserves_reflects_arr)
also have "... = A_B.Map (A_B.comp (\<chi> (J.cod j)) (\<chi>.A.map j)) a"
using a j \<chi>.naturality by simp
also have "... = A_B.Map (\<chi> (J.cod j)) (A.cod a) \<cdot>\<^sub>B A_B.Map x a"
using a j x A_B.Map_comp
by (metis (no_types, lifting) A.comp_cod_arr \<chi>.A.map_simp \<chi>.is_natural_2
\<chi>.preserves_reflects_arr)
also have "... = at (A.cod a) \<chi> (J.cod j) \<cdot>\<^sub>B A_B.Map x a"
using a j at_simp by simp
also have "... = at (A.cod a) \<chi> j \<cdot>\<^sub>B A_B.Map x a"
using a j \<chi>_cod_a.is_natural_2 \<chi>_cod_a.A.map_simp
by (metis J.arr_cod_iff_arr J.cod_cod)
also have "... = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
using a j x \<chi>_cod_a.cone_axioms preserves_cod by simp
finally show ?thesis by blast
qed
ultimately show "Dao\<chi>_dom_a.map j = D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>) j"
by blast
qed
have B: "Dao\<chi>'_dom_a.map = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>')"
proof
fix j
have "\<not>J.arr j \<Longrightarrow>
Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
using Dao\<chi>'_dom_a.is_extensional \<chi>'_cod_a.cone_axioms x'_a by force
moreover have
"J.arr j \<Longrightarrow> Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
proof -
assume j: "J.arr j"
have "Dao\<chi>'_dom_a.map j = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi>' (J.dom j)"
using j Dao\<chi>'_dom_a.map_simp_2 by simp
also have "... = A_B.Map (D j) a \<cdot>\<^sub>B A_B.Map (\<chi>' (J.dom j)) (A.dom a)"
using a j at_simp by simp
also have "... = A_B.Map (A_B.comp (D j) (\<chi>' (J.dom j))) a"
using a j A_B.Map_comp
by (metis (no_types, lifting) A.comp_arr_dom \<chi>'.is_natural_1
\<chi>'.preserves_reflects_arr)
also have "... = A_B.Map (A_B.comp (\<chi>' (J.cod j)) (\<chi>'.A.map j)) a"
using a j \<chi>'.naturality by simp
also have "... = A_B.Map (\<chi>' (J.cod j)) (A.cod a) \<cdot>\<^sub>B A_B.Map x' a"
using a j x' A_B.Map_comp
by (metis (no_types, lifting) A.comp_cod_arr \<chi>'.A.map_simp \<chi>'.is_natural_2
\<chi>'.preserves_reflects_arr)
also have "... = at (A.cod a) \<chi>' (J.cod j) \<cdot>\<^sub>B A_B.Map x' a"
using a j at_simp by simp
also have "... = at (A.cod a) \<chi>' j \<cdot>\<^sub>B A_B.Map x' a"
using a j \<chi>'_cod_a.is_natural_2 \<chi>'_cod_a.A.map_simp
by (metis J.arr_cod_iff_arr J.cod_cod)
also have "... = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
using a j x' \<chi>'_cod_a.cone_axioms preserves_cod by simp
finally show ?thesis by blast
qed
ultimately show
"Dao\<chi>'_dom_a.map j = D_cod_a.cones_map ?x'_a (at (A.cod a) \<chi>') j"
by blast
qed
text\<open>
Next, we show that \<open>f_dom_a\<close>, which is the unique arrow that transforms
\<open>\<chi>_dom_a\<close> into \<open>\<chi>'_dom_a\<close>, is also the unique arrow that transforms
\<open>Dao\<chi>_dom_a\<close> into \<open>Dao\<chi>'_dom_a\<close>.
\<close>
have C: "D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map = Dao\<chi>'_dom_a.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation
J B \<chi>'_dom_a.A.map (at (A.cod a) D) Dao\<chi>'_dom_a.map" ..
show "natural_transformation J B \<chi>'_dom_a.A.map (at (A.cod a) D)
(D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map)"
proof -
interpret \<kappa>: cone J B \<open>at (A.cod a) D\<close> ?x'_dom_a
\<open>D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map\<close>
proof -
have "\<And>b b' f. \<lbrakk> f \<in> B.hom b' b; D_cod_a.cone b Dao\<chi>_dom_a.map \<rbrakk>
\<Longrightarrow> D_cod_a.cone b' (D_cod_a.cones_map f Dao\<chi>_dom_a.map)"
using D_cod_a.cones_map_mapsto by blast
moreover have "D_cod_a.cone ?x_dom_a Dao\<chi>_dom_a.map" ..
ultimately show "D_cod_a.cone ?x'_dom_a
(D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map)"
using f_dom_a by simp
qed
show ?thesis ..
qed
show "\<And>j. J.ide j \<Longrightarrow>
D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j = Dao\<chi>'_dom_a.map j"
proof -
fix j
assume j: "J.ide j"
have "D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j =
Dao\<chi>_dom_a.map j \<cdot>\<^sub>B ?f_dom_a"
using j f_dom_a Dao\<chi>_dom_a.cone_axioms
by (elim B.in_homE, auto)
also have "... = (at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> j) \<cdot>\<^sub>B ?f_dom_a"
using j Dao\<chi>_dom_a.map_simp_ide by simp
also have "... = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi> j \<cdot>\<^sub>B ?f_dom_a"
using B.comp_assoc by simp
also have "... = at a D j \<cdot>\<^sub>B D_dom_a.cones_map ?f_dom_a (at (A.dom a) \<chi>) j"
using j \<chi>_dom_a.cone_axioms f_dom_a
by (elim B.in_homE, auto)
also have "... = at a D j \<cdot>\<^sub>B at (A.dom a) \<chi>' j"
using a AaPa A.ide_dom by presburger
also have "... = Dao\<chi>'_dom_a.map j"
using j Dao\<chi>'_dom_a.map_simp_ide by simp
finally show
"D_cod_a.cones_map ?f_dom_a Dao\<chi>_dom_a.map j = Dao\<chi>'_dom_a.map j"
by auto
qed
qed
text\<open>
Naturality amounts to showing that \<open>C f_cod_a x'_a = C x_a f_dom_a\<close>.
To do this, we show that both arrows transform @{term "at (A.cod a) \<chi>"}
into \<open>Dao\<chi>'_cod_a\<close>, thus they are equal by the universality of
@{term "at (A.cod a) \<chi>"}.
\<close>
have "\<exists>!fa. \<guillemotleft>fa : ?x'_dom_a \<rightarrow>\<^sub>B ?x_cod_a\<guillemotright> \<and>
D_cod_a.cones_map fa (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
using Dao\<chi>'_dom_a.cone_axioms a \<chi>_cod_a.is_universal [of ?x'_dom_a Dao\<chi>'_dom_a.map]
by fast
moreover have
"?f_cod_a \<cdot>\<^sub>B ?x'_a \<in> B.hom ?x'_dom_a ?x_cod_a \<and>
D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof
show "?f_cod_a \<cdot>\<^sub>B ?x'_a \<in> B.hom ?x'_dom_a ?x_cod_a"
using f_cod_a x'_a by blast
show "D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof -
have "D_cod_a.cones_map (?f_cod_a \<cdot>\<^sub>B ?x'_a) (at (A.cod a) \<chi>)
= restrict (D_cod_a.cones_map ?x'_a o D_cod_a.cones_map ?f_cod_a)
(D_cod_a.cones (?x_cod_a))
(at (A.cod a) \<chi>)"
using x'_a D_cod_a.cones_map_comp [of ?f_cod_a ?x'_a] f_cod_a
by (elim B.in_homE, auto)
also have "... = D_cod_a.cones_map ?x'_a
(D_cod_a.cones_map ?f_cod_a (at (A.cod a) \<chi>))"
using \<chi>_cod_a.cone_axioms by simp
also have "... = Dao\<chi>'_dom_a.map"
using a B AaPa_map A.ide_cod by presburger
finally show ?thesis by auto
qed
qed
moreover have
"?x_a \<cdot>\<^sub>B ?f_dom_a \<in> B.hom ?x'_dom_a ?x_cod_a \<and>
D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof
show "?x_a \<cdot>\<^sub>B ?f_dom_a \<in> B.hom ?x'_dom_a ?x_cod_a"
using f_dom_a x_a by blast
show "D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>) = Dao\<chi>'_dom_a.map"
proof -
have
"D_cod_a.cones (B.cod (A_B.Map x a)) = D_cod_a.cones (A_B.Map x (A.cod a))"
using a x by simp
moreover have "B.seq ?x_a ?f_dom_a"
using f_dom_a x_a by (elim B.in_homE, auto)
ultimately have
"D_cod_a.cones_map (?x_a \<cdot>\<^sub>B ?f_dom_a) (at (A.cod a) \<chi>)
= restrict (D_cod_a.cones_map ?f_dom_a o D_cod_a.cones_map ?x_a)
(D_cod_a.cones (?x_cod_a))
(at (A.cod a) \<chi>)"
using D_cod_a.cones_map_comp [of ?x_a ?f_dom_a] x_a by argo
also have "... = D_cod_a.cones_map ?f_dom_a
(D_cod_a.cones_map ?x_a (at (A.cod a) \<chi>))"
using \<chi>_cod_a.cone_axioms by simp
also have "... = Dao\<chi>'_dom_a.map"
using A C a AaPa by argo
finally show ?thesis by blast
qed
qed
ultimately show "?f_cod_a \<cdot>\<^sub>B ?x'_a = ?x_a \<cdot>\<^sub>B ?f_dom_a"
using a \<chi>_cod_a.is_universal by blast
qed
text\<open>
The arrow from @{term x'} to @{term x} in \<open>[A, B]\<close> determined by
the natural transformation \<open>\<phi>\<close> transforms @{term \<chi>} into @{term \<chi>'}.
Moreover, it is the unique such arrow, since the components of \<open>\<phi>\<close>
are each determined by universality.
\<close>
let ?f = "A_B.MkArr (\<lambda>a. A_B.Map x' a) (\<lambda>a. A_B.Map x a) \<phi>.map"
have f_in_hom: "?f \<in> A_B.hom x' x"
proof -
have arr_f: "A_B.arr ?f"
using x' x A_B.arr_MkArr \<phi>.natural_transformation_axioms by simp
moreover have "A_B.MkIde (\<lambda>a. A_B.Map x a) = x"
using x A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis
moreover have "A_B.MkIde (\<lambda>a. A_B.Map x' a) = x'"
using x' A_B.ide_char A_B.MkArr_Map A_B.in_homE A_B.ide_in_hom by metis
ultimately show ?thesis
using A_B.dom_char A_B.cod_char by auto
qed
have Fun_f: "\<And>a. A.ide a \<Longrightarrow> A_B.Map ?f a = (THE fa. ?P a fa)"
using f_in_hom \<phi>.map_simp_ide by fastforce
have cones_map_f: "cones_map ?f \<chi> = \<chi>'"
using AaPa Fun_f at_ide_is_diagram assms(2) x x' cone_\<chi> cone_\<chi>' f_in_hom Fun_f
cones_map_pointwise
by presburger
show "\<guillemotleft>?f : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map ?f \<chi> = \<chi>'" using f_in_hom cones_map_f by auto
show "\<And>f'. \<guillemotleft>f' : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f' \<chi> = \<chi>' \<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : x' \<rightarrow>\<^sub>[\<^sub>A\<^sub>,\<^sub>B\<^sub>] x\<guillemotright> \<and> cones_map f' \<chi> = \<chi>'"
have 0: "\<And>a. A.ide a \<Longrightarrow>
diagram.cones_map J B (at a D) (A_B.Map f' a) (at a \<chi>) = at a \<chi>'"
using f' cone_\<chi> cone_\<chi>' cones_map_pointwise by blast
have "f' = A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f')"
using f' A_B.MkArr_Map by auto
also have "... = ?f"
proof (intro A_B.MkArr_eqI)
show "A_B.arr (A_B.MkArr (A_B.Dom f') (A_B.Cod f') (A_B.Map f'))"
using f' calculation by blast
show 1: "A_B.Dom f' = A_B.Map x'" using f' A_B.Map_dom by auto
show 2: "A_B.Cod f' = A_B.Map x" using f' A_B.Map_cod by auto
show "A_B.Map f' = \<phi>.map"
proof (intro NaturalTransformation.eqI)
show "natural_transformation A B (A_B.Map x') (A_B.Map x) \<phi>.map" ..
show "natural_transformation A B (A_B.Map x') (A_B.Map x) (A_B.Map f')"
using f' 1 2 A_B.arr_char [of f'] by auto
show "\<And>a. A.ide a \<Longrightarrow> A_B.Map f' a = \<phi>.map a"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>at a D\<close> using a at_ide_is_diagram by auto
interpret Fun_f': natural_transformation A B \<open>A_B.Dom f'\<close> \<open>A_B.Cod f'\<close>
\<open>A_B.Map f'\<close>
using f' A_B.arr_char by fast
have "A_B.Map f' a \<in> B.hom (A_B.Map x' a) (A_B.Map x a)"
using a f' Fun_f'.preserves_hom A.ide_in_hom by auto
hence "?P a (A_B.Map f' a)" using a 0 [of a] by simp
moreover have "?P a (\<phi>.map a)"
using a \<phi>.map_simp_ide Fun_f AaPa by presburger
ultimately show "A_B.Map f' a = \<phi>.map a" using a EU by blast
qed
qed
qed
finally show "f' = ?f" by auto
qed
qed
qed
qed
end
context functor_category
begin
text\<open>
A functor category \<open>[A, B]\<close> has limits of shape @{term[source=true] J}
whenever @{term B} has limits of shape @{term[source=true] J}.
\<close>
lemma has_limits_of_shape_if_target_does:
assumes "category (J :: 'j comp)"
and "B.has_limits_of_shape J"
shows "has_limits_of_shape J"
proof (unfold has_limits_of_shape_def)
have "\<And>D. diagram J comp D \<Longrightarrow> (\<exists>x \<chi>. limit_cone J comp D x \<chi>)"
proof -
fix D
assume D: "diagram J comp D"
interpret J: category J using assms(1) by auto
interpret JxA: product_category J A ..
interpret D: diagram J comp D using D by auto
interpret D: diagram_in_functor_category A B J D ..
interpret Curry: currying J A B ..
text\<open>
Given diagram @{term D} in \<open>[A, B]\<close>, choose for each object \<open>a\<close>
of \<open>A\<close> a limit cone \<open>(la, \<chi>a)\<close> for \<open>at a D\<close> in \<open>B\<close>.
\<close>
let ?l = "\<lambda>a. diagram.some_limit J B (D.at a D)"
let ?\<chi> = "\<lambda>a. diagram.some_limit_cone J B (D.at a D)"
have l\<chi>: "\<And>a. A.ide a \<Longrightarrow> diagram.limit_cone J B (D.at a D) (?l a) (?\<chi> a)"
proof -
fix a
assume a: "A.ide a"
interpret Da: diagram J B \<open>D.at a D\<close>
using a D.at_ide_is_diagram by blast
show "limit_cone J B (D.at a D) (?l a) (?\<chi> a)"
using assms(2) B.has_limits_of_shape_def Da.diagram_axioms
Da.limit_cone_some_limit_cone
by auto
qed
text\<open>
The choice of limit cones induces a limit functor from \<open>A\<close> to \<open>B\<close>.
\<close>
interpret uncurry_D: diagram JxA.comp B "Curry.uncurry D"
proof -
interpret "functor" JxA.comp B \<open>Curry.uncurry D\<close>
using D.functor_axioms Curry.uncurry_preserves_functors by simp
interpret binary_functor J A B \<open>Curry.uncurry D\<close> ..
show "diagram JxA.comp B (Curry.uncurry D)" ..
qed
interpret uncurry_D: parametrized_diagram J A B \<open>Curry.uncurry D\<close> ..
let ?L = "uncurry_D.L ?l ?\<chi>"
let ?P = "uncurry_D.P ?l ?\<chi>"
interpret L: "functor" A B ?L
using l\<chi> uncurry_D.chosen_limits_induce_functor [of ?l ?\<chi>] by simp
have L_ide: "\<And>a. A.ide a \<Longrightarrow> ?L a = ?l a"
using uncurry_D.L_ide [of ?l ?\<chi>] l\<chi> by blast
have L_arr: "\<And>a. A.arr a \<Longrightarrow> (\<exists>!f. ?P a f) \<and> ?P a (?L a)"
using uncurry_D.L_arr [of ?l ?\<chi>] l\<chi> by blast
text\<open>
The functor \<open>L\<close> extends to a functor \<open>L'\<close> from \<open>JxA\<close>
to \<open>B\<close> that is constant on \<open>J\<close>.
\<close>
let ?L' = "\<lambda>ja. if JxA.arr ja then ?L (snd ja) else B.null"
let ?P' = "\<lambda>ja. ?P (snd ja)"
interpret L': "functor" JxA.comp B ?L'
apply unfold_locales
using L.preserves_arr L.preserves_dom L.preserves_cod
apply auto[4]
using L.preserves_comp JxA.comp_char by (elim JxA.seqE, auto)
have "\<And>ja. JxA.arr ja \<Longrightarrow> (\<exists>!f. ?P' ja f) \<and> ?P' ja (?L' ja)"
proof -
fix ja
assume ja: "JxA.arr ja"
have "A.arr (snd ja)" using ja by blast
thus "(\<exists>!f. ?P' ja f) \<and> ?P' ja (?L' ja)"
using ja L_arr by presburger
qed
hence L'_arr: "\<And>ja. JxA.arr ja \<Longrightarrow> ?P' ja (?L' ja)" by blast
have L'_ide: "\<And>ja. \<lbrakk> J.arr (fst ja); A.ide (snd ja) \<rbrakk> \<Longrightarrow> ?L' ja = ?l (snd ja)"
using L_ide l\<chi> by force
have L'_arr_map:
"\<And>ja. JxA.arr ja \<Longrightarrow> uncurry_D.P ?l ?\<chi> (snd ja) (uncurry_D.L ?l ?\<chi> (snd ja))"
using L'_arr by presburger
text\<open>
The map that takes an object \<open>(j, a)\<close> of \<open>JxA\<close> to the component
\<open>\<chi> a j\<close> of the limit cone \<open>\<chi> a\<close> is a natural transformation
from \<open>L\<close> to uncurry \<open>D\<close>.
\<close>
let ?\<chi>' = "\<lambda>ja. ?\<chi> (snd ja) (fst ja)"
interpret \<chi>': transformation_by_components JxA.comp B ?L' \<open>Curry.uncurry D\<close> ?\<chi>'
proof
fix ja
assume ja: "JxA.ide ja"
let ?j = "fst ja"
let ?a = "snd ja"
interpret \<chi>a: limit_cone J B \<open>D.at ?a D\<close> \<open>?l ?a\<close> \<open>?\<chi> ?a\<close>
using ja l\<chi> by blast
show "\<guillemotleft>?\<chi>' ja : ?L' ja \<rightarrow>\<^sub>B Curry.uncurry D ja\<guillemotright>"
using ja L'_ide [of ja] by force
next
fix ja
assume ja: "JxA.arr ja"
let ?j = "fst ja"
let ?a = "snd ja"
have j: "J.arr ?j" using ja by simp
have a: "A.arr ?a" using ja by simp
interpret D_dom_a: diagram J B \<open>D.at (A.dom ?a) D\<close>
using a D.at_ide_is_diagram by auto
interpret D_cod_a: diagram J B \<open>D.at (A.cod ?a) D\<close>
using a D.at_ide_is_diagram by auto
interpret Da: natural_transformation J B
\<open>D.at (A.dom ?a) D\<close> \<open>D.at (A.cod ?a) D\<close> \<open>D.at ?a D\<close>
using a D.functor_axioms D.functor_at_arr_is_transformation by simp
interpret \<chi>_dom_a: limit_cone J B \<open>D.at (A.dom ?a) D\<close> \<open>?l (A.dom ?a)\<close>
\<open>?\<chi> (A.dom ?a)\<close>
using a l\<chi> by simp
interpret \<chi>_cod_a: limit_cone J B \<open>D.at (A.cod ?a) D\<close> \<open>?l (A.cod ?a)\<close>
\<open>?\<chi> (A.cod ?a)\<close>
using a l\<chi> by simp
interpret Dao\<chi>_dom_a: vertical_composite J B
\<chi>_dom_a.A.map \<open>D.at (A.dom ?a) D\<close> \<open>D.at (A.cod ?a) D\<close>
\<open>?\<chi> (A.dom ?a)\<close> \<open>D.at ?a D\<close>
..
interpret Dao\<chi>_dom_a: cone J B \<open>D.at (A.cod ?a) D\<close> \<open>?l (A.dom ?a)\<close> Dao\<chi>_dom_a.map ..
show "?\<chi>' (JxA.cod ja) \<cdot>\<^sub>B ?L' ja = B (Curry.uncurry D ja) (?\<chi>' (JxA.dom ja))"
proof -
have "?\<chi>' (JxA.cod ja) \<cdot>\<^sub>B ?L' ja = ?\<chi> (A.cod ?a) (J.cod ?j) \<cdot>\<^sub>B ?L' ja"
using ja by fastforce
also have "... = D_cod_a.cones_map (?L' ja) (?\<chi> (A.cod ?a)) (J.cod ?j)"
using ja L'_arr_map [of ja] \<chi>_cod_a.cone_axioms by auto
also have "... = Dao\<chi>_dom_a.map (J.cod ?j)"
using ja \<chi>_cod_a.induced_arrowI Dao\<chi>_dom_a.cone_axioms L'_arr by presburger
also have "... = D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D_dom_a.some_limit_cone (J.cod ?j)"
using ja Dao\<chi>_dom_a.map_simp_ide by fastforce
also have "... = D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D.at (A.dom ?a) D ?j \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using ja \<chi>_dom_a.naturality \<chi>_dom_a.ide_apex apply simp
by (metis B.comp_arr_ide \<chi>_dom_a.preserves_reflects_arr)
also have "... = (D.at ?a D (J.cod ?j) \<cdot>\<^sub>B D.at (A.dom ?a) D ?j) \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using j ja B.comp_assoc by presburger
also have "... = B (D.at ?a D ?j) (?\<chi>' (JxA.dom ja))"
using a j ja Map_comp A.comp_arr_dom D.is_natural_2 by simp
also have "... = Curry.uncurry D ja \<cdot>\<^sub>B ?\<chi>' (JxA.dom ja)"
using Curry.uncurry_def by simp
finally show ?thesis by auto
qed
qed
text\<open>
Since \<open>\<chi>'\<close> is constant on \<open>J\<close>, \<open>curry \<chi>'\<close> is a cone over \<open>D\<close>.
\<close>
interpret constL: constant_functor J comp \<open>MkIde ?L\<close>
using L.natural_transformation_axioms MkArr_in_hom ide_in_hom L.functor_axioms
by unfold_locales blast
(* TODO: This seems a little too involved. *)
have curry_L': "constL.map = Curry.curry ?L' ?L' ?L'"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> constL.map j = Curry.curry ?L' ?L' ?L' j"
using Curry.curry_def constL.is_extensional by simp
moreover have "J.arr j \<Longrightarrow> constL.map j = Curry.curry ?L' ?L' ?L' j"
using Curry.curry_def constL.value_is_ide in_homE ide_in_hom by auto
ultimately show "constL.map j = Curry.curry ?L' ?L' ?L' j" by blast
qed
hence uncurry_constL: "Curry.uncurry constL.map = ?L'"
using L'.natural_transformation_axioms Curry.uncurry_curry by simp
interpret curry_\<chi>': natural_transformation J comp constL.map D
\<open>Curry.curry ?L' (Curry.uncurry D) \<chi>'.map\<close>
proof -
have "Curry.curry (Curry.uncurry D) (Curry.uncurry D) (Curry.uncurry D) = D"
using Curry.curry_uncurry D.functor_axioms D.natural_transformation_axioms
by blast
thus "natural_transformation J comp constL.map D
(Curry.curry ?L' (Curry.uncurry D) \<chi>'.map)"
using Curry.curry_preserves_transformations curry_L' \<chi>'.natural_transformation_axioms
by force
qed
interpret curry_\<chi>': cone J comp D \<open>MkIde ?L\<close> \<open>Curry.curry ?L' (Curry.uncurry D) \<chi>'.map\<close>
..
text\<open>
The value of \<open>curry_\<chi>'\<close> at each object \<open>a\<close> of \<open>A\<close> is the
limit cone \<open>\<chi> a\<close>, hence \<open>curry_\<chi>'\<close> is a limit cone.
\<close>
have 1: "\<And>a. A.ide a \<Longrightarrow> D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) = ?\<chi> a"
proof -
fix a
assume a: "A.ide a"
have "D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) =
(\<lambda>j. Curry.uncurry (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) (j, a))"
using a by simp
moreover have "... = (\<lambda>j. \<chi>'.map (j, a))"
using a Curry.uncurry_curry \<chi>'.natural_transformation_axioms by simp
moreover have "... = ?\<chi> a"
proof (intro NaturalTransformation.eqI)
interpret \<chi>a: limit_cone J B \<open>D.at a D\<close> \<open>?l a\<close> \<open>?\<chi> a\<close> using a l\<chi> by simp
interpret \<chi>': binary_functor_transformation J A B ?L' \<open>Curry.uncurry D\<close> \<chi>'.map ..
show "natural_transformation J B \<chi>a.A.map (D.at a D) (?\<chi> a)" ..
show "natural_transformation J B \<chi>a.A.map (D.at a D) (\<lambda>j. \<chi>'.map (j, a))"
proof -
have "\<chi>a.A.map = (\<lambda>j. ?L' (j, a))"
using a \<chi>a.A.map_def L'_ide by auto
thus ?thesis
using a \<chi>'.fixing_ide_gives_natural_transformation_2 by simp
qed
fix j
assume j: "J.ide j"
show "\<chi>'.map (j, a) = ?\<chi> a j"
using a j \<chi>'.map_simp_ide by simp
qed
ultimately show "D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map) = ?\<chi> a" by simp
qed
hence 2: "\<And>a. A.ide a \<Longrightarrow> diagram.limit_cone J B (D.at a D) (?l a)
(D.at a (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map))"
using l\<chi> by simp
hence "limit_cone J comp D (MkIde ?L) (Curry.curry ?L' (Curry.uncurry D) \<chi>'.map)"
using 1 2 L.functor_axioms L_ide curry_\<chi>'.cone_axioms curry_L'
D.cone_is_limit_if_pointwise_limit
by simp
thus "\<exists>x \<chi>. limit_cone J comp D x \<chi>" by blast
qed
thus "\<forall>D. diagram J comp D \<longrightarrow> (\<exists>x \<chi>. limit_cone J comp D x \<chi>)" by blast
qed
lemma has_limits_if_target_does:
assumes "B.has_limits (undefined :: 'j)"
shows "has_limits (undefined :: 'j)"
using assms B.has_limits_def has_limits_def has_limits_of_shape_if_target_does by fast
end
section "The Yoneda Functor Preserves Limits"
text\<open>
In this section, we show that the Yoneda functor from \<open>C\<close> to \<open>[Cop, S]\<close>
preserves limits.
\<close>
context yoneda_functor
begin
lemma preserves_limits:
fixes J :: "'j comp"
assumes "diagram J C D" and "diagram.has_as_limit J C D a"
shows "diagram.has_as_limit J Cop_S.comp (map o D) (map a)"
proof -
text\<open>
The basic idea of the proof is as follows:
If \<open>\<chi>\<close> is a limit cone in \<open>C\<close>, then for every object \<open>a'\<close>
of \<open>Cop\<close> the evaluation of \<open>Y o \<chi>\<close> at \<open>a'\<close> is a limit cone
in \<open>S\<close>. By the results on limits in functor categories,
this implies that \<open>Y o \<chi>\<close> is a limit cone in \<open>[Cop, S]\<close>.
\<close>
interpret J: category J using assms(1) diagram_def by auto
interpret D: diagram J C D using assms(1) by auto
from assms(2) obtain \<chi> where \<chi>: "D.limit_cone a \<chi>" by blast
interpret \<chi>: limit_cone J C D a \<chi> using \<chi> by auto
have a: "C.ide a" using \<chi>.ide_apex by auto
interpret YoD: diagram J Cop_S.comp \<open>map o D\<close>
using D.diagram_axioms functor_axioms preserves_diagrams [of J D] by simp
interpret YoD: diagram_in_functor_category Cop.comp S J \<open>map o D\<close> ..
interpret Yo\<chi>: cone J Cop_S.comp \<open>map o D\<close> \<open>map a\<close> \<open>map o \<chi>\<close>
using \<chi>.cone_axioms preserves_cones by blast
have "\<And>a'. C.ide a' \<Longrightarrow>
limit_cone J S (YoD.at a' (map o D))
(Cop_S.Map (map a) a') (YoD.at a' (map o \<chi>))"
proof -
fix a'
assume a': "C.ide a'"
interpret A': constant_functor J C a'
using a' by (unfold_locales, auto)
interpret YoD_a': diagram J S \<open>YoD.at a' (map o D)\<close>
using a' YoD.at_ide_is_diagram by simp
interpret Yo\<chi>_a': cone J S \<open>YoD.at a' (map o D)\<close>
\<open>Cop_S.Map (map a) a'\<close> \<open>YoD.at a' (map o \<chi>)\<close>
using a' YoD.cone_at_ide_is_cone Yo\<chi>.cone_axioms by fastforce
have eval_at_ide: "\<And>j. J.ide j \<Longrightarrow> YoD.at a' (map \<circ> D) j = Hom.map (a', D j)"
proof -
fix j
assume j: "J.ide j"
have "YoD.at a' (map \<circ> D) j = Cop_S.Map (map (D j)) a'"
using a' j YoD.at_simp YoD.preserves_arr [of j] by auto
also have "... = Y (D j) a'" using Y_def by simp
also have "... = Hom.map (a', D j)" using a' j D.preserves_arr by simp
finally show "YoD.at a' (map \<circ> D) j = Hom.map (a', D j)" by auto
qed
have eval_at_arr: "\<And>j. J.arr j \<Longrightarrow> YoD.at a' (map \<circ> \<chi>) j = Hom.map (a', \<chi> j)"
proof -
fix j
assume j: "J.arr j"
have "YoD.at a' (map \<circ> \<chi>) j = Cop_S.Map ((map o \<chi>) j) a'"
using a' j YoD.at_simp [of a' j "map o \<chi>"] preserves_arr by fastforce
also have "... = Y (\<chi> j) a'" using Y_def by simp
also have "... = Hom.map (a', \<chi> j)" using a' j by simp
finally show "YoD.at a' (map \<circ> \<chi>) j = Hom.map (a', \<chi> j)" by auto
qed
have Fun_map_a_a': "Cop_S.Map (map a) a' = Hom.map (a', a)"
using a a' map_simp preserves_arr [of a] by simp
show "limit_cone J S (YoD.at a' (map o D))
(Cop_S.Map (map a) a') (YoD.at a' (map o \<chi>))"
proof
fix x \<sigma>
assume \<sigma>: "YoD_a'.cone x \<sigma>"
interpret \<sigma>: cone J S \<open>YoD.at a' (map o D)\<close> x \<sigma> using \<sigma> by auto
have x: "S.ide x" using \<sigma>.ide_apex by simp
text\<open>
For each object \<open>j\<close> of \<open>J\<close>, the component \<open>\<sigma> j\<close>
is an arrow in \<open>S.hom x (Hom.map (a', D j))\<close>.
Each element \<open>e \<in> S.set x\<close> therefore determines an arrow
\<open>\<psi> (a', D j) (S.Fun (\<sigma> j) e) \<in> C.hom a' (D j)\<close>.
These arrows are the components of a cone \<open>\<kappa> e\<close> over @{term D}
with apex @{term a'}.
\<close>
have \<sigma>j: "\<And>j. J.ide j \<Longrightarrow> \<guillemotleft>\<sigma> j : x \<rightarrow>\<^sub>S Hom.map (a', D j)\<guillemotright>"
using eval_at_ide \<sigma>.preserves_hom J.ide_in_hom by force
have \<kappa>: "\<And>e. e \<in> S.set x \<Longrightarrow>
transformation_by_components
J C A'.map D (\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
proof -
fix e
assume e: "e \<in> S.set x"
show "transformation_by_components J C A'.map D (\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
proof
fix j
assume j: "J.ide j"
show "\<guillemotleft>\<psi> (a', D j) (S.Fun (\<sigma> j) e) : A'.map j \<rightarrow> D j\<guillemotright>"
using e j S.Fun_mapsto [of "\<sigma> j"] A'.preserves_ide Hom.set_map eval_at_ide
Hom.\<psi>_mapsto [of "A'.map j" "D j"]
by force
next
fix j
assume j: "J.arr j"
show "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> A'.map j =
D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> A'.map j =
\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<cdot> a'"
using A'.map_simp j by simp
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (\<sigma> (J.cod j)) e) \<in> C.hom a' (D (J.cod j))"
using a' e j Hom.\<psi>_mapsto [of "A'.map j" "D (J.cod j)"] A'.map_simp
S.Fun_mapsto [of "\<sigma> (J.cod j)"] Hom.set_map eval_at_ide
by auto
thus ?thesis
using C.comp_arr_dom by fastforce
qed
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e))"
proof -
have "S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) =
(S.Fun (Y (D j) a') o S.Fun (\<sigma> (J.dom j))) e"
by simp
also have "... = S.Fun (Y (D j) a' \<cdot>\<^sub>S \<sigma> (J.dom j)) e"
using a' e j Y_arr_ide(1) S.in_homE \<sigma>j eval_at_ide S.Fun_comp by force
also have "... = S.Fun (\<sigma> (J.cod j)) e"
using a' j x \<sigma>.is_natural_2 \<sigma>.A.map_simp S.comp_arr_dom J.arr_cod_iff_arr
J.cod_cod YoD.preserves_arr \<sigma>.is_natural_1 YoD.at_simp
by auto
finally have
"S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) = S.Fun (\<sigma> (J.cod j)) e"
by auto
thus ?thesis by simp
qed
also have "... = D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)"
proof -
have "S.Fun (Y (D j) a') (S.Fun (\<sigma> (J.dom j)) e) =
\<phi> (a', D (J.cod j)) (D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e))"
proof -
have "S.Fun (\<sigma> (J.dom j)) e \<in> Hom.set (a', D (J.dom j))"
using a' e j \<sigma>j S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map
YoD.at_simp eval_at_ide
by auto
moreover have "C.arr (\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)) \<and>
C.dom (\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)) = a'"
using a' e j \<sigma>j S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map eval_at_ide
Hom.\<psi>_mapsto [of a' "D (J.dom j)"]
by auto
ultimately show ?thesis
using a' e j Hom.Fun_map C.comp_arr_dom by force
qed
moreover have "D j \<cdot> \<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e)
\<in> C.hom a' (D (J.cod j))"
proof -
have "\<psi> (a', D (J.dom j)) (S.Fun (\<sigma> (J.dom j)) e) \<in> C.hom a' (D (J.dom j))"
using a' e j Hom.\<psi>_mapsto [of a' "D (J.dom j)"] eval_at_ide
S.Fun_mapsto [of "\<sigma> (J.dom j)"] Hom.set_map
by auto
thus ?thesis using j D.preserves_hom by blast
qed
ultimately show ?thesis using a' j Hom.\<psi>_\<phi> by simp
qed
finally show ?thesis by auto
qed
qed
qed
let ?\<kappa> = "\<lambda>e. transformation_by_components.map J C A'.map
(\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e))"
have cone_\<kappa>e: "\<And>e. e \<in> S.set x \<Longrightarrow> D.cone a' (?\<kappa> e)"
proof -
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: transformation_by_components J C A'.map D
\<open>\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e)\<close>
using e \<kappa> by blast
show "D.cone a' (?\<kappa> e)" ..
qed
text\<open>
Since \<open>\<kappa> e\<close> is a cone for each element \<open>e\<close> of \<open>S.set x\<close>,
by the universal property of the limit cone \<open>\<chi>\<close> there is a unique arrow
\<open>fe \<in> C.hom a' a\<close> that transforms \<open>\<chi>\<close> to \<open>\<kappa> e\<close>.
\<close>
have ex_fe: "\<And>e. e \<in> S.set x \<Longrightarrow> \<exists>!fe. \<guillemotleft>fe : a' \<rightarrow> a\<guillemotright> \<and> D.cones_map fe \<chi> = ?\<kappa> e"
using cone_\<kappa>e \<chi>.is_universal by simp
text\<open>
The map taking \<open>e \<in> S.set x\<close> to \<open>fe \<in> C.hom a' a\<close>
determines an arrow \<open>f \<in> S.hom x (Hom (a', a))\<close> that
transforms the cone obtained by evaluating \<open>Y o \<chi>\<close> at \<open>a'\<close>
to the cone \<open>\<sigma>\<close>.
\<close>
let ?f = "S.mkArr (S.set x) (Hom.set (a', a))
(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)))"
have 0: "(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e))) \<in> S.set x \<rightarrow> Hom.set (a', a)"
proof
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
thus "\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)) \<in> Hom.set (a', a)"
using a a' Hom.\<phi>_mapsto by auto
qed
have f: "\<guillemotleft>?f : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright>"
proof -
have "(\<lambda>e. \<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e))) \<in> S.set x \<rightarrow> Hom.set (a', a)"
proof
fix e
assume e: "e \<in> S.set x"
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
thus "\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)) \<in> Hom.set (a', a)"
using a a' Hom.\<phi>_mapsto by auto
qed
thus ?thesis
using a a' x \<sigma>.ide_apex S.mkArr_in_hom [of "S.set x" "Hom.set (a', a)"]
Hom.set_subset_Univ S.mkIde_set
by simp
qed
have "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) = \<sigma>"
proof (intro NaturalTransformation.eqI)
show "natural_transformation J S \<sigma>.A.map (YoD.at a' (map o D)) \<sigma>"
using \<sigma>.natural_transformation_axioms by auto
have 1: "S.cod ?f = Cop_S.Map (map a) a'"
using f Fun_map_a_a' by force
interpret YoD_a'of: cone J S \<open>YoD.at a' (map o D)\<close> x
\<open>YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>))\<close>
proof -
have "YoD_a'.cone (S.cod ?f) (YoD.at a' (map o \<chi>))"
using a a' f Yo\<chi>_a'.cone_axioms preserves_arr [of a] by auto
hence "YoD_a'.cone (S.dom ?f) (YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))"
using f YoD_a'.cones_map_mapsto S.arrI by blast
thus "cone J S (YoD.at a' (map o D)) x
(YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))"
using f by auto
qed
show "natural_transformation J S \<sigma>.A.map (YoD.at a' (map o D))
(YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)))" ..
fix j
assume j: "J.ide j"
have "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j = YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f"
using f j Fun_map_a_a' Yo\<chi>_a'.cone_axioms by fastforce
also have "... = \<sigma> j"
proof (intro S.arr_eqI)
show "S.par (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) (\<sigma> j)"
using 1 f j x YoD_a'.preserves_hom by fastforce
show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) = S.Fun (\<sigma> j)"
proof
fix e
have "e \<notin> S.set x \<Longrightarrow> S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
using 1 f j x S.Fun_mapsto [of "\<sigma> j"] \<sigma>.A.map_simp
extensional_arb [of "S.Fun (\<sigma> j)"]
by auto
moreover have "e \<in> S.set x \<Longrightarrow>
S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
proof -
assume e: "e \<in> S.set x"
interpret \<kappa>e: transformation_by_components J C A'.map D
\<open>\<lambda>j. \<psi> (a', D j) (S.Fun (\<sigma> j) e)\<close>
using e \<kappa> by blast
interpret \<kappa>e: cone J C D a' \<open>?\<kappa> e\<close> using e cone_\<kappa>e by simp
have induced_arrow: "\<chi>.induced_arrow a' (?\<kappa> e) \<in> C.hom a' a"
using a a' e ex_fe \<chi>.induced_arrowI \<kappa>e.cone_axioms by simp
have "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e =
restrict (S.Fun (YoD.at a' (map o \<chi>) j) o S.Fun ?f) (S.set x) e"
using 1 e f j S.Fun_comp YoD_a'.preserves_hom by force
also have "... = (\<phi> (a', D j) o C (\<chi> j) o \<psi> (a', a)) (S.Fun ?f e)"
using j a' f e Hom.map_simp_2 S.Fun_mkArr Hom.preserves_arr [of "(a', \<chi> j)"]
eval_at_arr S.arr_mkArr
by (elim S.in_homE, auto)
also have "... = (\<phi> (a', D j) o C (\<chi> j) o \<psi> (a', a))
(\<phi> (a', a) (\<chi>.induced_arrow a' (?\<kappa> e)))"
using e f S.Fun_mkArr by fastforce
also have "... = \<phi> (a', D j) (D.cones_map (\<chi>.induced_arrow a' (?\<kappa> e)) \<chi> j)"
using a a' e j 0 Hom.\<psi>_\<phi> induced_arrow \<chi>.cone_axioms by auto
also have "... = \<phi> (a', D j) (?\<kappa> e j)"
using \<chi>.induced_arrowI \<kappa>e.cone_axioms by fastforce
also have "... = \<phi> (a', D j) (\<psi> (a', D j) (S.Fun (\<sigma> j) e))"
using j \<kappa>e.map_def [of j] by simp
also have "... = S.Fun (\<sigma> j) e"
proof -
have "S.Fun (\<sigma> j) e \<in> Hom.set (a', D j)"
using a' e j S.Fun_mapsto [of "\<sigma> j"] eval_at_ide Hom.set_map by auto
thus ?thesis
using a' j Hom.\<phi>_\<psi> C.ide_in_hom J.ide_in_hom by blast
qed
finally show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
by auto
qed
ultimately show "S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e = S.Fun (\<sigma> j) e"
by auto
qed
qed
finally show "YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j = \<sigma> j" by auto
qed
hence ff: "?f \<in> S.hom x (Hom.map (a', a)) \<and>
YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) = \<sigma>"
using f by auto
text\<open>
Any other arrow \<open>f' \<in> S.hom x (Hom.map (a', a))\<close> that
transforms the cone obtained by evaluating \<open>Y o \<chi>\<close> at @{term a'}
to the cone @{term \<sigma>}, must equal \<open>f\<close>, showing that \<open>f\<close>
is unique.
\<close>
moreover have "\<And>f'. \<guillemotleft>f' : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f' (YoD.at a' (map o \<chi>)) = \<sigma>
\<Longrightarrow> f' = ?f"
proof -
fix f'
assume f': "\<guillemotleft>f' : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f' (YoD.at a' (map o \<chi>)) = \<sigma>"
show "f' = ?f"
proof (intro S.arr_eqI)
show par: "S.par f' ?f" using f f' by (elim S.in_homE, auto)
show "S.Fun f' = S.Fun ?f"
proof
fix e
have "e \<notin> S.set x \<Longrightarrow> S.Fun f' e = S.Fun ?f e"
using f f' x S.Fun_mapsto extensional_arb S.Fun_in_terms_of_comp
by fastforce
moreover have "e \<in> S.set x \<Longrightarrow> S.Fun f' e = S.Fun ?f e"
proof -
assume e: "e \<in> S.set x"
have fe: "S.Fun ?f e \<in> Hom.set (a', a)"
using e f S.arr_mkArr par by auto
have f'e: "S.Fun f' e \<in> Hom.set (a', a)"
using a a' e f' S.Fun_mapsto Hom.set_map by fastforce
have 1: "\<guillemotleft>\<psi> (a', a) (S.Fun f' e) : a' \<rightarrow> a\<guillemotright>"
using a a' e f' f'e S.Fun_mapsto Hom.\<psi>_mapsto Hom.set_map by blast
have 2: "\<guillemotleft>\<psi> (a', a) (S.Fun ?f e) : a' \<rightarrow> a\<guillemotright>"
using a a' e f' fe S.Fun_mapsto Hom.\<psi>_mapsto Hom.set_map by blast
interpret \<chi>ofe: cone J C D a' \<open>D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>\<close>
proof -
have "D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<in> D.cones a \<rightarrow> D.cones a'"
using 2 D.cones_map_mapsto [of "\<psi> (a', a) (S.Fun ?f e)"]
by (elim C.in_homE, auto)
thus "cone J C D a' (D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>)"
using \<chi>.cone_axioms by blast
qed
have A: "\<And>h j. h \<in> C.hom a' a \<Longrightarrow> J.arr j \<Longrightarrow>
S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
proof -
fix h j
assume j: "J.arr j"
assume h: "h \<in> C.hom a' a"
have "S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a)) (\<phi> (a', a) h)"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j)
= restrict (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a))
(Hom.set (a', a))"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) = S.Fun (Y (\<chi> j) a')"
using a' j YoD.at_simp Y_def Yo\<chi>.preserves_reflects_arr [of j]
by simp
also have "... = restrict (\<phi> (a', D (J.cod j)) \<circ> C (\<chi> j) \<circ> \<psi> (a', a))
(Hom.set (a', a))"
using a' j \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
Y_arr_ide [of a' "\<chi> j" a "D (J.cod j)"] \<chi>.A.map_simp S.Fun_mkArr
by fastforce
finally show ?thesis by blast
qed
thus ?thesis
using a a' h Hom.\<phi>_mapsto by auto
qed
also have "... = \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
using a a' h Hom.\<psi>_\<phi> by simp
finally show "S.Fun (YoD.at a' (map o \<chi>) j) (\<phi> (a', a) h)
= \<phi> (a', D (J.cod j)) (\<chi> j \<cdot> h)"
by auto
qed
have "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi>"
proof
fix j
have "\<not>J.arr j \<Longrightarrow> D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
using 1 2 \<chi>.cone_axioms by (elim C.in_homE, auto)
moreover have "J.arr j \<Longrightarrow> D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
proof -
assume j: "J.arr j"
have "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e)"
using j 1 \<chi>.cone_axioms by auto
also have "... = \<psi> (a', D (J.cod j)) (S.Fun (\<sigma> j) e)"
proof -
have "\<psi> (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun f' e)) =
\<psi> (a', D (J.cod j))
(\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e)))"
using j a a' f'e A Hom.\<phi>_\<psi> Hom.\<psi>_mapsto by force
moreover have "\<chi> j \<cdot> \<psi> (a', a) (S.Fun f' e) \<in> C.hom a' (D (J.cod j))"
using a a' j f'e Hom.\<psi>_mapsto \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
\<chi>.A.map_simp
by auto
moreover have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun f' e) =
S.Fun (\<sigma> j) e"
using Fun_map_a_a' a a' j f' e x Yo\<chi>_a'.A.map_simp eval_at_ide
Yo\<chi>_a'.cone_axioms
by auto
ultimately show ?thesis
using a a' Hom.\<psi>_\<phi> by auto
qed
also have "... = \<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e)"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e) =
\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e))"
using j a a' fe A [of "\<psi> (a', a) (S.Fun ?f e)" j] Hom.\<phi>_\<psi> Hom.\<psi>_mapsto
by auto
hence "\<psi> (a', D (J.cod j)) (S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e)) =
\<psi> (a', D (J.cod j))
(\<phi> (a', D (J.cod j)) (\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e)))"
by simp
moreover have "\<chi> j \<cdot> \<psi> (a', a) (S.Fun ?f e) \<in> C.hom a' (D (J.cod j))"
using a a' j fe Hom.\<psi>_mapsto \<chi>.preserves_hom [of j "J.dom j" "J.cod j"]
\<chi>.A.map_simp
by auto
moreover have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e) =
S.Fun (\<sigma> j) e"
proof -
have "S.Fun (YoD.at a' (map o \<chi>) j) (S.Fun ?f e)
= (S.Fun (YoD.at a' (map o \<chi>) j) o S.Fun ?f) e"
by simp
also have "... = S.Fun (YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f) e"
using Fun_map_a_a' a a' j f e x Yo\<chi>_a'.A.map_simp eval_at_ide
by auto
also have "... = S.Fun (\<sigma> j) e"
proof -
have "YoD.at a' (map o \<chi>) j \<cdot>\<^sub>S ?f =
YoD_a'.cones_map ?f (YoD.at a' (map o \<chi>)) j"
using j f Yo\<chi>_a'.cone_axioms Fun_map_a_a' by auto
thus ?thesis using j ff by argo
qed
finally show ?thesis by auto
qed
ultimately show ?thesis
using a a' Hom.\<psi>_\<phi> by auto
qed
also have "... = D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
using j 2 \<chi>.cone_axioms by force
finally show "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
by auto
qed
ultimately show "D.cones_map (\<psi> (a', a) (S.Fun f' e)) \<chi> j =
D.cones_map (\<psi> (a', a) (S.Fun ?f e)) \<chi> j"
by auto
qed
hence "\<psi> (a', a) (S.Fun f' e) = \<psi> (a', a) (S.Fun ?f e)"
using 1 2 \<chi>ofe.cone_axioms \<chi>.cone_axioms \<chi>.is_universal by blast
hence "\<phi> (a', a) (\<psi> (a', a) (S.Fun f' e)) = \<phi> (a', a) (\<psi> (a', a) (S.Fun ?f e))"
by simp
thus "S.Fun f' e = S.Fun ?f e"
using a a' fe f'e Hom.\<phi>_\<psi> by force
qed
ultimately show "S.Fun f' e = S.Fun ?f e" by auto
qed
qed
qed
ultimately have "\<exists>!f. \<guillemotleft>f : x \<rightarrow>\<^sub>S Hom.map (a', a)\<guillemotright> \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"
using ex1I [of "\<lambda>f. S.in_hom x (Hom.map (a', a)) f \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"]
by blast
thus "\<exists>!f. \<guillemotleft>f : x \<rightarrow>\<^sub>S Cop_S.Map (map a) a'\<guillemotright> \<and>
YoD_a'.cones_map f (YoD.at a' (map o \<chi>)) = \<sigma>"
using a a' Y_def by simp
qed
qed
thus "YoD.has_as_limit (map a)"
using YoD.cone_is_limit_if_pointwise_limit Yo\<chi>.cone_axioms by auto
qed
end
end
diff --git a/thys/Chandy_Lamport/Distributed_System.thy b/thys/Chandy_Lamport/Distributed_System.thy
--- a/thys/Chandy_Lamport/Distributed_System.thy
+++ b/thys/Chandy_Lamport/Distributed_System.thy
@@ -1,1829 +1,1829 @@
section \<open>Modelling distributed systems\<close>
text \<open>We assume familiarity with Chandy and Lamport's
paper \emph{Distributed Snapshots: Determining Global States of
Distributed Systems}~\cite{chandy}.\<close>
theory Distributed_System
imports Main
begin
type_synonym 'a fifo = "'a list"
type_synonym channel_id = nat
datatype 'm message =
Marker
| Msg 'm
datatype recording_state =
NotStarted
| Recording
| Done
text \<open>We characterize distributed systems by three underlying type variables:
Type variable 'p captures the processes of the underlying system.
Type variable 's describes the possible states of the processes.
Finally, type variable 'm describes all possible messages in said system.
Each process is in exactly one state at any point in time of the system.
Processes are interconnected by directed channels, which hold messages in-flight
between connected processes. There can be an arbitrary number of channels between
different processes. The entire state of the system including the (potentially unfinished)
snapshot state is called \emph{configuration}.\<close>
record ('p, 's, 'm) configuration =
states :: "'p \<Rightarrow> 's"
msgs :: "channel_id \<Rightarrow> 'm message fifo"
process_snapshot :: "'p \<Rightarrow> 's option"
channel_snapshot :: "channel_id \<Rightarrow> 'm fifo * recording_state"
text \<open>An event in Chandy and Lamport's formalization describes a
process' state transition, optionally producing or consuming
(but not both) a message on a channel. Additionally, a process may either initiate
a snapshot spontaneously, or is forced to do so by receiving a snapshot \emph{marker}
on one of it's incoming channels.\<close>
datatype ('p, 's, 'm) event =
isTrans: Trans (occurs_on: 'p) 's 's
| isSend: Send (getId: channel_id)
(occurs_on: 'p)
(partner: 'p)
's 's (getMsg: 'm)
| isRecv: Recv (getId: channel_id)
(occurs_on: 'p)
(partner: 'p)
's 's (getMsg: 'm)
| isSnapshot: Snapshot (occurs_on: 'p)
| isRecvMarker: RecvMarker (getId: channel_id)
(occurs_on: 'p)
(partner: 'p)
text \<open>We introduce abbreviations and type synoyms for commonly used terms.\<close>
type_synonym ('p, 's, 'm) trace = "('p, 's, 'm) event list"
abbreviation ps where "ps \<equiv> process_snapshot"
abbreviation cs where "cs \<equiv> channel_snapshot"
abbreviation no_snapshot_change where
"no_snapshot_change c c' \<equiv> ((\<forall>p'. ps c p' = ps c' p') \<and> (\<forall>i'. cs c i' = cs c' i'))"
abbreviation has_snapshotted where
"has_snapshotted c p \<equiv> process_snapshot c p \<noteq> None"
text \<open>A regular event is an event as described in Chandy and Lamport's
original paper: A state transition accompanied by the emission
or receiving of a message. Nonregular events are related to
snapshotting and receiving markers along communication channels.\<close>
definition regular_event[simp]:
"regular_event ev \<equiv> (isTrans ev \<or> isSend ev \<or> isRecv ev)"
lemma nonregular_event:
"~ regular_event ev = (isSnapshot ev \<or> isRecvMarker ev)"
by (meson event.distinct_disc event.exhaust_disc regular_event)
lemma event_occurs_on_unique:
assumes
"p \<noteq> q"
"occurs_on ev = p"
shows
"occurs_on ev \<noteq> q"
using assms by (cases ev, auto)
subsection \<open>The distributed system locale\<close>
text \<open>In order to capture Chandy and Lamport's computation system
we introduce two locales. The distributed system locale describes
global truths, such as the mapping from channel IDs to sender and
receiver processes, the transition relations for the underlying
computation system and the core assumption that no process has
a channel to itself. While not explicitly mentioned in Chandy's
and Lamport's work, it makes sense to assume that a channel need
not communicate to itself via messages, since it shares memory with
itself.\<close>
locale distributed_system =
fixes
channel :: "channel_id \<Rightarrow> ('p * 'p) option" and
trans :: "'p \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> bool" and
send :: "channel_id \<Rightarrow> 'p \<Rightarrow> 'p \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 'm \<Rightarrow> bool" and
recv :: "channel_id \<Rightarrow> 'p \<Rightarrow> 'p \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 'm \<Rightarrow> bool"
assumes
no_self_channel:
"\<forall>i. \<nexists>p. channel i = Some (p, p)"
begin
subsubsection \<open>State transitions\<close>
definition can_occur :: "('p, 's, 'm) event \<Rightarrow> ('p, 's, 'm) configuration \<Rightarrow> bool" where
"can_occur ev c \<equiv> (case ev of
Trans p s s' \<Rightarrow> states c p = s
\<and> trans p s s'
| Send i p q s s' msg \<Rightarrow> states c p = s
\<and> channel i = Some (p, q)
\<and> send i p q s s' msg
| Recv i p q s s' msg \<Rightarrow> states c p = s
\<and> channel i = Some (q, p)
\<and> length (msgs c i) > 0
\<and> hd (msgs c i) = Msg msg
\<and> recv i p q s s' msg
| Snapshot p \<Rightarrow> \<not> has_snapshotted c p
| RecvMarker i p q \<Rightarrow> channel i = Some (q, p)
\<and> length (msgs c i) > 0
\<and> hd (msgs c i) = Marker)"
definition src where
"src i p \<equiv> (\<exists>q. channel i = Some (p, q))"
definition dest where
"dest i q \<equiv> (\<exists>p. channel i = Some (p, q))"
lemma can_occur_Recv:
assumes
"can_occur (Recv i p q s s' m) c"
shows
"states c p = s \<and> channel i = Some (q, p) \<and> (\<exists>xs. msgs c i = Msg m # xs) \<and> recv i p q s s' m"
proof -
have "\<exists>xs. msgs c i = Msg m # xs"
using assms can_occur_def
by (metis (mono_tags, lifting) event.case(3) hd_Cons_tl length_greater_0_conv)
then show ?thesis using assms can_occur_def by auto
qed
abbreviation check_snapshot_occur where
"check_snapshot_occur c c' p \<equiv>
(can_occur (Snapshot p) c \<and>
(ps c' p = Some (states c p))
\<and> (\<forall>p'. states c p' = states c' p')
\<and> (\<forall>p'. (p' \<noteq> p) \<longrightarrow> ps c' p' = ps c p')
\<and> (\<forall>i. (\<exists>q. channel i = Some (p, q)) \<longrightarrow> msgs c' i = msgs c i @ [Marker])
\<and> (\<forall>i. (\<exists>q. channel i = Some (q, p)) \<longrightarrow> channel_snapshot c' i = (fst (channel_snapshot c i), Recording))
\<and> (\<forall>i. (\<nexists>q. channel i = Some (p, q)) \<longrightarrow> msgs c' i = msgs c i)
\<and> (\<forall>i. (\<nexists>q. channel i = Some (q, p)) \<longrightarrow> channel_snapshot c' i = channel_snapshot c i))"
abbreviation check_recv_marker_occur where
"check_recv_marker_occur c c' i p q \<equiv>
(can_occur (RecvMarker i p q) c
\<and> (\<forall>r. states c r = states c' r)
\<and> (\<forall>r. (r \<noteq> p) \<longrightarrow> process_snapshot c r = process_snapshot c' r)
\<and> (Marker # msgs c' i = msgs c i)
\<and> (channel_snapshot c' i = (fst (channel_snapshot c i), Done))
\<and> (if has_snapshotted c p
then (process_snapshot c p = process_snapshot c' p)
\<and> (\<forall>i'. (i' \<noteq> i) \<longrightarrow> msgs c' i' = msgs c i')
\<and> (\<forall>i'. (i' \<noteq> i) \<longrightarrow> channel_snapshot c i' = channel_snapshot c' i')
else (process_snapshot c' p = Some (states c p))
\<and> (\<forall>i'. i' \<noteq> i \<and> (\<exists>r. channel i' = Some (p, r))
\<longrightarrow> msgs c' i' = msgs c i' @ [Marker])
\<and> (\<forall>i'. i' \<noteq> i \<and> (\<exists>r. channel i' = Some (r, p))
\<longrightarrow> channel_snapshot c' i' = (fst (channel_snapshot c i'), Recording))
\<and> (\<forall>i'. i' \<noteq> i \<and> (\<nexists>r. channel i' = Some (p, r))
\<longrightarrow> msgs c' i' = msgs c i')
\<and> (\<forall>i'. i' \<noteq> i \<and> (\<nexists>r. channel i' = Some (r, p))
\<longrightarrow> channel_snapshot c' i' = channel_snapshot c i')))"
abbreviation check_trans_occur where
"check_trans_occur c c' p s s'\<equiv>
(can_occur (Trans p s s') c
\<and> (states c' p = s')
\<and> (\<forall>r. (r \<noteq> p) \<longrightarrow> states c' r = states c r)
\<and> (\<forall>i. msgs c' i = msgs c i)
\<and> (no_snapshot_change c c'))"
abbreviation check_send_occur where
"check_send_occur c c' i p q s s' msg \<equiv>
(can_occur (Send i p q s s' msg) c
\<and> (states c' p = s')
\<and> (\<forall>r. (r \<noteq> p) \<longrightarrow> states c' r = states c r)
\<and> (msgs c' i = msgs c i @ [Msg msg])
\<and> (\<forall>i'. i \<noteq> i' \<longrightarrow> msgs c' i' = msgs c i')
\<and> (no_snapshot_change c c'))"
abbreviation check_recv_occur where
"check_recv_occur c c' i p q s s' msg \<equiv>
(can_occur (Recv i p q s s' msg) c
\<and> (states c p = s \<and> states c' p = s')
\<and> (\<forall>r. (r \<noteq> p) \<longrightarrow> states c' r = states c r)
\<and> (msgs c i = Msg msg # msgs c' i)
\<and> (\<forall>i'. i \<noteq> i' \<longrightarrow> msgs c' i' = msgs c i')
\<and> (\<forall>r. process_snapshot c r = process_snapshot c' r)
\<and> (\<forall>i'. i' \<noteq> i \<longrightarrow> channel_snapshot c i' = channel_snapshot c' i')
\<and> (if snd (channel_snapshot c i) = Recording
then channel_snapshot c' i = (fst (channel_snapshot c i) @ [msg], Recording)
else channel_snapshot c i = channel_snapshot c' i))"
text \<open>The \emph{next} predicate lets us express configuration transitions
using events. The predicate $next(s_1, e, s_2)$ denotes the transition
of the configuration $s_1$ to $s_2$ via the event $e$. It ensures that
$e$ can occur in state $s_1$ and the state $s_2$ is correctly constructed
from $s_1$.\<close>
primrec "next" ::
"('p, 's, 'm) configuration
\<Rightarrow> ('p, 's, 'm) event
\<Rightarrow> ('p, 's, 'm) configuration
\<Rightarrow> bool"
("_ \<turnstile> _ \<mapsto> _" [70, 70, 70]) where
next_snapshot: "c \<turnstile> Snapshot p \<mapsto> c' =
check_snapshot_occur c c' p"
| next_recv_marker: "c \<turnstile> RecvMarker i p q \<mapsto> c' =
check_recv_marker_occur c c' i p q"
| next_trans: "c \<turnstile> Trans p s s' \<mapsto> c' =
check_trans_occur c c' p s s'"
| next_send: "c \<turnstile> Send i p q s s' msg \<mapsto> c' =
check_send_occur c c' i p q s s' msg"
| next_recv: "c \<turnstile> Recv i p q s s' msg \<mapsto> c' =
check_recv_occur c c' i p q s s' msg"
text \<open>Useful lemmas about state transitions\<close>
lemma state_and_event_determine_next:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"c \<turnstile> ev \<mapsto> c''"
shows
"c' = c''"
proof (cases ev)
case (Snapshot p)
then have "states c' = states c''" using assms by auto
moreover have "msgs c' = msgs c''"
proof (rule ext)
fix i
show "msgs c' i = msgs c'' i"
proof (cases "channel i = None")
case True
then show ?thesis using Snapshot assms by auto
next
case False
then obtain r s where "channel i = Some (r, s)" by auto
with assms Snapshot show ?thesis by (cases "r = p", simp_all)
qed
qed
moreover have "process_snapshot c' = process_snapshot c''" by (metis Snapshot assms next_snapshot ext)
moreover have "channel_snapshot c' = channel_snapshot c''"
proof (rule ext)
fix i
show "channel_snapshot c' i = channel_snapshot c'' i"
proof (cases "channel i = None")
case True
then show ?thesis using assms Snapshot by simp
next
case False
then obtain r s where "channel i = Some (r, s)" by auto
with assms Snapshot show ?thesis by (cases "s = p", simp_all)
qed
qed
ultimately show "c' = c''" by simp
next
case (RecvMarker i p)
then have "states c' = states c''" using assms by auto
moreover have "msgs c' = msgs c''"
proof (rule ext)
fix i'
show "msgs c' i' = msgs c'' i'"
proof (cases "i' = i")
case True
then have "Marker # msgs c' i' = msgs c i'" using assms RecvMarker by simp
- also have "... = Marker # msgs c'' i'" using assms RecvMarker `i' = i` by simp
+ also have "... = Marker # msgs c'' i'" using assms RecvMarker \<open>i' = i\<close> by simp
finally show ?thesis by simp
next
case False
then show ?thesis
proof (cases "has_snapshotted c p")
case True
- then show ?thesis using assms RecvMarker `i' \<noteq> i` by simp
+ then show ?thesis using assms RecvMarker \<open>i' \<noteq> i\<close> by simp
next
case no_snap: False
then show ?thesis
proof (cases "channel i' = None")
case True
- then show ?thesis using assms RecvMarker `i' \<noteq> i` no_snap by simp
+ then show ?thesis using assms RecvMarker \<open>i' \<noteq> i\<close> no_snap by simp
next
case False
then obtain r s where "channel i' = Some (r, s)" by auto
- with assms RecvMarker no_snap `i' \<noteq> i` show ?thesis by (cases "r = p"; simp_all)
+ with assms RecvMarker no_snap \<open>i' \<noteq> i\<close> show ?thesis by (cases "r = p"; simp_all)
qed
qed
qed
qed
moreover have "process_snapshot c' = process_snapshot c''"
proof (rule ext)
fix r
show "ps c' r = ps c'' r"
proof (cases "r \<noteq> p")
case True
then show ?thesis using assms RecvMarker by simp
next
case False
- with assms RecvMarker `~ r \<noteq> p` show ?thesis by (cases "has_snapshotted c r", auto)
+ with assms RecvMarker \<open>~ r \<noteq> p\<close> show ?thesis by (cases "has_snapshotted c r", auto)
qed
qed
moreover have "channel_snapshot c' = channel_snapshot c''"
proof (rule ext)
fix i'
show "cs c' i' = cs c'' i'"
proof (cases "i' = i")
case True
then show ?thesis using assms RecvMarker by simp
next
case False
then show ?thesis
proof (cases "has_snapshotted c p")
case True
- then show ?thesis using assms RecvMarker `i' \<noteq> i` by simp
+ then show ?thesis using assms RecvMarker \<open>i' \<noteq> i\<close> by simp
next
case no_snap: False
then show ?thesis
proof (cases "channel i' = None")
case True
- then show ?thesis using assms RecvMarker `i' \<noteq> i` no_snap by simp
+ then show ?thesis using assms RecvMarker \<open>i' \<noteq> i\<close> no_snap by simp
next
case False
then obtain r s where "channel i' = Some (r, s)" by auto
- with assms RecvMarker no_snap `i' \<noteq> i` show ?thesis by (cases "s = p"; simp_all)
+ with assms RecvMarker no_snap \<open>i' \<noteq> i\<close> show ?thesis by (cases "s = p"; simp_all)
qed
qed
qed
qed
ultimately show "c' = c''" by simp
next
case (Trans p s s')
then have "states c' = states c''" by (metis (no_types, lifting) assms next_trans ext)
moreover have "msgs c' = msgs c''" using assms Trans by auto
moreover have "process_snapshot c' = process_snapshot c''" using assms Trans by auto
moreover have "channel_snapshot c' = channel_snapshot c''" using assms Trans by auto
ultimately show "c' = c''" by simp
next
case (Send i p s s' m)
then have "states c' = states c''" by (metis (no_types, lifting) assms next_send ext)
moreover have "msgs c' = msgs c''"
proof (rule ext)
fix i'
from assms Send show "msgs c' i' = msgs c'' i'" by (cases "i' = i", simp_all)
qed
moreover have "process_snapshot c' = process_snapshot c''" using assms Send by auto
moreover have "channel_snapshot c' = channel_snapshot c''" using assms Send by auto
ultimately show "c' = c''" by simp
next
case (Recv i p s s' m)
then have "states c' = states c''" by (metis (no_types, lifting) assms next_recv ext)
moreover have "msgs c' = msgs c''"
proof (rule ext)
fix i'
from assms Recv show "msgs c' i' = msgs c'' i'" by (cases "i' = i", simp_all)
qed
moreover have "process_snapshot c' = process_snapshot c''" using assms Recv by auto
moreover have "channel_snapshot c' = channel_snapshot c''"
proof (rule ext)
fix i'
show "cs c' i' = cs c'' i'"
proof (cases "i' \<noteq> i")
case True
then show ?thesis using assms Recv by simp
next
case False
with assms Recv show ?thesis by (cases "snd (cs c i') = Recording", auto)
qed
qed
ultimately show "c' = c''" by simp
qed
lemma exists_next_if_can_occur:
assumes
"can_occur ev c"
shows
"\<exists>c'. c \<turnstile> ev \<mapsto> c'"
proof (cases ev)
case (Snapshot p)
let ?c = "\<lparr> states = states c,
msgs = %i. if (\<exists>q. channel i = Some (p, q)) then msgs c i @ [Marker] else msgs c i,
process_snapshot = %r. if r = p then Some (states c p) else ps c r,
channel_snapshot = %i. if (\<exists>q. channel i = Some (q, p)) then (fst (cs c i), Recording) else cs c i \<rparr>"
have "c \<turnstile> ev \<mapsto> ?c" using Snapshot assms by auto
then show ?thesis by blast
next
case (RecvMarker i p q)
show ?thesis
proof (cases "has_snapshotted c p")
case True
let ?c = "\<lparr> states = states c,
msgs = %i'. if i = i' then tl (msgs c i') else msgs c i',
process_snapshot = ps c,
channel_snapshot = %i'. if i = i' then (fst (cs c i'), Done) else cs c i' \<rparr>"
have "msgs c i = Marker # msgs ?c i"
using assms can_occur_def RecvMarker hd_Cons_tl by fastforce
then have "c \<turnstile> ev \<mapsto> ?c" using True RecvMarker assms by auto
then show ?thesis by blast
next
case False
let ?c = "\<lparr> states = states c,
msgs = %i'. if i' = i
then tl (msgs c i')
else if (\<exists>r. channel i' = Some (p, r))
then msgs c i' @ [Marker]
else msgs c i',
process_snapshot = %r. if r = p then Some (states c r) else ps c r,
channel_snapshot = %i'. if i = i' then (fst (cs c i'), Done)
else if (\<exists>r. channel i' = Some (r, p))
then (fst (cs c i'), Recording)
else cs c i' \<rparr>"
have "msgs c i = Marker # msgs ?c i"
using assms can_occur_def RecvMarker hd_Cons_tl by fastforce
moreover have "ps ?c p = Some (states c p)" by simp
ultimately have "c \<turnstile> ev \<mapsto> ?c" using RecvMarker assms False by auto
then show ?thesis by blast
qed
next
case (Trans p s s')
let ?c = "\<lparr> states = %r. if r = p then s' else states c r,
msgs = msgs c,
process_snapshot = ps c,
channel_snapshot = cs c \<rparr>"
have "c \<turnstile> ev \<mapsto> ?c"
using Trans assms by auto
then show ?thesis by blast
next
case (Send i p q s s' msg)
let ?c = "\<lparr> states = %r. if r = p then s' else states c r,
msgs = %i'. if i = i' then msgs c i' @ [Msg msg] else msgs c i',
process_snapshot = ps c,
channel_snapshot = cs c \<rparr>"
have "c \<turnstile> ev \<mapsto> ?c"
using Send assms by auto
then show ?thesis by blast
next
case (Recv i p q s s' msg)
then show ?thesis
proof (cases "snd (cs c i)")
case Recording
let ?c = "\<lparr> states = %r. if r = p then s' else states c r,
msgs = %i'. if i = i' then tl (msgs c i') else msgs c i',
process_snapshot = ps c,
channel_snapshot = %i'. if i = i'
then (fst (cs c i') @ [msg], Recording)
else cs c i'\<rparr>"
have "c \<turnstile> ev \<mapsto> ?c"
using Recv Recording assms can_occur_Recv by fastforce
then show ?thesis by blast
next
case Done
let ?c = "\<lparr> states = %r. if r = p then s' else states c r,
msgs = %i'. if i = i' then tl (msgs c i') else msgs c i',
process_snapshot = ps c,
channel_snapshot = cs c \<rparr>"
have "c \<turnstile> ev \<mapsto> ?c"
using Done Recv assms can_occur_Recv by fastforce
then show ?thesis by blast
next
case NotStarted
let ?c = "\<lparr> states = %r. if r = p then s' else states c r,
msgs = %i'. if i = i' then tl (msgs c i') else msgs c i',
process_snapshot = ps c,
channel_snapshot = cs c \<rparr>"
have "c \<turnstile> ev \<mapsto> ?c"
using NotStarted Recv assms can_occur_Recv by fastforce
then show ?thesis by blast
qed
qed
lemma exists_exactly_one_following_state:
"can_occur ev c \<Longrightarrow> \<exists>!c'. c \<turnstile> ev \<mapsto> c'"
using exists_next_if_can_occur state_and_event_determine_next by blast
lemma no_state_change_if_no_event:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"occurs_on ev \<noteq> p"
shows
"states c p = states c' p \<and> process_snapshot c p = process_snapshot c' p"
using assms by (cases ev, auto)
lemma no_msgs_change_if_no_channel:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"channel i = None"
shows
"msgs c i = msgs c' i"
using assms proof (cases ev)
case (RecvMarker cid p)
then have "cid \<noteq> i" using assms RecvMarker can_occur_def by fastforce
with assms RecvMarker show ?thesis by (cases "has_snapshotted c p", auto)
next
case (Send cid p s s' m)
then have "cid \<noteq> i" using assms Send can_occur_def by fastforce
then show ?thesis using assms Send by auto
next
case (Recv cid p s s' m)
then have "cid \<noteq> i" using assms Recv can_occur_def by fastforce
then show ?thesis using assms Recv by simp
qed simp_all
lemma no_cs_change_if_no_channel:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"channel i = None"
shows
"cs c i = cs c' i"
using assms proof (cases ev)
case (RecvMarker cid p)
then have "cid \<noteq> i" using assms RecvMarker can_occur_def by fastforce
with assms RecvMarker show ?thesis by (cases "has_snapshotted c p", auto)
next
case (Send cid p s s' m)
then have "cid \<noteq> i" using assms Send can_occur_def by fastforce
then show ?thesis using assms Send by auto
next
case (Recv cid p s s' m)
then have "cid \<noteq> i" using assms Recv can_occur_def by fastforce
then show ?thesis using assms Recv by simp
qed simp_all
lemma no_msg_change_if_no_event:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"isSend ev \<longrightarrow> getId ev \<noteq> i" and
"isRecv ev \<longrightarrow> getId ev \<noteq> i" and
"regular_event ev"
shows
"msgs c i = msgs c' i"
proof (cases "channel i = None")
case True
then show ?thesis using assms no_msgs_change_if_no_channel by simp
next
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms by simp
then show ?thesis
proof (elim disjE)
assume "isTrans ev"
then show ?thesis
by (metis assms(1) event.collapse(1) next_trans)
next
assume "isSend ev"
then obtain i' r s u u' m where Send: "ev = Send i' r s u u' m" by (meson isSend_def)
then show ?thesis using Send assms by auto
next
assume "isRecv ev"
then obtain i' r s u u' m where "ev = Recv i' r s u u' m" by (meson isRecv_def)
then show ?thesis using assms by auto
qed
qed
lemma no_cs_change_if_no_event:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"isRecv ev \<longrightarrow> getId ev \<noteq> i" and
"regular_event ev"
shows
"cs c i = cs c' i"
proof -
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms by simp
then show ?thesis
proof (elim disjE)
assume "isTrans ev"
then show ?thesis
by (metis assms(1) event.collapse(1) next_trans)
next
assume "isSend ev"
then obtain i' r s u u' m where "ev = Send i' r s u u' m" by (meson isSend_def)
then show ?thesis using assms by auto
next
assume "isRecv ev"
then obtain i r s u u' m where "ev = Recv i r s u u' m" by (meson isRecv_def)
then show ?thesis using assms by auto
qed
qed
lemma happen_implies_can_occur:
assumes
"c \<turnstile> ev \<mapsto> c'"
shows
"can_occur ev c"
proof -
show ?thesis using assms by (cases ev, auto)
qed
lemma snapshot_increases_message_length:
assumes
"ev = Snapshot p" and
"c \<turnstile> ev \<mapsto> c'" and
"channel i = Some (q, r)"
shows
"length (msgs c i) \<le> length (msgs c' i)"
using assms by (cases "p = q", auto)
lemma recv_marker_changes_head_only_at_i:
assumes
"ev = RecvMarker i p q" and
"c \<turnstile> ev \<mapsto> c'" and
"i' \<noteq> i"
shows
"msgs c i' = [] \<or> hd (msgs c i') = hd (msgs c' i')"
proof (cases "channel i' = None")
case True
then show ?thesis using assms no_msgs_change_if_no_channel by presburger
next
case False
then show ?thesis
proof (cases "msgs c i'")
case Nil
then show ?thesis by simp
next
case (Cons m xs)
then obtain r s where "channel i' = Some (r, s)" using False by auto
then show ?thesis
proof (cases "has_snapshotted c p")
case True
then show ?thesis using assms by auto
next
case False
with assms show ?thesis by (cases "r = p", auto)
qed
qed
qed
lemma recv_marker_other_channels_not_shrinking:
assumes
"ev = RecvMarker i p q" and
"c \<turnstile> ev \<mapsto> c'"
shows
"length (msgs c i') \<le> length (msgs c' i') \<longleftrightarrow> i \<noteq> i'"
proof (rule iffI)
show "length (msgs c i') \<le> length (msgs c' i') \<Longrightarrow> i \<noteq> i'"
proof (rule ccontr)
assume asm: "~ i \<noteq> i'" "length (msgs c i') \<le> length (msgs c' i')"
then have "msgs c i = Marker # msgs c' i" using assms by auto
then have "length (msgs c i) > length (msgs c' i)" by simp
then have "length (msgs c i') > length (msgs c' i')" using asm by simp
then show False using asm by simp
qed
next
show "i \<noteq> i' \<Longrightarrow> length (msgs c i') \<le> length (msgs c' i')"
proof -
assume "i \<noteq> i'"
then show ?thesis
proof (cases "channel i' = None")
case True
then show ?thesis using assms no_msgs_change_if_no_channel by presburger
next
case False
then obtain r s where chan: "channel i' = Some (r, s)" by auto
then show ?thesis
proof (cases "has_snapshotted c p")
case True
- with assms `i \<noteq> i'` show ?thesis by auto
+ with assms \<open>i \<noteq> i'\<close> show ?thesis by auto
next
case no_snap: False
then show ?thesis
proof (cases "p = r")
case True
- then have "msgs c' i' = msgs c i' @ [Marker]" using `i \<noteq> i'` assms no_snap chan by auto
+ then have "msgs c' i' = msgs c i' @ [Marker]" using \<open>i \<noteq> i'\<close> assms no_snap chan by auto
then show ?thesis by auto
next
case False
- then show ?thesis using assms `i \<noteq> i'` chan no_snap by auto
+ then show ?thesis using assms \<open>i \<noteq> i'\<close> chan no_snap by auto
qed
qed
qed
qed
qed
lemma regular_event_cannot_induce_snapshot:
assumes
"~ has_snapshotted c p" and
"c \<turnstile> ev \<mapsto> c'"
shows
"regular_event ev \<longrightarrow> ~ has_snapshotted c' p"
proof (cases ev)
case (Trans q s s')
then show ?thesis using assms(1) assms(2) by auto
next
case (Send q r s s' m)
then show ?thesis using assms by auto
next
case (Recv q r s s' m)
then show ?thesis using assms by auto
qed simp_all
lemma regular_event_preserves_process_snapshots:
assumes
"c \<turnstile> ev \<mapsto> c'"
shows
"regular_event ev \<Longrightarrow> ps c r = ps c' r"
proof (cases ev)
case (Trans p s s')
then show ?thesis
using assms by auto
next
case (Send p q s s' m)
then show ?thesis
using assms by auto
next
case (Recv p q s s' m)
then show ?thesis
using assms by auto
qed simp_all
lemma no_state_change_if_nonregular_event:
assumes
"~ regular_event ev" and
"c \<turnstile> ev \<mapsto> c'"
shows
"states c p = states c' p"
proof -
have "isSnapshot ev \<or> isRecvMarker ev" using nonregular_event assms by auto
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then obtain q where "ev = Snapshot q"
by (meson isSnapshot_def)
then show ?thesis
using assms(2) by auto
next
case 2
then obtain i q r where "ev = RecvMarker i q r"
by (meson isRecvMarker_def)
then show ?thesis using assms(2) by auto
qed
qed
lemma nonregular_event_induces_snapshot:
assumes
"~ has_snapshotted c p" and
"c \<turnstile> ev \<mapsto> c'" and
"occurs_on ev = p" and
"~ regular_event ev"
shows
"~ regular_event ev \<longrightarrow> has_snapshotted c' p"
proof (cases ev)
case (Snapshot q)
then have "q = p" using assms by auto
then show ?thesis using Snapshot assms(2) by auto
next
case (RecvMarker i q r)
then have "q = p" using assms by auto
then show ?thesis using RecvMarker assms by auto
qed (simp_all add: assms)
lemma snapshot_state_unchanged:
assumes
step: "c \<turnstile> ev \<mapsto> c'" and
"has_snapshotted c p"
shows
"ps c p = ps c' p"
proof (cases "occurs_on ev = p")
case False
then show ?thesis
using local.step no_state_change_if_no_event by auto
next
case True
then show ?thesis
proof (cases "regular_event ev")
case True
then show ?thesis
using local.step regular_event_preserves_process_snapshots by auto
next
case False
have "isRecvMarker ev"
proof (rule ccontr)
have "isSnapshot ev \<or> isRecvMarker ev"
using False nonregular_event by blast
moreover assume "~ isRecvMarker ev"
ultimately have "isSnapshot ev" by simp
then have "ev = Snapshot p" by (metis True event.collapse(4))
then have "can_occur ev c"
using happen_implies_can_occur local.step by blast
then have "~ has_snapshotted c p" unfolding can_occur_def
by (simp add: \<open>ev = Snapshot p\<close>)
then show False using assms by auto
qed
then show ?thesis (* z3 sledgehammer fails for Isabelle2019 *)
proof -
have "\<exists>n pa. c \<turnstile> RecvMarker n p pa \<mapsto> c'"
by (metis True \<open>isRecvMarker ev\<close> event.collapse(5) local.step)
then show ?thesis
using assms(2) by force
qed
qed
qed
lemma message_must_be_delivered:
assumes
valid: "c \<turnstile> ev \<mapsto> c'" and
delivered: "(msgs c i \<noteq> [] \<and> hd (msgs c i) = m) \<and> (msgs c' i = [] \<or> hd (msgs c' i) \<noteq> m)"
shows
"(\<exists>p q. ev = RecvMarker i p q \<and> m = Marker)
\<or> (\<exists>p q s s' m'. ev = Recv i p q s s' m' \<and> m = Msg m')"
proof (cases ev)
case (Snapshot p)
then show ?thesis
proof (cases "msgs c i")
case Nil
then show ?thesis using delivered by simp
next
case (Cons m xs)
with assms Snapshot show ?thesis
proof (cases "channel i = None")
case True
then show ?thesis using assms Snapshot by auto
next
case False
then obtain r s where chan: "channel i = Some (r, s)" by auto
then show ?thesis
proof (cases "r = p")
case True
then have "msgs c' i = msgs c i @ [Marker]" using assms(1) Snapshot chan by auto
then show ?thesis using delivered by auto
next
case False
then have "msgs c' i = msgs c i" using assms Snapshot chan by simp
then show ?thesis using delivered Cons by simp
qed
qed
qed
next
case (RecvMarker i' p q)
then have "i' = i"
by (metis assms(1) delivered le_0_eq length_greater_0_conv list.size(3) recv_marker_changes_head_only_at_i recv_marker_other_channels_not_shrinking)
moreover have "Marker = m"
- using `i' = i` RecvMarker assms(1) can_occur_def delivered by auto
+ using \<open>i' = i\<close> RecvMarker assms(1) can_occur_def delivered by auto
moreover have "channel i = Some (q, p)"
using RecvMarker assms(1) calculation(1) can_occur_def by auto
ultimately show ?thesis using RecvMarker by simp
next
case (Trans p' s s')
then show ?thesis
using valid delivered by auto
next
case (Send p' q' s s' m')
then show ?thesis
by (metis (no_types, lifting) delivered distributed_system.next.simps(4) distributed_system_axioms hd_append2 snoc_eq_iff_butlast valid)
next
case (Recv i' p q s s' m')
then have "i = i'"
using assms(1) delivered by auto
also have "m = Msg m'"
by (metis (no_types, lifting) Recv delivered list.sel(1) next_recv valid)
ultimately show ?thesis using Recv by auto
qed
lemma message_must_be_delivered_2:
assumes
"c \<turnstile> ev \<mapsto> c'"
"m : set (msgs c i)"
"m \<notin> set (msgs c' i)"
shows
"(\<exists>p q. ev = RecvMarker i p q \<and> m = Marker) \<or> (\<exists>p q s s' m'. ev = Recv i p q s s' m' \<and> m = Msg m')"
proof -
have uneq_sets: "set (msgs c i) \<noteq> set (msgs c' i)"
using assms(2) assms(3) by blast
then obtain p q where chan: "channel i = Some (p, q)"
using assms no_msgs_change_if_no_channel by fastforce
then show ?thesis
proof (cases ev)
case (Snapshot p')
with Snapshot assms chan have "set (msgs c' i) = set (msgs c i)" by (cases "p' = p", auto)
then show ?thesis using uneq_sets by simp
next
case (Trans p' s s')
then show ?thesis using uneq_sets assms by simp
next
case (Send i' p' q' s s' m)
then show ?thesis
by (metis (no_types, lifting) UnCI assms(1) assms(2) assms(3) local.next.simps(4) set_append)
next
case (RecvMarker i' p' q')
have "i' = i"
proof (rule ccontr)
assume "~ i' = i"
show False using assms chan RecvMarker
proof (cases "has_snapshotted c p'")
case True
- then show False using assms chan RecvMarker `~ i' = i` by simp
+ then show False using assms chan RecvMarker \<open>~ i' = i\<close> by simp
next
case False
- then show False using assms chan RecvMarker `~ i' = i` by (cases "p' = p", simp_all)
+ then show False using assms chan RecvMarker \<open>~ i' = i\<close> by (cases "p' = p", simp_all)
qed
qed
moreover have "m = Marker"
proof -
have "msgs c i' = Marker # msgs c' i'" using assms chan RecvMarker by auto
- then show ?thesis using assms `i' = i` by simp
+ then show ?thesis using assms \<open>i' = i\<close> by simp
qed
ultimately show ?thesis using RecvMarker by simp
next
case (Recv i' p' q' s s' m')
have "i' = i"
proof (rule ccontr)
assume "~ i' = i"
then show False
using Recv assms(1) uneq_sets by auto
qed
then have "i' = i \<and> m = Msg m'"
using Recv assms by auto
then show ?thesis using Recv by simp
qed
qed
lemma recv_marker_means_snapshotted_1:
assumes
"ev = RecvMarker i p q" and
"c \<turnstile> ev \<mapsto> c'"
shows
"has_snapshotted c' p"
using assms snapshot_state_unchanged by (cases "has_snapshotted c p", auto)
lemma recv_marker_means_snapshotted_2:
fixes
c c' :: "('p, 's, 'm) configuration" and
ev :: "('p, 's, 'm) event" and
i :: channel_id
assumes
"c \<turnstile> ev \<mapsto> c'" and
"Marker : set (msgs c i)" and
"Marker \<notin> set (msgs c' i)" and
"channel i = Some (q, p)"
shows
"has_snapshotted c' p"
proof -
have "\<exists>p q. ev = RecvMarker i p q"
using assms message_must_be_delivered_2 by blast
then obtain r s where RecvMarker: "ev = RecvMarker i r s"
by blast
then have "r = p"
using assms(1) assms(4) can_occur_def by auto
then show ?thesis
using recv_marker_means_snapshotted_1 assms RecvMarker by blast
qed
lemma event_stays_valid_if_no_occurrence:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"occurs_on ev \<noteq> occurs_on ev'" and
"can_occur ev' c"
shows
"can_occur ev' c'"
proof (cases ev')
case (Trans p s s')
have "states c p = states c' p"
using Trans assms(1) assms(2) no_state_change_if_no_event by auto
moreover have "states c p = s" using can_occur_def assms Trans by simp
ultimately have "states c' p = s" by simp
moreover have "trans p s s'"
using Trans assms(3) can_occur_def by auto
ultimately show ?thesis
by (simp add: Trans can_occur_def)
next
case (Recv i p q s s' m)
then have "hd (msgs c i) = Msg m"
proof -
from Recv have "length (msgs c i) > 0" using assms(3) can_occur_def by auto
then obtain m' xs where mcqp: "msgs c i = m' # xs"
by (metis list.size(3) nat_less_le neq_Nil_conv)
then have "Msg m = m'"
proof (cases m', auto)
case Marker
then have "msgs c i = Marker # xs" by (simp add:mcqp)
then have "~ can_occur ev' c" using Recv can_occur_def by simp
then show False using assms(3) by simp
next
case (Msg msg)
then have "msgs c i = Msg msg # xs" by (simp add: mcqp)
then show "m = msg" using Recv can_occur_def assms(3) by simp
qed
then show ?thesis by (simp add: mcqp)
qed
show ?thesis
proof (rule ccontr)
assume asm: "~ can_occur ev' c'"
then have "msgs c' i = [] \<or> hd (msgs c' i) \<noteq> Msg m"
using Recv assms can_occur_def no_state_change_if_no_event distributed_system_axioms list.case_eq_if by fastforce
then obtain i' p' q' s'' s''' m' where RMoR: "ev = RecvMarker i' p' q' \<or> ev = Recv i p' q' s'' s''' m'"
by (metis Recv \<open>hd (msgs c i) = Msg m\<close> assms(1) assms(3) can_occur_Recv list.discI message_must_be_delivered)
then have "occurs_on ev = p"
proof -
have f1: "states c p = s \<and> channel i = Some (q, p) \<and> recv i p q s s' m \<and> 0 < length (msgs c i) \<and> hd (msgs c i) = Msg m"
using Recv assms(3) can_occur_def by force
have f2: "RecvMarker i' p' q' = ev \<or> states c p' = s'' \<and> channel i = Some (q', p') \<and> recv i p' q' s'' s''' m' \<and> 0 < length (msgs c i) \<and> hd (msgs c i) = Msg m'"
using RMoR assms(1) can_occur_def by force
have "\<forall>e n c. \<exists>p pa s sa m. \<forall>ca cb. (\<not> c \<turnstile> e \<mapsto> ca \<or> msgs ca n \<noteq> [] \<or> hd (msgs c n) = Marker \<or> msgs c n = [] \<or> Recv n p pa s sa m = e) \<and> (\<not> c \<turnstile> e \<mapsto> cb \<or> hd (msgs c n) = Marker \<or> hd (msgs cb n) = hd (msgs c n) \<or> msgs c n = [] \<or> Recv n p pa s sa m = e)"
by (metis (no_types) message_must_be_delivered)
then show ?thesis
using f2 f1 by (metis RMoR \<open>msgs c' i = [] \<or> hd (msgs c' i) \<noteq> Msg m\<close> assms(1) event.disc(13,15) event.sel(3,5) length_greater_0_conv message.distinct(1) option.inject prod.inject)
qed
then show False using assms Recv by simp
qed
next
case (Send i p q s s' m)
then have "states c p = states c' p" using assms no_state_change_if_no_event by auto
then show "can_occur ev' c'" using assms assms(3) can_occur_def Send by auto
next
case (RecvMarker i p q)
then have msgs_ci: "hd (msgs c i) = Marker \<and> length (msgs c i) > 0"
proof -
from RecvMarker have "length (msgs c i) > 0" using assms(3) can_occur_def by auto
then obtain m' xs where mci: "msgs c i = m' # xs"
by (metis list.size(3) nat_less_le neq_Nil_conv)
then have m_mark: "Marker = m'"
proof (cases m', auto)
case (Msg msg)
then have "msgs c i = Msg msg # xs" by (simp add:mci)
then have "~ can_occur ev' c" using RecvMarker can_occur_def by simp
then show False using assms(3) by simp
qed
then show ?thesis by (simp add: mci)
qed
show ?thesis
proof (rule ccontr)
assume asm: "~ can_occur ev' c'"
then have "msgs c' i = [] \<or> hd (msgs c' i) \<noteq> Marker"
using RecvMarker assms(3) can_occur_def list.case_eq_if by fastforce
then have "\<exists>p q. ev = RecvMarker i p q \<and> Marker = Marker" using message_must_be_delivered msgs_ci assms by blast
then obtain r s where RecvMarker_ev: "ev = RecvMarker i r s" by blast
then have "p = r \<and> q = s"
using RecvMarker assms(1) assms(3) can_occur_def by auto
then have "occurs_on ev = p" using assms RecvMarker_ev by auto
then show False using assms using RecvMarker by auto
qed
next
case (Snapshot p)
then have "~ has_snapshotted c p" using assms assms(3) can_occur_def by simp
show ?thesis
proof (rule ccontr)
assume asm: "~ can_occur ev' c'"
then have "has_snapshotted c' p" using can_occur_def Snapshot by simp
then have "occurs_on ev = p"
using \<open>\<not> has_snapshotted c p\<close> assms(1) no_state_change_if_no_event by fastforce
then show False using assms(2) Snapshot by auto
qed
qed
lemma msgs_unchanged_for_other_is:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"regular_event ev" and
"getId ev = i" and
"i' \<noteq> i"
shows
"msgs c i' = msgs c' i'"
proof -
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms by simp
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then obtain p s s' where "ev = Trans p s s'" by (meson isTrans_def)
then show ?thesis using assms by simp
next
case 2
then obtain i' p q s s' m where "ev = Send i' p q s s' m" by (meson isSend_def)
then show ?thesis using assms by simp
next
case 3
then obtain i' p q s s' m where "ev = Recv i' p q s s' m" by (meson isRecv_def)
with assms show ?thesis by auto
qed
qed
lemma msgs_unchanged_if_snapshotted_RecvMarker_for_other_is:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"ev = RecvMarker i p q" and
"has_snapshotted c p" and
"i' \<noteq> i"
shows
"msgs c i' = msgs c' i'"
using assms by auto
lemma event_can_go_back_if_no_sender:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"occurs_on ev \<noteq> occurs_on ev'" and
"can_occur ev' c'" and
"~ isRecvMarker ev'" and
"~ isSend ev"
shows
"can_occur ev' c"
proof (cases ev')
case (Snapshot p)
then have "~ has_snapshotted c' p" using assms(3) can_occur_def by simp
then have "~ has_snapshotted c p" using assms(1) snapshot_state_unchanged by force
then show ?thesis using can_occur_def Snapshot by simp
next
case (RecvMarker i p q)
then show ?thesis using assms(4) by auto
next
case (Trans p s s')
then show ?thesis
using assms(1) assms(2) can_occur_def no_state_change_if_no_event assms(3) by auto
next
case (Send p q s s' m)
then show ?thesis
using assms(1) assms(2) can_occur_def no_state_change_if_no_event assms(3) by auto
next
case (Recv i p q s s' m)
have "msgs c' i \<noteq> Nil" using Recv can_occur_def assms by auto
moreover have "hd (msgs c' i) = Msg m \<and> length (msgs c' i) > 0"
proof -
from Recv have "length (msgs c' i) > 0" using assms(3) can_occur_def by auto
then obtain m' xs where mcqp: "msgs c' i = m' # xs"
by (metis list.size(3) nat_less_le neq_Nil_conv)
then have "Msg m = m'"
proof (cases m', auto)
case Marker
then have "msgs c' i = Marker # xs" by (simp add:mcqp)
then have "~ can_occur ev' c'" using Recv can_occur_def by simp
then show False using assms(3) by simp
next
case (Msg msg)
then have "msgs c' i = Msg msg # xs" by (simp add: mcqp)
then show "m = msg" using Recv can_occur_def assms(3) by simp
qed
then show ?thesis by (simp add: mcqp)
qed
moreover have "msgs c i \<noteq> Nil \<and> hd (msgs c' i) = hd (msgs c i)"
proof (cases ev)
case (Snapshot p')
then have "p' \<noteq> p" using assms Recv by simp
have chan: "channel i = Some (q, p)"
by (metis Recv assms(3) distributed_system.can_occur_Recv distributed_system_axioms)
with Snapshot assms have "length (msgs c i) > 0 \<and> hd (msgs c i) = hd (msgs c' i)"
proof (cases "q = p'")
case True
then have "msgs c' i = msgs c i @ [Marker]" using Snapshot chan assms by simp
then show ?thesis
by (metis append_self_conv2 calculation(2) hd_append2 length_greater_0_conv list.sel(1) message.simps(3))
next
case False
then have "msgs c' i = msgs c i" using Snapshot chan assms by simp
then show ?thesis using calculation by simp
qed
then show ?thesis by simp
next
case (RecvMarker i' p' q')
then have "i' \<noteq> i"
using Recv assms(1) assms(2) assms(3) can_occur_def by force
then show ?thesis
proof (cases "has_snapshotted c p'")
case True
- then have "msgs c i = msgs c' i" using `i' \<noteq> i` RecvMarker assms by simp
+ then have "msgs c i = msgs c' i" using \<open>i' \<noteq> i\<close> RecvMarker assms by simp
then show ?thesis using calculation by simp
next
case no_snap: False
then have chan: "channel i = Some (q, p)"
by (metis Recv assms(3) distributed_system.can_occur_Recv distributed_system_axioms)
then show ?thesis
proof (cases "q = p'")
case True
then have "msgs c' i = msgs c i @ [Marker]"
using no_snap RecvMarker \<open>i' \<noteq> i\<close> assms(1) chan by auto
then show ?thesis
by (metis append_self_conv2 calculation(2) hd_append2 list.sel(1) message.simps(3))
next
case False
- then have "msgs c' i = msgs c i" using RecvMarker no_snap False chan assms `i' \<noteq> i` by simp
+ then have "msgs c' i = msgs c i" using RecvMarker no_snap False chan assms \<open>i' \<noteq> i\<close> by simp
then show ?thesis using calculation by simp
qed
qed
next
case (Trans p' s'' s''')
- then show ?thesis using assms(1) `msgs c' i \<noteq> Nil` by auto
+ then show ?thesis using assms(1) \<open>msgs c' i \<noteq> Nil\<close> by auto
next
case (Send i' p' q' s'' s''' m'')
have "p' \<noteq> p"
using Recv Send assms(2) by auto
then show ?thesis
using Recv Send assms(1) assms(5) calculation(1) by auto
next
case (Recv i' p' q' s'' s''' m'')
- then have "i' \<noteq> i" using assms `ev' = Recv i p q s s' m`
+ then have "i' \<noteq> i" using assms \<open>ev' = Recv i p q s s' m\<close>
by (metis distributed_system.can_occur_Recv distributed_system_axioms event.sel(3) next_recv option.inject prod.inject)
have "msgs c i = msgs c' i" using msgs_unchanged_for_other_is Recv \<open>i' \<noteq> i\<close> assms(1) by auto
- then show ?thesis using `msgs c' i \<noteq> Nil` by simp
+ then show ?thesis using \<open>msgs c' i \<noteq> Nil\<close> by simp
qed
moreover have "states c p = states c' p" using no_state_change_if_no_event assms Recv by simp
ultimately show ?thesis
using Recv assms(3) can_occur_def list.case_eq_if by fastforce
qed
lemma nonregular_event_can_go_back_if_in_distinct_processes:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"regular_event ev" and
"~ regular_event ev'" and
"can_occur ev' c'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"can_occur ev' c"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms by simp
moreover have "isSnapshot ev' \<or> isRecvMarker ev'" using assms nonregular_event by auto
ultimately show ?thesis
proof (elim disjE, goal_cases)
case 1
then show ?case
using assms(1) assms(4) assms(5) event_can_go_back_if_no_sender by blast
next
case 2
then obtain s s' where Trans: "ev = Trans ?p s s'"
by (metis event.collapse(1))
obtain i r where RecvMarker: "ev' = RecvMarker i ?q r"
using 2 by (metis event.collapse(5))
have "msgs c i = msgs c' i"
using "2"(1) assms(1) assms(2) no_msg_change_if_no_event by blast
moreover have "can_occur ev' c'" using assms by simp
ultimately show ?thesis using can_occur_def RecvMarker
by (metis (mono_tags, lifting) "2"(2) event.case_eq_if event.distinct_disc(13) event.distinct_disc(17) event.distinct_disc(19) event.distinct_disc(7) event.sel(10))
next
case 3
then have "ev' = Snapshot ?q"
by (metis event.collapse(4))
have "~ has_snapshotted c' ?q"
by (metis (mono_tags, lifting) "3"(1) assms(4) can_occur_def event.case_eq_if event.distinct_disc(11) event.distinct_disc(16) event.distinct_disc(6))
then have "~ has_snapshotted c ?q"
using assms(1) assms(2) regular_event_preserves_process_snapshots by auto
- then show ?case unfolding can_occur_def using `ev' = Snapshot ?q`
+ then show ?case unfolding can_occur_def using \<open>ev' = Snapshot ?q\<close>
by (metis (mono_tags, lifting) event.simps(29))
next
case 4
then have "ev' = Snapshot ?q"
by (metis event.collapse(4))
have "~ has_snapshotted c' ?q"
by (metis (mono_tags, lifting) \<open>ev' = Snapshot (occurs_on ev')\<close> assms(4) can_occur_def event.simps(29))
then have "~ has_snapshotted c ?q"
using assms(1) assms(2) regular_event_preserves_process_snapshots by auto
then show ?case unfolding can_occur_def
by (metis (mono_tags, lifting) \<open>ev' = Snapshot (occurs_on ev')\<close> event.simps(29))
next
case 5
then obtain i s u u' m where "ev = Send i ?p s u u' m"
by (metis event.collapse(2))
from 5 obtain i' r where "ev' = RecvMarker i' ?q r"
by (metis event.collapse(5))
then have pre: "hd (msgs c' i') = Marker \<and> length (msgs c' i') > 0"
by (metis (mono_tags, lifting) assms(4) can_occur_def event.simps(30))
have "hd (msgs c i') = Marker \<and> length (msgs c i') > 0"
proof (cases "i' = i")
case False
then have "msgs c i' = msgs c' i'"
by (metis \<open>ev = Send i (occurs_on ev) s u u' m\<close> assms(1) assms(2) event.sel(8) msgs_unchanged_for_other_is)
then show ?thesis using pre by auto
next
case True
then have "msgs c' i' = msgs c i' @ [Msg m]"
by (metis \<open>ev = Send i (occurs_on ev) s u u' m\<close> assms(1) next_send)
then have "length (msgs c' i') > 1"
using pre by fastforce
then have "length (msgs c i') > 0"
by (simp add: \<open>msgs c' i' = msgs c i' @ [Msg m]\<close>)
then show ?thesis
using \<open>msgs c' i' = msgs c i' @ [Msg m]\<close> pre by auto
qed
- then show ?case unfolding can_occur_def using `ev' = RecvMarker i' ?q r`
+ then show ?case unfolding can_occur_def using \<open>ev' = RecvMarker i' ?q r\<close>
by (metis (mono_tags, lifting) assms(4) can_occur_def event.simps(30))
next
case 6
then obtain i s u u' m where "ev = Recv i ?p s u u' m"
by (metis event.collapse(3))
from 6 obtain i' r where "ev' = RecvMarker i' ?q r"
by (metis event.collapse(5))
then have "i' \<noteq> i"
proof -
have "?p \<noteq> ?q" using assms by simp
moreover have "channel i = Some (s, ?p)"
by (metis \<open>ev = Recv i (occurs_on ev) s u u' m\<close> assms(1) distributed_system.can_occur_Recv distributed_system_axioms happen_implies_can_occur)
moreover have "channel i' = Some (r, ?q)"
by (metis (mono_tags, lifting) \<open>ev' = RecvMarker i' (occurs_on ev') r\<close> assms(4) can_occur_def event.case_eq_if event.disc(5,10,15,20) event.sel(5,10,13))
ultimately show ?thesis by auto
qed
then show ?case
by (metis (mono_tags, lifting) "6"(1) \<open>ev = Recv i (occurs_on ev) s u u' m\<close> \<open>ev' = RecvMarker i' (occurs_on ev') r\<close> assms(1) assms(4) can_occur_def event.case_eq_if event.distinct_disc(13) event.distinct_disc(17) event.distinct_disc(7) event.sel(10) next_recv)
qed
qed
lemma same_state_implies_same_result_state:
assumes
"states c p = states d p"
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'"
shows
"states d' p = states c' p"
proof (cases "occurs_on ev = p")
case False
then show ?thesis
by (metis assms(1-3) distributed_system.no_state_change_if_no_event distributed_system_axioms)
next
case True
then show ?thesis
using assms by (cases ev, auto)
qed
lemma same_snapshot_state_implies_same_result_snapshot_state:
assumes
"ps c p = ps d p" and
"states c p = states d p" and
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'"
shows
"ps d' p = ps c' p"
proof (cases "occurs_on ev = p")
case False
then show ?thesis
using assms no_state_change_if_no_event by auto
next
case True
then show ?thesis
proof (cases ev)
case (Snapshot q)
then have "p = q" using True by auto
then show ?thesis
using Snapshot assms(2) assms(3) assms(4) by auto
next
case (RecvMarker i q r)
then have "p = q" using True by auto
then show ?thesis
proof -
have f1: "\<And>c ca. \<not> c \<turnstile> ev \<mapsto> ca \<or> ps c p = None \<or> ps c p = ps ca p"
using RecvMarker \<open>p = q\<close> by force
have "\<And>c ca. ps c p \<noteq> None \<or> \<not> c \<turnstile> ev \<mapsto> ca \<or> ps ca p = Some (states c p)"
using RecvMarker \<open>p = q\<close> by force
then show ?thesis
using f1 by (metis (no_types) assms(1) assms(2) assms(3) assms(4))
qed
next
case (Trans q s s')
then have "p = q"
using True by auto
then show ?thesis
using Trans assms(1) assms(3) assms(4) by auto
next
case (Send i q r u u' m)
then have "p = q" using True by auto
then show ?thesis
using Send assms(1) assms(3) assms(4) by auto
next
case (Recv i q r u u' m)
then have "p = q" using True by auto
then show ?thesis
using Recv assms(1) assms(3) assms(4) by auto
qed
qed
lemma same_messages_imply_same_resulting_messages:
assumes
"msgs c i = msgs d i"
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'" and
"regular_event ev"
shows
"msgs c' i = msgs d' i"
proof -
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms
by simp
then show ?thesis
proof (elim disjE)
assume "isTrans ev"
then show ?thesis
by (metis assms(1) assms(2) assms(3) isTrans_def next_trans)
next
assume "isSend ev"
then obtain i' r s u u' m where "ev = Send i' r s u u' m"
by (metis event.collapse(2))
with assms show ?thesis by (cases "i = i'", auto)
next
assume "isRecv ev"
then obtain i' r s u u' m where Recv: "ev = Recv i' r s u u' m"
by (metis event.collapse(3))
with assms show ?thesis by (cases "i = i'", auto)
qed
qed
lemma Trans_msg:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"isTrans ev"
shows
"msgs c i = msgs c' i"
using assms(1) assms(2) no_msg_change_if_no_event regular_event by blast
lemma new_msg_in_set_implies_occurrence:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"m \<notin> set (msgs c i)" and
"m \<in> set (msgs c' i)" and
"channel i = Some (p, q)"
shows
"occurs_on ev = p" (is ?P)
proof (rule ccontr)
assume "~ ?P"
have "set (msgs c' i) \<subseteq> set (msgs c i)"
proof (cases ev)
case (Snapshot r)
- then have "msgs c' i = msgs c i" using `~ ?P` assms by simp
+ then have "msgs c' i = msgs c i" using \<open>~ ?P\<close> assms by simp
then show ?thesis by auto
next
case (RecvMarker i' r s)
then show ?thesis
proof (cases "has_snapshotted c r")
case True
then show ?thesis
proof (cases "i' = i")
case True
then have "Marker # msgs c' i = msgs c i" using RecvMarker True assms by simp
then show ?thesis
by (metis set_subset_Cons)
next
case False
then show ?thesis using RecvMarker True assms by simp
qed
next
case no_snap: False
have chan: "channel i' = Some (s, r)"
using RecvMarker assms(1) can_occur_def by auto
then show ?thesis
proof (cases "i' = i")
case True
then have "Marker # msgs c' i = msgs c i" using RecvMarker assms by simp
then show ?thesis by (metis set_subset_Cons)
next
case False
- then have "msgs c' i = msgs c i" using `~ ?P` RecvMarker assms no_snap by simp
+ then have "msgs c' i = msgs c i" using \<open>~ ?P\<close> RecvMarker assms no_snap by simp
then show ?thesis by simp
qed
qed
next
case (Trans r u u')
- then show ?thesis using assms `~ ?P` by simp
+ then show ?thesis using assms \<open>~ ?P\<close> by simp
next
case (Send i' r s u u' m')
- then have "i' \<noteq> i" using `~ ?P` can_occur_def assms by auto
- then have "msgs c i = msgs c' i" using `~ ?P` assms Send by simp
+ then have "i' \<noteq> i" using \<open>~ ?P\<close> can_occur_def assms by auto
+ then have "msgs c i = msgs c' i" using \<open>~ ?P\<close> assms Send by simp
then show ?thesis by simp
next
case (Recv i' r s u u' m')
then show ?thesis
by (metis (no_types, lifting) assms(1) eq_iff local.next.simps(5) set_subset_Cons)
qed
moreover have "~ set (msgs c' i) \<subseteq> set (msgs c i)" using assms by blast
ultimately show False by simp
qed
lemma new_Marker_in_set_implies_nonregular_occurence:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"Marker \<notin> set (msgs c i)" and
"Marker \<in> set (msgs c' i)" and
"channel i = Some (p, q)"
shows
"~ regular_event ev" (is ?P)
proof (rule ccontr)
have "occurs_on ev = p"
using assms new_msg_in_set_implies_occurrence by blast
assume "~ ?P"
then have "isTrans ev \<or> isSend ev \<or> isRecv ev" by simp
then have "Marker \<notin> set (msgs c' i)"
proof (elim disjE, goal_cases)
case 1
then obtain r u u' where "ev = Trans r u u'"
by (metis event.collapse(1))
then show ?thesis
using assms(1) assms(2) by auto
next
case 2
then obtain i' r q u u' m where "ev = Send i' r q u u' m"
by (metis event.collapse(2))
then show ?thesis
by (metis (no_types, lifting) Un_iff assms(1) assms(2) empty_iff empty_set insert_iff list.set(2) message.distinct(1) next_send set_append)
next
case 3
then obtain i' r q u u' m where "ev = Recv i' r q u u' m"
by (metis event.collapse(3))
then show ?thesis
by (metis assms(1) assms(2) list.set_intros(2) next_recv)
qed
then show False using assms by simp
qed
lemma RecvMarker_implies_Marker_in_set:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"ev = RecvMarker cid p q"
shows
"Marker \<in> set (msgs c cid)"
by (metis (mono_tags, lifting) assms(1) assms(2) can_occur_def distributed_system.happen_implies_can_occur distributed_system_axioms event.simps(30) list.set_sel(1) list.size(3) nat_less_le)
lemma RecvMarker_given_channel:
assumes
"isRecvMarker ev" and
"getId ev = cid" and
"channel cid = Some (p, q)" and
"can_occur ev c"
shows
"ev = RecvMarker cid q p"
by (metis (mono_tags, lifting) assms(1) assms(2) assms(3) assms(4) can_occur_def event.case_eq_if event.collapse(5) event.distinct_disc(8,14,18,20) option.inject prod.inject)
lemma Recv_given_channel:
assumes
"isRecv ev" and
"getId ev = cid" and
"channel cid = Some (p, q)" and
"can_occur ev c"
shows
"\<exists>s s' m. ev = Recv cid q p s s' m"
by (metis assms(1) assms(2) assms(3) assms(4) distributed_system.can_occur_Recv distributed_system_axioms event.collapse(3) option.inject prod.inject)
lemma same_cs_if_not_recv:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"~ isRecv ev"
shows
"fst (cs c cid) = fst (cs c' cid)"
proof (cases "channel cid = None")
case True
then show ?thesis
using assms(1) no_cs_change_if_no_channel by auto
next
case False
then obtain p q where chan: "channel cid = Some (p, q)" by auto
then show ?thesis
proof (cases ev)
case (Snapshot r)
with Snapshot assms chan show ?thesis by (cases "r = q", auto)
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "has_snapshotted c r")
case True
with assms RecvMarker chan show ?thesis by (cases "cid' = cid", auto)
next
case no_snap: False
then show ?thesis
proof (cases "cid' = cid")
case True
then show ?thesis using RecvMarker assms chan by auto
next
case False
with assms RecvMarker chan no_snap show ?thesis by (cases "r = q", auto)
qed
qed
next
case (Trans r u u')
then show ?thesis using assms by auto
next
case (Send r s u u')
then show ?thesis using assms by auto
qed (metis assms(2) isRecv_def)
qed
lemma done_only_from_recv_marker:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"channel cid = Some (p, q)" and
"snd (cs c cid) \<noteq> Done" and
"snd (cs c' cid) = Done"
shows
"ev = RecvMarker cid q p"
proof (rule ccontr)
assume "~ ev = RecvMarker cid q p"
then show False
proof (cases "isRecvMarker ev")
case True
then obtain cid' s r where RecvMarker: "ev = RecvMarker cid' s r" by (meson isRecvMarker_def)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then show False
using \<open>ev = RecvMarker cid' s r\<close> \<open>ev \<noteq> RecvMarker cid q p\<close> assms(1) assms(2) can_occur_def by auto
qed
then have "snd (cs c' cid) \<noteq> Done"
proof (cases "has_snapshotted c s")
case True
- then show ?thesis using RecvMarker assms `cid \<noteq> cid'` by simp
+ then show ?thesis using RecvMarker assms \<open>cid \<noteq> cid'\<close> by simp
next
case False
- with RecvMarker assms `cid \<noteq> cid'` show ?thesis by (cases "s = q", auto)
+ with RecvMarker assms \<open>cid \<noteq> cid'\<close> show ?thesis by (cases "s = q", auto)
qed
then show False using assms by auto
next
case False
then have "isSnapshot ev \<or> isTrans ev \<or> isSend ev \<or> isRecv ev"
using event.exhaust_disc by blast
then have "snd (cs c' cid) \<noteq> Done"
proof (elim disjE, goal_cases)
case 1
then obtain r where Snapshot: "ev = Snapshot r"
by (meson isSnapshot_def)
with assms show ?thesis by (cases "q = r", auto)
next
case 2
then obtain r u u' where "ev = Trans r u u'"
by (meson isTrans_def)
then show ?case using assms by auto
next
case 3
then obtain cid' r s u u' m where "ev = Send cid' r s u u' m"
by (meson isSend_def)
then show ?thesis using assms by auto
next
case 4
then obtain cid' r s u u' m where Recv: "ev = Recv cid' r s u u' m"
by (meson isRecv_def)
show ?thesis
using Recv assms proof (cases "cid = cid'")
case True
then have "snd (cs c cid) = NotStarted \<or> snd (cs c cid) = Recording"
using assms(3) recording_state.exhaust by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have "snd (cs c' cid') = NotStarted"
using True Recv assms(1) by auto
then show ?case using True by auto
next
case 2
then have "snd (cs c' cid') = Recording"
using True Recv assms(1) by auto
then show ?case using True by auto
qed
qed auto
qed
then show False using assms by auto
qed
qed
lemma cs_not_not_started_stable:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"snd (cs c cid) \<noteq> NotStarted" and
"channel cid = Some (p, q)"
shows
"snd (cs c' cid) \<noteq> NotStarted"
using assms proof (cases ev)
case (Snapshot r)
then show ?thesis
by (metis assms(1) assms(2) next_snapshot recording_state.simps(2) sndI)
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "has_snapshotted c r")
case True
with RecvMarker assms show ?thesis by (cases "cid = cid'", auto)
next
case no_snap: False
then show ?thesis
proof (cases "cid = cid'")
case True
then show ?thesis using RecvMarker assms by auto
next
case False
with RecvMarker assms no_snap show ?thesis by (cases "s = p", auto)
qed
qed
next
case (Recv cid' r s u u' m)
then have "snd (cs c cid) = Recording \<or> snd (cs c cid) = Done"
using assms(2) recording_state.exhaust by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then show ?thesis
by (metis (no_types, lifting) Recv assms(1) eq_snd_iff next_recv recording_state.distinct(1))
next
case 2
with Recv assms show ?thesis by (cases "cid = cid'", auto)
qed
qed auto
lemma fst_cs_changed_by_recv_recording:
assumes
step: "c \<turnstile> ev \<mapsto> c'" and
"fst (cs c cid) \<noteq> fst (cs c' cid)" and
"channel cid = Some (p, q)"
shows
"snd (cs c cid) = Recording \<and> (\<exists>p q u u' m. ev = Recv cid q p u u' m)"
proof -
have oc_on: "occurs_on ev = q"
proof -
obtain nn :: "('p, 's, 'm) event \<Rightarrow> nat" and aa :: "('p, 's, 'm) event \<Rightarrow> 'p" and aaa :: "('p, 's, 'm) event \<Rightarrow> 'p" and bb :: "('p, 's, 'm) event \<Rightarrow> 's" and bba :: "('p, 's, 'm) event \<Rightarrow> 's" and cc :: "('p, 's, 'm) event \<Rightarrow> 'm" where
f1: "\<forall>e. (\<not> isRecv e \<or> e = Recv (nn e) (aa e) (aaa e) (bb e) (bba e) (cc e)) \<and> (isRecv e \<or> (\<forall>n a aa b ba c. e \<noteq> Recv n a aa b ba c))"
using isRecv_def by moura
then have f2: "c \<turnstile> Recv (nn ev) (aa ev) (aaa ev) (bb ev) (bba ev) (cc ev) \<mapsto> c'"
by (metis (no_types) assms(2) local.step same_cs_if_not_recv)
have f3: "\<forall>x0 x1 x7 x8. (x0 \<noteq> x7 \<longrightarrow> cs (x8::('p, 's, 'm) configuration) x0 = cs (x1::('p, 's, _) configuration) x0) = (x0 = x7 \<or> cs x8 x0 = cs x1 x0)"
by auto
have f4: "\<forall>x0 x1 x7 x8. (x7 \<noteq> x0 \<longrightarrow> msgs (x1::('p, 's, 'm) configuration) x0 = msgs (x8::('p, 's, _) configuration) x0) = (x7 = x0 \<or> msgs x1 x0 = msgs x8 x0)"
by auto
have "\<forall>x0 x1 x6 x8. (x0 \<noteq> x6 \<longrightarrow> states (x1::('p, 's, 'm) configuration) x0 = states (x8::(_, _, 'm) configuration) x0) = (x0 = x6 \<or> states x1 x0 = states x8 x0)"
by fastforce
then have "can_occur (Recv (nn ev) (aa ev) (aaa ev) (bb ev) (bba ev) (cc ev)) c \<and> states c (aa ev) = bb ev \<and> states c' (aa ev) = bba ev \<and> (\<forall>a. a = aa ev \<or> states c' a = states c a) \<and> msgs c (nn ev) = Msg (cc ev) # msgs c' (nn ev) \<and> (\<forall>n. nn ev = n \<or> msgs c' n = msgs c n) \<and> (\<forall>a. ps c a = ps c' a) \<and> (\<forall>n. n = nn ev \<or> cs c n = cs c' n) \<and> (if snd (cs c (nn ev)) = Recording then cs c' (nn ev) = (fst (cs c (nn ev)) @ [cc ev], Recording) else cs c (nn ev) = cs c' (nn ev))"
using f4 f3 f2 by force
then show ?thesis
using f1 by (metis (no_types) Pair_inject assms(2) assms(3) can_occur_Recv event.sel(3) local.step option.sel same_cs_if_not_recv)
qed
have "isRecv ev" (is ?P)
proof (rule ccontr)
assume "~ ?P"
then have "fst (cs c cid) = fst (cs c' cid)" by (metis local.step same_cs_if_not_recv)
then show False using assms by simp
qed
then obtain cid' r s u u' m where Recv: "ev = Recv cid' r s u u' m" by (meson isRecv_def)
have "cid = cid'"
proof (rule ccontr)
assume "~ cid = cid'"
then have "fst (cs c cid) = fst (cs c' cid)" using Recv step by auto
then show False using assms by simp
qed
moreover have "snd (cs c cid) = Recording"
proof (rule ccontr)
assume "~ snd (cs c cid) = Recording"
- then have "fst (cs c cid) = fst (cs c' cid)" using Recv step `cid = cid'` by auto
+ then have "fst (cs c cid) = fst (cs c' cid)" using Recv step \<open>cid = cid'\<close> by auto
then show False using assms by simp
qed
ultimately show ?thesis using Recv by simp
qed
lemma no_marker_and_snapshotted_implies_no_more_markers:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"has_snapshotted c p" and
"Marker \<notin> set (msgs c cid)" and
"channel cid = Some (p, q)"
shows
"Marker \<notin> set (msgs c' cid)"
proof (cases ev)
case (Snapshot r)
then have "r \<noteq> p"
using assms(1) assms(2) can_occur_def by auto
then have "msgs c cid = msgs c' cid" using assms Snapshot by simp
then show ?thesis using assms by simp
next
case (RecvMarker cid' r s)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
moreover have "can_occur ev c" using happen_implies_can_occur assms by blast
ultimately have "Marker : set (msgs c cid)" using can_occur_def RecvMarker
by (metis (mono_tags, lifting) assms(1) event.simps(30) hd_in_set list.size(3) recv_marker_other_channels_not_shrinking zero_order(1))
then show False using assms by simp
qed
then have "msgs c cid = msgs c' cid"
proof (cases "r = p")
case True
then show ?thesis
using RecvMarker \<open>cid \<noteq> cid'\<close> assms(1) assms(2) msgs_unchanged_if_snapshotted_RecvMarker_for_other_is by blast
next
case False
- with RecvMarker `cid \<noteq> cid'` step assms show ?thesis by (cases "has_snapshotted c r", auto)
+ with RecvMarker \<open>cid \<noteq> cid'\<close> step assms show ?thesis by (cases "has_snapshotted c r", auto)
qed
then show ?thesis using assms by simp
next
case (Trans r u u')
then show ?thesis using assms by auto
next
case (Send cid' r s u u' m)
with assms Send show ?thesis by (cases "cid = cid'", auto)
next
case (Recv cid' r s u u' m)
with assms Recv show ?thesis by (cases "cid = cid'", auto)
qed
lemma same_messages_if_no_occurrence:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"~ occurs_on ev = p" and
"~ occurs_on ev = q" and
"channel cid = Some (p, q)"
shows
"msgs c cid = msgs c' cid \<and> cs c cid = cs c' cid"
proof (cases ev)
case (Snapshot r)
then show ?thesis using assms by auto
next
case (RecvMarker cid' r s)
have "cid \<noteq> cid'"
by (metis RecvMarker_given_channel assms(1) assms(3) assms(4) RecvMarker event.sel(5,10) happen_implies_can_occur isRecvMarker_def)
have "\<nexists>a. channel cid = Some (r, q)"
using assms(2) assms(4) RecvMarker by auto
- with RecvMarker assms `cid \<noteq> cid'` show ?thesis by (cases "has_snapshotted c r", auto)
+ with RecvMarker assms \<open>cid \<noteq> cid'\<close> show ?thesis by (cases "has_snapshotted c r", auto)
next
case (Trans r u u')
then show ?thesis using assms by auto
next
case (Send cid' r s u u' m)
then have "cid \<noteq> cid'"
by (metis (mono_tags, lifting) Pair_inject assms(1) assms(2) assms(4) can_occur_def event.sel(2) event.simps(27) happen_implies_can_occur option.inject)
then show ?thesis using assms Send by simp
next
case (Recv cid' r s u u' m)
then have "cid \<noteq> cid'"
by (metis assms(1) assms(3) assms(4) distributed_system.can_occur_Recv distributed_system.happen_implies_can_occur distributed_system_axioms event.sel(3) option.inject prod.inject)
then show ?thesis using assms Recv by simp
qed
end (* locale distributed_system *)
end (* theory Distributed_System *)
diff --git a/thys/Chandy_Lamport/Example.thy b/thys/Chandy_Lamport/Example.thy
--- a/thys/Chandy_Lamport/Example.thy
+++ b/thys/Chandy_Lamport/Example.thy
@@ -1,274 +1,274 @@
section \<open>Example\<close>
text \<open>We provide an example in order to prove that our locale is non-vacuous.
This example corresponds to the computation and associated snapshot described
in Section 4 of~\cite{chandy}.\<close>
theory Example
imports
Snapshot
begin
datatype PType = P | Q
datatype MType = M | M'
datatype SType = S_Wait | S_Send | T_Wait | T_Send
fun trans :: "PType \<Rightarrow> SType \<Rightarrow> SType \<Rightarrow> bool" where
"trans p s s' = False"
fun send :: "channel_id \<Rightarrow> PType \<Rightarrow> PType \<Rightarrow> SType
\<Rightarrow> SType \<Rightarrow> MType \<Rightarrow> bool" where
"send c p q s s' m = ((c = 0 \<and> p = P \<and> q = Q
\<and> s = S_Send \<and> s' = S_Wait \<and> m = M)
\<or> (c = 1 \<and> p = Q \<and> q = P
\<and> s = T_Send \<and> s' = T_Wait \<and> m = M'))"
fun recv :: "channel_id \<Rightarrow> PType \<Rightarrow> PType \<Rightarrow> SType
\<Rightarrow> SType \<Rightarrow> MType \<Rightarrow> bool" where
"recv c p q s s' m = ((c = 1 \<and> p = P \<and> q = Q
\<and> s = S_Wait \<and> s' = S_Send \<and> m = M')
\<or> (c = 0 \<and> p = Q \<and> q = P
\<and> s = T_Wait \<and> s' = T_Send \<and> m = M))"
fun chan :: "nat \<Rightarrow> (PType * PType) option" where
"chan n = (if n = 0 then Some (P, Q)
else if n = 1 then Some (Q, P)
else None)"
abbreviation init :: "(PType, SType, MType) configuration" where
"init \<equiv> \<lparr>
states = (%p. if p = P then S_Send else T_Send),
msgs = (%d. []),
process_snapshot = (%p. None),
channel_snapshot = (%d. ([], NotStarted))
\<rparr>"
abbreviation t0 where "t0 \<equiv> Snapshot P"
abbreviation s1 :: "(PType, SType, MType) configuration" where
"s1 \<equiv> \<lparr>
states = (%p. if p = P then S_Send else T_Send),
msgs = (%d. if d = 0 then [Marker] else []),
process_snapshot = (%p. if p = P then Some S_Send else None),
channel_snapshot = (%d. if d = 1 then ([], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t1 where "t1 \<equiv> Send 0 P Q S_Send S_Wait M"
abbreviation s2 :: "(PType, SType, MType) configuration" where
"s2 \<equiv> \<lparr>
states = (%p. if p = P then S_Wait else T_Send),
msgs = (%d. if d = 0 then [Marker, Msg M] else []),
process_snapshot = (%p. if p = P then Some S_Send else None),
channel_snapshot = (%d. if d = 1 then ([], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t2 where "t2 \<equiv> Send 1 Q P T_Send T_Wait M'"
abbreviation s3 :: "(PType, SType, MType) configuration" where
"s3 \<equiv> \<lparr>
states = (%p. if p = P then S_Wait else T_Wait),
msgs = (%d. if d = 0 then [Marker, Msg M] else if d = 1 then [Msg M'] else []),
process_snapshot = (%p. if p = P then Some S_Send else None),
channel_snapshot = (%d. if d = 1 then ([], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t3 where "t3 \<equiv> Snapshot Q"
abbreviation s4 :: "(PType, SType, MType) configuration" where
"s4 \<equiv> \<lparr>
states = (%p. if p = P then S_Wait else T_Wait),
msgs = (%d. if d = 0 then [Marker, Msg M] else if d = 1 then [Msg M', Marker] else []),
process_snapshot = (%p. if p = P then Some S_Send else Some T_Wait),
channel_snapshot = (%d. if d = 1 then ([], Recording) else if d = 0 then ([], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t4 where "t4 \<equiv> RecvMarker 0 Q P"
abbreviation s5 :: "(PType, SType, MType) configuration" where
"s5 \<equiv> \<lparr>
states = (%p. if p = P then S_Wait else T_Wait),
msgs = (%d. if d = 0 then [Msg M] else if d = 1 then [Msg M', Marker] else []),
process_snapshot = (%p. if p = P then Some S_Send else Some T_Wait),
channel_snapshot = (%d. if d = 0 then ([], Done) else if d = 1 then ([], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t5 where "t5 \<equiv> Recv 1 P Q S_Wait S_Send M'"
abbreviation s6 :: "(PType, SType, MType) configuration" where
"s6 \<equiv> \<lparr>
states = (%p. if p = P then S_Send else T_Wait),
msgs = (%d. if d = 0 then [Msg M] else if d = 1 then [Marker] else []),
process_snapshot = (%p. if p = P then Some S_Send else Some T_Wait),
channel_snapshot = (%d. if d = 0 then ([], Done) else if d = 1 then ([M'], Recording) else ([], NotStarted))
\<rparr>"
abbreviation t6 where "t6 \<equiv> RecvMarker 1 P Q"
abbreviation s7 :: "(PType, SType, MType) configuration" where
"s7 \<equiv> \<lparr>
states = (%p. if p = P then S_Send else T_Wait),
msgs = (%d. if d = 0 then [Msg M] else if d = 1 then [] else []),
process_snapshot = (%p. if p = P then Some S_Send else Some T_Wait),
channel_snapshot = (%d. if d = 0 then ([], Done) else if d = 1 then ([M'], Done) else ([], NotStarted))
\<rparr>"
lemma s7_no_marker:
shows
"\<forall>cid. Marker \<notin> set (msgs s7 cid)"
by simp
interpretation computation chan trans send recv init s7
proof
have "distributed_system chan"
proof
show "\<forall>i. \<nexists>p. chan i = Some (p, p)" by simp
qed
show "\<forall>p q. p \<noteq> q \<longrightarrow> (\<lambda>p q. \<exists>i. chan i = Some (p, q))\<^sup>+\<^sup>+ p q"
proof ((rule allI)+, rule impI)
fix p q :: PType assume "p \<noteq> q"
then have "(p = P \<and> q = Q) \<or> (p = Q \<and> q = P)"
using PType.exhaust by auto
then have "\<exists>i. chan i = Some (p, q)" by (elim disjE) auto
then show "(\<lambda>p q. \<exists>i. chan i = Some (p, q))\<^sup>+\<^sup>+ p q" by blast
qed
show "finite {i. \<exists>p q. chan i = Some (p, q)}"
proof -
have "{i. \<exists>p q. chan i = Some (p, q)} = {0,1}" by auto
then show ?thesis by simp
qed
show "1 < card (UNIV :: PType set)"
proof -
have "(UNIV :: PType set) = {P, Q}"
using PType.exhaust by blast
then have "card (UNIV :: PType set) = 2"
by (metis One_nat_def PType.distinct(1) Suc_1 card.insert card.empty finite.emptyI finite.insertI insert_absorb insert_not_empty singletonD)
then show ?thesis by auto
qed
show "finite (UNIV :: PType set)"
proof -
have "(UNIV :: PType set) = {P, Q}"
using PType.exhaust by blast
then show ?thesis
by (metis finite.emptyI finite.insertI)
qed
show "\<forall>i. \<nexists>p. chan i = Some (p, p)" by simp
show "\<forall>i. (\<exists>p q. chan i = Some (p, q)) \<longrightarrow> Marker \<notin> set (msgs init i)" by auto
show "\<forall>i. chan i = None \<longrightarrow> msgs init i = []" by auto
show "\<forall>p. \<not> ps init p \<noteq> None" by auto
show "\<forall>i. cs init i = ([], NotStarted)" by auto
show "\<exists>t. distributed_system.trace chan Example.trans send recv init t s7"
proof -
let ?t = "[t0, t1, t2, t3, t4, t5, t6]"
have "distributed_system.next chan trans send recv init t0 s1"
proof -
have "distributed_system.can_occur chan trans send recv t0 init"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
then show ?thesis
by (simp add: \<open>distributed_system chan\<close> distributed_system.next_snapshot)
qed
moreover have "distributed_system.next chan trans send recv s1 t1 s2"
proof -
have "distributed_system.can_occur chan trans send recv t1 s1"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
then show ?thesis
by (simp add: \<open>distributed_system chan\<close> distributed_system.next_send)
qed
moreover have "distributed_system.next chan trans send recv s2 t2 s3"
proof -
have "distributed_system.can_occur chan trans send recv t2 s2"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
moreover have "\<forall>r. r \<noteq> P \<longrightarrow> r = Q" using PType.exhaust by auto
- ultimately show ?thesis by (simp add: `distributed_system chan` distributed_system.next_send)
+ ultimately show ?thesis by (simp add: \<open>distributed_system chan\<close> distributed_system.next_send)
qed
moreover have "distributed_system.next chan trans send recv s3 t3 s4"
proof -
have "distributed_system.can_occur chan trans send recv t3 s3"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
moreover have "\<forall>p'. p' \<noteq> P \<longrightarrow> p' = Q" using PType.exhaust by auto
- ultimately show ?thesis by (simp add: `distributed_system chan` distributed_system.next_snapshot)
+ ultimately show ?thesis by (simp add: \<open>distributed_system chan\<close> distributed_system.next_snapshot)
qed
moreover have "distributed_system.next chan trans send recv s4 t4 s5"
proof -
have "distributed_system.can_occur chan trans send recv t4 s4"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
then show ?thesis
- by (simp add: `distributed_system chan` distributed_system.next_def)
+ by (simp add: \<open>distributed_system chan\<close> distributed_system.next_def)
qed
moreover have "distributed_system.next chan trans send recv s5 t5 s6"
proof -
have "distributed_system.can_occur chan trans send recv t5 s5"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
then show ?thesis
- by (simp add: `distributed_system chan` distributed_system.next_def)
+ by (simp add: \<open>distributed_system chan\<close> distributed_system.next_def)
qed
moreover have "distributed_system.next chan trans send recv s6 t6 s7"
proof -
have "distributed_system.can_occur chan trans send recv t6 s6"
using \<open>distributed_system chan\<close> distributed_system.can_occur_def by fastforce
then show ?thesis
- by (simp add: `distributed_system chan` distributed_system.next_def)
+ by (simp add: \<open>distributed_system chan\<close> distributed_system.next_def)
qed
ultimately have "distributed_system.trace chan trans send recv init ?t s7"
by (meson \<open>distributed_system chan\<close> distributed_system.trace.simps)
then show ?thesis by blast
qed
show "\<forall>t i cid. distributed_system.trace chan Example.trans send recv init t s7 \<and>
Marker \<in> set (msgs (distributed_system.s chan Example.trans send recv init t i) cid) \<longrightarrow>
(\<exists>j\<ge>i. Marker \<notin> set (msgs (distributed_system.s chan Example.trans send recv init t j) cid))"
proof ((rule allI)+, (rule impI)+)
fix t i cid
assume asm: "distributed_system.trace chan Example.trans send recv init t s7 \<and>
Marker \<in> set (msgs (distributed_system.s chan Example.trans send recv init t i) cid)"
have tr_exists: "distributed_system.trace chan Example.trans send recv init t s7" using asm by blast
have marker_in_channel: "Marker \<in> set (msgs (distributed_system.s chan Example.trans send recv init t i) cid)" using asm by simp
have s7_is_fin: "s7 = (distributed_system.s chan Example.trans send recv init t (length t))"
by (metis (no_types, lifting) \<open>distributed_system chan\<close> \<open>distributed_system.trace chan Example.trans send recv init t s7\<close> distributed_system.exists_trace_for_any_i distributed_system.trace_and_start_determines_end order_refl take_all)
have "i < length t"
proof (rule ccontr)
assume "~ i < length t"
then have "distributed_system.trace chan Example.trans send recv
(distributed_system.s chan Example.trans send recv init t (length t))
[]
(distributed_system.s chan Example.trans send recv init t i)"
by (metis (no_types, lifting) \<open>distributed_system chan\<close> distributed_system.exists_trace_for_any_i distributed_system.trace.simps distributed_system.trace_and_start_determines_end not_less s7_is_fin take_all tr_exists)
then have "Marker \<notin> set (msgs (distributed_system.s chan Example.trans send recv init t i) cid)"
proof -
have "distributed_system.s chan Example.trans send recv init t i = s7"
using \<open>distributed_system chan\<close> \<open>distributed_system.trace chan Example.trans send recv (distributed_system.s chan Example.trans send recv init t (length t)) [] (distributed_system.s chan Example.trans send recv init t i)\<close> distributed_system.trace.simps s7_is_fin by fastforce
then show ?thesis using s7_no_marker by simp
qed
then show False using marker_in_channel by simp
qed
then show "(\<exists>j\<ge>i. Marker \<notin> set (msgs (distributed_system.s chan Example.trans send recv init t j) cid))"
proof -
have "distributed_system.trace chan Example.trans send recv
(distributed_system.s chan Example.trans send recv init t i)
(take ((length t) - i) (drop i t))
(distributed_system.s chan Example.trans send recv init t (length t))"
using \<open>distributed_system chan\<close> \<open>i < length t\<close> distributed_system.exists_trace_for_any_i_j less_imp_le_nat tr_exists by blast
then have "Marker \<notin> set (msgs (distributed_system.s chan Example.trans send recv init t (length t)) cid)"
proof -
have "distributed_system.s chan Example.trans send recv init t (length t) = s7"
by (simp add: s7_is_fin)
then show ?thesis using s7_no_marker by simp
qed
then show ?thesis
using \<open>i < length t\<close> less_imp_le_nat by blast
qed
qed
show "\<forall>t p. distributed_system.trace chan Example.trans send recv init t s7 \<longrightarrow>
(\<exists>i. ps (distributed_system.s chan Example.trans send recv init t i) p \<noteq> None \<and> i \<le> length t)"
proof ((rule allI)+, rule impI)
fix t p
assume "distributed_system.trace chan Example.trans send recv init t s7"
have s7_is_fin: "s7 = (distributed_system.s chan Example.trans send recv init t (length t))"
by (metis (no_types, lifting) \<open>distributed_system chan\<close> \<open>distributed_system.trace chan Example.trans send recv init t s7\<close> distributed_system.exists_trace_for_any_i distributed_system.trace_and_start_determines_end order_refl take_all)
moreover have "has_snapshotted s7 p" by simp
ultimately show "(\<exists>i. ps (distributed_system.s chan Example.trans send recv init t i) p \<noteq> None \<and> i \<le> length t)"
by auto
qed
qed
end
diff --git a/thys/Chandy_Lamport/Snapshot.thy b/thys/Chandy_Lamport/Snapshot.thy
--- a/thys/Chandy_Lamport/Snapshot.thy
+++ b/thys/Chandy_Lamport/Snapshot.thy
@@ -1,5272 +1,5272 @@
section \<open>The Chandy--Lamport algorithm\<close>
theory Snapshot
imports
"HOL-Library.Sublist"
"HOL-Library.Permutation"
Distributed_System
Trace
Util
Swap
begin
subsection \<open>The computation locale\<close>
text \<open>We extend the distributed system locale presented
earlier: Now we are given a trace t of the distributed system between
two configurations, the initial and final configuartions of t. Our objective
is to show that the Chandy--Lamport algorithm terminated successfully and
exhibits the same properties as claimed in~\cite{chandy}. In the initial state
no snapshotting must have taken place yet, however the computation itself may
have progressed arbitrarily far already.
We assume that there exists at least one process, that the
total number of processes in the system is finite, and that there
are only finitely many channels between the processes. The process graph
is strongly connected. Finally there are Chandy and Lamport's core assumptions:
every process snapshots at some time and no marker may remain in a channel forever.\<close>
locale computation = distributed_system +
fixes
init final :: "('a, 'b, 'c) configuration"
assumes
finite_channels:
"finite {i. \<exists>p q. channel i = Some (p, q)}" and
strongly_connected_raw:
"\<forall>p q. (p \<noteq> q) \<longrightarrow>
(tranclp (\<lambda>p q. (\<exists>i. channel i = Some (p, q)))) p q" and
at_least_two_processes:
"card (UNIV :: 'a set) > 1" and
finite_processes:
"finite (UNIV :: 'a set)" and
no_initial_Marker:
"\<forall>i. (\<exists>p q. channel i = Some (p, q))
\<longrightarrow> Marker \<notin> set (msgs init i)" and
no_msgs_if_no_channel:
"\<forall>i. channel i = None \<longrightarrow> msgs init i = []" and
no_initial_process_snapshot:
"\<forall>p. ~ has_snapshotted init p" and
no_initial_channel_snapshot:
"\<forall>i. channel_snapshot init i = ([], NotStarted)" and
valid: "\<exists>t. trace init t final" and
l1: "\<forall>t i cid. trace init t final
\<and> Marker \<in> set (msgs (s init t i) cid)
\<longrightarrow> (\<exists>j. j \<ge> i \<and> Marker \<notin> set (msgs (s init t j) cid))" and
l2: "\<forall>t p. trace init t final
\<longrightarrow> (\<exists>i. has_snapshotted (s init t i) p \<and> i \<le> length t)"
begin
definition has_channel where
"has_channel p q \<longleftrightarrow> (\<exists>i. channel i = Some (p, q))"
lemmas strongly_connected = strongly_connected_raw[folded has_channel_def]
lemma exists_some_channel:
shows "\<exists>i p q. channel i = Some (p, q)"
proof -
obtain p q where "p : (UNIV :: 'a set) \<and> q : (UNIV :: 'a set) \<and> p \<noteq> q"
by (metis (mono_tags) One_nat_def UNIV_eq_I all_not_in_conv at_least_two_processes card_Suc_Diff1 card.empty finite_processes insert_iff iso_tuple_UNIV_I less_numeral_extra(4) n_not_Suc_n)
then have "(tranclp has_channel) p q" using strongly_connected by simp
then obtain r s where "has_channel r s"
by (meson tranclpD)
then show ?thesis using has_channel_def by auto
qed
abbreviation S where
"S \<equiv> s init"
lemma no_messages_if_no_channel:
assumes "trace init t final"
shows "channel cid = None \<Longrightarrow> msgs (s init t i) cid = []"
using no_messages_introduced_if_no_channel[OF assms no_msgs_if_no_channel] by blast
lemma S_induct [consumes 3, case_names S_init S_step]:
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
\<And>i. P i i;
\<And>i j. i < j \<Longrightarrow> j \<le> length t \<Longrightarrow> (S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i)) \<Longrightarrow> P (Suc i) j \<Longrightarrow> P i j
\<rbrakk> \<Longrightarrow> P i j"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by simp
next
case (Suc n)
then have "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i))" using Suc step_Suc by simp
then show ?case using Suc by simp
qed
lemma exists_index:
assumes
"trace init t final" and
"ev \<in> set (take (j - i) (drop i t))"
shows
"\<exists>k. i \<le> k \<and> k < j \<and> ev = t ! k"
proof -
have "trace (S t i) (take (j - i) (drop i t)) (S t j)"
by (metis assms(1) assms(2) diff_is_0_eq' exists_trace_for_any_i_j list.distinct(1) list.set_cases nat_le_linear take_eq_Nil)
obtain l where "ev = (take (j - i) (drop i t)) ! l" "l < length (take (j - i) (drop i t))"
by (metis assms(2) in_set_conv_nth)
let ?k = "l + i"
have "(take (j - i) (drop i t)) ! l = drop i t ! l"
using \<open>l < length (take (j - i) (drop i t))\<close> by auto
also have "... = t ! ?k"
by (metis add.commute assms(2) drop_all empty_iff list.set(1) nat_le_linear nth_drop take_Nil)
finally have "ev = t ! ?k"
using \<open>ev = take (j - i) (drop i t) ! l\<close> by blast
moreover have "i \<le> ?k \<and> ?k < j"
using \<open>l < length (take (j - i) (drop i t))\<close> by auto
ultimately show ?thesis by blast
qed
lemma no_change_if_ge_length_t:
assumes
"trace init t final" and
"i \<ge> length t" and
"j \<ge> i"
shows
"S t i = S t j"
proof -
have "trace (S t i) (take (j - i) (drop i t)) (S t j)"
using assms(1) assms(3) exists_trace_for_any_i_j by blast
moreover have "(take (j - i) (drop i t)) = Nil"
by (simp add: assms(2))
ultimately show ?thesis
by (metis tr_init trace_and_start_determines_end)
qed
lemma no_marker_if_no_snapshot:
shows
"\<lbrakk> trace init t final; channel cid = Some (p, q);
~ has_snapshotted (S t i) p \<rbrakk>
\<Longrightarrow> Marker \<notin> set (msgs (S t i) cid)"
proof (induct i)
case 0
then show ?case
by (metis exists_trace_for_any_i no_initial_Marker take_eq_Nil tr_init trace_and_start_determines_end)
next
case (Suc n)
then have IH: "Marker \<notin> set (msgs (S t n) cid)"
by (meson distributed_system.exists_trace_for_any_i_j distributed_system.snapshot_stable_2 distributed_system_axioms eq_iff le_Suc_eq)
then obtain tr where decomp: "trace (S t n) tr (S t (Suc n))" "tr = take (Suc n - n) (drop n t)"
using Suc exists_trace_for_any_i_j le_Suc_eq by blast
have "Marker \<notin> set (msgs (S t (Suc n)) cid)"
proof (cases "tr = []")
case True
then show ?thesis
by (metis IH decomp(1) tr_init trace_and_start_determines_end)
next
case False
then obtain ev where step: "tr = [ev]" "(S t n) \<turnstile> ev \<mapsto> (S t (Suc n))"
by (metis One_nat_def Suc_eq_plus1 Suc_leI \<open>tr = take (Suc n - n) (drop n t)\<close> \<open>trace (S t n) tr (S t (Suc n))\<close> add_diff_cancel_left' append.simps(1) butlast_take cancel_comm_monoid_add_class.diff_cancel length_greater_0_conv list.distinct(1) list.sel(3) snoc_eq_iff_butlast take0 take_Nil trace.cases)
then show ?thesis
proof (cases ev)
case (Snapshot p')
then show ?thesis
by (metis IH Suc.prems(2) Suc.prems(3) local.step(2) new_Marker_in_set_implies_nonregular_occurence new_msg_in_set_implies_occurrence nonregular_event_induces_snapshot snapshot_state_unchanged)
next
case (RecvMarker cid' p' q')
have "p' \<noteq> p"
proof (rule ccontr)
assume asm: "~ p' \<noteq> p"
then have "has_snapshotted (S t (Suc n)) p"
proof -
have "~ regular_event ev" using RecvMarker by auto
moreover have "occurs_on ev = p" using asm RecvMarker by auto
ultimately show ?thesis using step(2) Suc.hyps Suc.prems
by (metis nonregular_event_induces_snapshot snapshot_state_unchanged)
qed
then show False using Suc.prems by blast
qed
moreover have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "hd (msgs (S t n) cid) = Marker \<and> length (msgs (S t n) cid) > 0"
using step RecvMarker can_occur_def by auto
then have "Marker : set (msgs (S t n) cid)"
using list.set_sel(1) by fastforce
then show False using IH by simp
qed
ultimately have "msgs (S t (Suc n)) cid = msgs (S t n) cid"
proof -
have "\<nexists>r. channel cid = Some (p', r)"
using Suc.prems(2) \<open>p' \<noteq> p\<close> by auto
- with `cid \<noteq> cid'` RecvMarker step show ?thesis by (cases "has_snapshotted (S t n) p'", auto)
+ with \<open>cid \<noteq> cid'\<close> RecvMarker step show ?thesis by (cases "has_snapshotted (S t n) p'", auto)
qed
then show ?thesis by (simp add: IH)
next
case (Trans p' s s')
then show ?thesis
using IH local.step(2) by force
next
case (Send cid' p' q' s s' m)
with step IH show ?thesis by (cases "cid' = cid", auto)
next
case (Recv cid' p' q' s s' m)
with step IH show ?thesis by (cases "cid' = cid", auto)
qed
qed
then show ?case by blast
qed
subsection \<open>Termination\<close>
text \<open>We prove that the snapshot algorithm terminates, as exhibited
by lemma \texttt{snapshot\_algorithm\_must\_terminate}. In the final configuration all
processes have snapshotted, and no markers remain in the channels.\<close>
lemma must_exist_snapshot:
assumes
"trace init t final"
shows
"\<exists>p i. Snapshot p = t ! i"
proof (rule ccontr)
assume "\<nexists>p i. Snapshot p = t ! i"
have "\<forall>i p. ~ has_snapshotted (S t i) p"
proof (rule allI)
fix i
show "\<forall>p. ~ has_snapshotted (S t i) p"
proof (induct i)
case 0
then show ?case
by (metis assms distributed_system.trace_and_start_determines_end distributed_system_axioms exists_trace_for_any_i computation.no_initial_process_snapshot computation_axioms take0 tr_init)
next
case (Suc n)
then have IH: "\<forall>p. ~ has_snapshotted (S t n) p" by auto
then obtain tr where "trace (S t n) tr (S t (Suc n))" "tr = take (Suc n - n) (drop n t)"
using assms exists_trace_for_any_i_j le_Suc_eq by blast
show "\<forall>p. ~ has_snapshotted (S t (Suc n)) p"
proof (cases "tr = []")
case True
then show ?thesis
by (metis IH \<open>trace (S t n) tr (S t (Suc n))\<close> tr_init trace_and_start_determines_end)
next
case False
then obtain ev where step: "tr = [ev]" "(S t n) \<turnstile> ev \<mapsto> (S t (Suc n))"
by (metis One_nat_def Suc_eq_plus1 Suc_leI \<open>tr = take (Suc n - n) (drop n t)\<close> \<open>trace (S t n) tr (S t (Suc n))\<close> add_diff_cancel_left' append.simps(1) butlast_take cancel_comm_monoid_add_class.diff_cancel length_greater_0_conv list.distinct(1) list.sel(3) snoc_eq_iff_butlast take0 take_Nil trace.cases)
then show ?thesis
using step Suc.hyps proof (cases ev)
case (Snapshot q)
then show ?thesis
by (metis \<open>\<nexists>p i. Snapshot p = t ! i\<close> \<open>tr = [ev]\<close> \<open>tr = take (Suc n - n) (drop n t)\<close> append_Cons append_take_drop_id nth_append_length)
next
case (RecvMarker cid' q r)
then have m: "Marker \<in> set (msgs (S t n) cid')"
using RecvMarker_implies_Marker_in_set step by blast
have "~ has_snapshotted (S t n) q" using Suc by auto
then have "Marker \<notin> set (msgs (S t n) cid')"
proof -
have "channel cid' = Some (r, q)" using step can_occur_def RecvMarker by auto
then show ?thesis
using IH assms no_marker_if_no_snapshot by blast
qed
then show ?thesis using m by auto
qed auto
qed
qed
qed
obtain j p where "has_snapshotted (S t j) p" using l2 assms by blast
then show False
using \<open>\<forall>i p. \<not> has_snapshotted (S t i) p\<close> by blast
qed
lemma recv_marker_means_snapshotted:
assumes
"trace init t final" and
"ev = RecvMarker cid p q" and
"(S t i) \<turnstile> ev \<mapsto> (S t (Suc i))"
shows
"has_snapshotted (S t i) q"
proof -
have "Marker = hd (msgs (S t i) cid) \<and> length (msgs (S t i) cid) > 0"
proof -
have "Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid"
using assms(2) assms(3) next_recv_marker by blast
then show ?thesis
by (metis length_greater_0_conv list.discI list.sel(1))
qed
then have "Marker \<in> set (msgs (S t i) cid)"
using hd_in_set by fastforce
then show "has_snapshotted (S t i) q"
proof -
have "channel cid = Some (q, p)" using assms can_occur_def by auto
then show ?thesis
using \<open>Marker \<in> set (msgs (S t i) cid)\<close> assms(1) no_marker_if_no_snapshot by blast
qed
qed
lemma recv_marker_means_cs_Done:
assumes
"trace init t final" and
"t ! i = RecvMarker cid p q" and
"i < length t"
shows
"snd (cs (S t (i+1)) cid) = Done"
proof -
have "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
using assms(1) assms(3) step_Suc by auto
then show ?thesis
by (simp add: assms(2))
qed
lemma snapshot_produces_marker:
assumes
"trace init t final" and
"~ has_snapshotted (S t i) p" and
"has_snapshotted (S t (Suc i)) p" and
"channel cid = Some (p, q)"
shows
"Marker : set (msgs (S t (Suc i)) cid) \<or> has_snapshotted (S t i) q"
proof -
obtain ev where ex_ev: "(S t i) \<turnstile> ev \<mapsto> (S t (Suc i))"
by (metis append_Nil2 append_take_drop_id assms(1) assms(2) assms(3) distributed_system.step_Suc distributed_system_axioms drop_eq_Nil less_Suc_eq_le nat_le_linear not_less_eq s_def)
then have "occurs_on ev = p"
using assms(2) assms(3) no_state_change_if_no_event by force
then show ?thesis
using assms ex_ev proof (cases ev)
case (Snapshot r)
then have "Marker \<in> set (msgs (S t (Suc i)) cid)"
using ex_ev assms(2) assms(3) assms(4) by fastforce
then show ?thesis by simp
next
case (RecvMarker cid' r s)
- have "r = p" using `occurs_on ev = p`
+ have "r = p" using \<open>occurs_on ev = p\<close>
by (simp add: RecvMarker)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "has_snapshotted (S t i) q"
using RecvMarker RecvMarker_implies_Marker_in_set assms(1) assms(2) assms(4) ex_ev no_marker_if_no_snapshot by blast
then show ?thesis by simp
next
case False
- then have "\<exists>s. channel cid = Some (r, s)" using RecvMarker assms can_occur_def `r = p` by simp
+ then have "\<exists>s. channel cid = Some (r, s)" using RecvMarker assms can_occur_def \<open>r = p\<close> by simp
then have "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Marker]"
- using RecvMarker assms ex_ev `r = p` False by simp
+ using RecvMarker assms ex_ev \<open>r = p\<close> False by simp
then show ?thesis by simp
qed
qed auto
qed
lemma exists_snapshot_for_all_p:
assumes
"trace init t final"
shows
"\<exists>i. ~ has_snapshotted (S t i) p \<and> has_snapshotted (S t (Suc i)) p" (is ?Q)
proof -
obtain i where "has_snapshotted (S t i) p" using l2 assms by blast
let ?j = "LEAST j. has_snapshotted (S t j) p"
have "?j \<noteq> 0"
proof -
have "~ has_snapshotted (S t 0) p"
by (metis exists_trace_for_any_i list.discI no_initial_process_snapshot s_def take_eq_Nil trace.simps)
then show ?thesis
by (metis (mono_tags, lifting) \<open>has_snapshotted (S t i) p\<close> wellorder_Least_lemma(1))
qed
have "?j \<le> i"
by (meson Least_le \<open>has_snapshotted (S t i) p\<close>)
have "\<not> has_snapshotted (S t (?j - 1)) p" (is ?P)
proof (rule ccontr)
assume "\<not> ?P"
then have "has_snapshotted (S t (?j - 1)) p" by simp
then have "\<exists>j. j < ?j \<and> has_snapshotted (S t j) p"
by (metis One_nat_def \<open>(LEAST j. ps (S t j) p \<noteq> None) \<noteq> 0\<close> diff_less lessI not_gr_zero)
then show False
using not_less_Least by blast
qed
show ?thesis
proof (rule ccontr)
assume "\<not> ?Q"
have "\<forall>i. \<not> has_snapshotted (S t i) p"
proof (rule allI)
fix i'
show "\<not> has_snapshotted (S t i') p"
proof (induct i')
case 0
then show ?case
using \<open>(LEAST j. ps (S t j) p \<noteq> None) \<noteq> 0\<close> by force
next
case (Suc i'')
then show ?case
using \<open>\<nexists>i. \<not> ps (S t i) p \<noteq> None \<and> ps (S t (Suc i)) p \<noteq> None\<close> by blast
qed
qed
then show False
using \<open>ps (S t i) p \<noteq> None\<close> by blast
qed
qed
lemma all_processes_snapshotted_in_final_state:
assumes
"trace init t final"
shows
"has_snapshotted final p"
proof -
obtain i where "has_snapshotted (S t i) p \<and> i \<le> length t"
using assms l2 by blast
moreover have "final = (S t (length t))"
by (metis (no_types, lifting) assms exists_trace_for_any_i le_Suc_eq length_Cons take_Nil take_all trace.simps trace_and_start_determines_end)
ultimately show ?thesis
using assms exists_trace_for_any_i_j snapshot_stable by blast
qed
definition next_marker_free_state where
"next_marker_free_state t i cid = (LEAST j. j \<ge> i \<and> Marker \<notin> set (msgs (S t j) cid))"
lemma exists_next_marker_free_state:
assumes
"channel cid = Some (p, q)"
"trace init t final"
shows
"\<exists>!j. next_marker_free_state t i cid = j \<and> j \<ge> i \<and> Marker \<notin> set (msgs (S t j) cid)"
proof (cases "Marker \<in> set (msgs (S t i) cid)")
case False
then have "next_marker_free_state t i cid = i" unfolding next_marker_free_state_def
by (metis (no_types, lifting) Least_equality order_refl)
then show ?thesis using False assms by blast
next
case True
then obtain j where "j \<ge> i" "Marker \<notin> set (msgs (S t j) cid)" using l1 assms by blast
then show ?thesis
by (metis (no_types, lifting) LeastI_ex next_marker_free_state_def)
qed
theorem snapshot_algorithm_must_terminate:
assumes
"trace init t final"
shows
"\<exists>phi. ((\<forall>p. has_snapshotted (S t phi) p)
\<and> (\<forall>cid. Marker \<notin> set (msgs (S t phi) cid)))"
proof -
let ?i = "{i. i \<le> length t \<and> (\<forall>p. has_snapshotted (S t i) p)}"
have fin_i: "finite ?i" by auto
moreover have "?i \<noteq> empty"
proof -
have "\<forall>p. has_snapshotted (S t (length t)) p"
by (meson assms exists_trace_for_any_i_j l2 snapshot_stable_2)
then show ?thesis by blast
qed
then obtain i where asm: "\<forall>p. has_snapshotted (S t i) p" by blast
have f: "\<forall>j. j \<ge> i \<longrightarrow> (\<forall>p. has_snapshotted (S t j) p)"
using snapshot_stable asm exists_trace_for_any_i_j valid assms by blast
let ?s = "(\<lambda>cid. (next_marker_free_state t i cid)) ` { cid. channel cid \<noteq> None }"
have "?s \<noteq> empty" using exists_some_channel by auto
have fin_s: "finite ?s" using finite_channels by simp
let ?phi = "Max ?s"
have "?phi \<ge> i"
proof (rule ccontr)
assume asm: "\<not> ?phi \<ge> i"
obtain cid p q where g: "channel cid = Some (p, q)" using exists_some_channel by auto
then have "next_marker_free_state t i cid \<ge> i" using exists_next_marker_free_state assms by blast
then have "Max ?s \<ge> i" using Max_ge_iff g fin_s by fast
then show False using asm by simp
qed
then have "\<And>cid. Marker \<notin> set (msgs (S t ?phi) cid)"
proof -
fix cid
show "Marker \<notin> set (msgs (S t ?phi) cid)"
proof (cases "Marker : set (msgs (S t i) cid)")
case False
then show ?thesis
using \<open>i \<le> Max ?s\<close> asm assms exists_trace_for_any_i_j no_markers_if_all_snapshotted by blast
next
case True
then have cpq: "channel cid \<noteq> None" using no_messages_if_no_channel assms by fastforce
then obtain p q where chan: "channel cid = Some (p, q)" by auto
then obtain j where i: "j = next_marker_free_state t i cid" "Marker \<notin> set (msgs (S t j) cid)"
using exists_next_marker_free_state assms by fast
have "j \<le> ?phi" using cpq fin_s i(1) pair_imageI by simp
then show "Marker \<notin> set (msgs (S t ?phi) cid)"
proof -
have "trace (S t j) (take (?phi - j) (drop j t)) (S t ?phi)"
using \<open>j \<le> ?phi\<close> assms exists_trace_for_any_i_j by blast
moreover have "\<forall>p. has_snapshotted (S t j) p"
by (metis assms chan f computation.exists_next_marker_free_state computation_axioms i(1))
ultimately show ?thesis
using i(2) no_markers_if_all_snapshotted by blast
qed
qed
qed
- thus ?thesis using f `?phi \<ge> i` by blast
+ thus ?thesis using f \<open>?phi \<ge> i\<close> by blast
qed
subsection \<open>Correctness\<close>
text \<open>The greatest part of this work is spent on the correctness
of the Chandy-Lamport algorithm. We prove that the snapshot is
consistent, i.e.\ there exists a permutation $t'$ of the trace $t$ and an intermediate
configuration $c'$ of $t'$ such that the configuration recorded in the snapshot
corresponds to the snapshot taken during execution of $t$, which is given as Theorem 1
in~\cite{chandy}.\<close>
lemma snapshot_stable_ver_2:
shows "trace init t final \<Longrightarrow> has_snapshotted (S t i) p \<Longrightarrow> j \<ge> i \<Longrightarrow> has_snapshotted (S t j) p"
using exists_trace_for_any_i_j snapshot_stable by blast
lemma snapshot_stable_ver_3:
shows "trace init t final \<Longrightarrow> ~ has_snapshotted (S t i) p \<Longrightarrow> i \<ge> j \<Longrightarrow> ~ has_snapshotted (S t j) p"
using snapshot_stable_ver_2 by blast
lemma marker_must_stay_if_no_snapshot:
assumes
"trace init t final" and
"has_snapshotted (S t i) p" and
"~ has_snapshotted (S t i) q" and
"channel cid = Some (p, q)"
shows
"Marker : set (msgs (S t i) cid)"
proof -
obtain j where "~ has_snapshotted (S t j) p \<and> has_snapshotted (S t (Suc j)) p"
using exists_snapshot_for_all_p assms by blast
have "j \<le> i"
proof (rule ccontr)
assume asm: "~ j \<le> i"
then have "~ has_snapshotted (S t i) p"
using \<open>\<not> has_snapshotted (S t j) p \<and> has_snapshotted (S t (Suc j)) p\<close> assms(1) less_imp_le_nat snapshot_stable_ver_3
by (meson nat_le_linear)
then show False using assms(2) by simp
qed
have "i \<le> length t"
proof (rule ccontr)
assume "~ i \<le> length t"
then have "i > length t"
using not_less by blast
obtain i' where a: "\<forall>p. has_snapshotted (S t i') p" using assms snapshot_algorithm_must_terminate by blast
have "i' \<ge> i"
using \<open>\<forall>p. has_snapshotted (S t i') p\<close> assms(1) assms(3) nat_le_linear snapshot_stable_ver_3 by blast
have "(S t i') \<noteq> (S t i)" using assms a by force
then have "i \<le> length t"
using \<open>i \<le> i'\<close> assms(1) computation.no_change_if_ge_length_t computation_axioms nat_le_linear by fastforce
- then show False using `~ i \<le> length t` by simp
+ then show False using \<open>~ i \<le> length t\<close> by simp
qed
have marker_in_set: "Marker : set (msgs (S t (Suc j)) cid)"
using \<open>\<not> has_snapshotted (S t j) p \<and> has_snapshotted (S t (Suc j)) p\<close> \<open>j \<le> i\<close> assms(1) assms(3) assms(4) snapshot_produces_marker snapshot_stable_ver_3 by blast
show ?thesis
proof (rule ccontr)
assume asm: "Marker \<notin> set (msgs (S t i) cid)"
then have range: "(Suc j) < i"
by (metis Suc_lessI \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (Suc j)) p \<noteq> None\<close> \<open>j \<le> i\<close> assms(2) marker_in_set order.order_iff_strict)
let ?k = "LEAST k. k \<ge> (Suc j) \<and> Marker \<notin> set (msgs (S t k) cid)"
have range_k: "(Suc j) < ?k \<and> ?k \<le> i"
proof -
have "j < (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid))"
by (metis (full_types) Suc_le_eq assms(1) assms(4) exists_next_marker_free_state next_marker_free_state_def)
then show ?thesis
proof -
assume a1: "j < (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid))"
have "j < i"
using local.range by linarith (* 4 ms *)
then have "(Suc j \<le> i \<and> Marker \<notin> set (msgs (S t i) cid)) \<and> (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) \<noteq> Suc j"
by (metis (lifting) Suc_leI asm marker_in_set wellorder_Least_lemma(1)) (* 64 ms *)
then show ?thesis
using a1 by (simp add: wellorder_Least_lemma(2)) (* 16 ms *)
qed
qed
have a: "Marker : set (msgs (S t (?k-1)) cid)"
proof -
obtain nn :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"\<forall>x0 x1. (\<exists>v2. x0 = Suc (x1 + v2)) = (x0 = Suc (x1 + nn x0 x1))"
by moura
then have f1: "(LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) = Suc (Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j))"
using \<open>Suc j < (LEAST k. Suc j \<le> k \<and> Marker \<notin> set (msgs (S t k) cid)) \<and> (LEAST k. Suc j \<le> k \<and> Marker \<notin> set (msgs (S t k) cid)) \<le> i\<close> less_iff_Suc_add by fastforce
have f2: "Suc j \<le> Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j)"
by simp
have f3: "\<forall>p n. \<not> p (n::nat) \<or> Least p \<le> n"
by (meson wellorder_Least_lemma(2))
have "\<not> (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) \<le> Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j)"
using f1 by linarith
then have f4: "\<not> (Suc j \<le> Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j) \<and> Marker \<notin> set (msgs (S t (Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j))) cid))"
using f3 by force
have "Suc j + nn (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) (Suc j) = (LEAST n. Suc j \<le> n \<and> Marker \<notin> set (msgs (S t n) cid)) - 1"
using f1 by linarith
then show ?thesis
using f4 f2 by presburger
qed
have b: "Marker \<notin> set (msgs (S t ?k) cid)"
using assms(1) assms(4) exists_next_marker_free_state next_marker_free_state_def by fastforce
have "?k - 1 < i" using range_k by auto
then obtain ev where step: "(S t (?k-1)) \<turnstile> ev \<mapsto> (S t (Suc (?k-1)))"
by (meson Suc_le_eq \<open>i \<le> length t\<close> assms(1) le_trans step_Suc)
then show False
using a assms(1) assms(3) assms(4) b computation.snapshot_stable_ver_3 computation_axioms less_iff_Suc_add range_k recv_marker_means_snapshotted_2 by fastforce
qed
qed
subsubsection \<open>Pre- and postrecording events\<close>
definition prerecording_event:
"prerecording_event t i \<equiv>
i < length t \<and> regular_event (t ! i)
\<and> ~ has_snapshotted (S t i) (occurs_on (t ! i))"
definition postrecording_event:
"postrecording_event t i \<equiv>
i < length t \<and> regular_event (t ! i)
\<and> has_snapshotted (S t i) (occurs_on (t ! i))"
abbreviation neighboring where
"neighboring t i j \<equiv> i < j \<and> j < length t \<and> regular_event (t ! i) \<and> regular_event (t ! j)
\<and> (\<forall>k. i < k \<and> k < j \<longrightarrow> ~ regular_event (t ! k))"
lemma pre_if_regular_and_not_post:
assumes
"regular_event (t ! i)" and
"~ postrecording_event t i" and
"i < length t"
shows
"prerecording_event t i"
using assms computation.postrecording_event computation_axioms prerecording_event by metis
lemma post_if_regular_and_not_pre:
assumes
"regular_event (t ! i)" and
"~ prerecording_event t i" and
"i < length t"
shows
"postrecording_event t i"
using assms computation.postrecording_event computation_axioms prerecording_event by metis
lemma post_before_pre_different_processes:
assumes
"i < j" and
"j < length t" and
neighboring: "\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
post_ei: "postrecording_event t i" and
pre_ej: "prerecording_event t j" and
valid: "trace init t final"
shows
"occurs_on (t ! i) \<noteq> occurs_on (t ! j)"
proof -
let ?p = "occurs_on (t ! i)"
let ?q = "occurs_on (t ! j)"
have sp: "has_snapshotted (S t i) ?p"
using assms postrecording_event prerecording_event by blast
have nsq: "~ has_snapshotted (S t j) ?q"
using assms postrecording_event prerecording_event by blast
show "?p \<noteq> ?q"
proof -
have "~ has_snapshotted (S t i) ?q"
proof (rule ccontr)
assume sq: "~ ~ has_snapshotted (S t i) ?q"
- from `i < j` have "i \<le> j" using less_imp_le by blast
+ from \<open>i < j\<close> have "i \<le> j" using less_imp_le by blast
then obtain tr where ex_trace: "trace (S t i) tr (S t j)"
using exists_trace_for_any_i_j valid by blast
then have "has_snapshotted (S t j) ?q" using ex_trace snapshot_stable sq by blast
then show False using nsq by simp
qed
then show ?thesis using sp by auto
qed
qed
lemma post_before_pre_neighbors:
assumes
"i < j" and
"j < length t" and
neighboring: "\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
post_ei: "postrecording_event t i" and
pre_ej: "prerecording_event t j" and
valid: "trace init t final"
shows
"Ball (set (take (j - (i+1)) (drop (i+1) t))) (%ev. ~ regular_event ev \<and> ~ occurs_on ev = occurs_on (t ! j))"
proof -
let ?p = "occurs_on (t ! i)"
let ?q = "occurs_on (t ! j)"
let ?between = "take (j - (i+1)) (drop (i+1) t)"
show ?thesis
proof (unfold Ball_def, rule allI, rule impI)
fix ev
assume "ev : set ?between"
have len_nr: "length ?between = (j - (i+1))" using assms(2) by auto
then obtain l where "?between ! l = ev" and range_l: "0 \<le> l \<and> l < (j - (i+1))"
by (metis \<open>ev \<in> set (take (j - (i + 1)) (drop (i + 1) t))\<close> gr_zeroI in_set_conv_nth le_numeral_extra(3) less_le)
let ?k = "l + (i+1)"
have "?between ! l = (t ! ?k)"
proof -
have "j < length t"
by (metis assms(2))
then show ?thesis
by (metis (no_types) Suc_eq_plus1 Suc_leI add.commute assms(1) drop_take length_take less_diff_conv less_imp_le_nat min.absorb2 nth_drop nth_take range_l)
qed
have "~ regular_event ev"
by (metis (no_types, lifting) assms(3) range_l One_nat_def Suc_eq_plus1 \<open>take (j - (i + 1)) (drop (i + 1) t) ! l = ev\<close> \<open>take (j - (i + 1)) (drop (i + 1) t) ! l = t ! (l + (i + 1))\<close> add.left_commute add_lessD1 lessI less_add_same_cancel2 less_diff_conv order_le_less)
have step_ev: "(S t ?k) \<turnstile> ev \<mapsto> (S t (?k+1))"
proof -
have "j \<le> length t"
by (metis assms(2) less_or_eq_imp_le)
then have "l + (i + 1) < length t"
by (meson less_diff_conv less_le_trans range_l)
then show ?thesis
by (metis (no_types) Suc_eq_plus1 \<open>take (j - (i + 1)) (drop (i + 1) t) ! l = ev\<close> \<open>take (j - (i + 1)) (drop (i + 1) t) ! l = t ! (l + (i + 1))\<close> distributed_system.step_Suc distributed_system_axioms valid)
qed
- obtain cid s r where f: "ev = RecvMarker cid s r \<or> ev = Snapshot r" using `~ regular_event ev`
+ obtain cid s r where f: "ev = RecvMarker cid s r \<or> ev = Snapshot r" using \<open>~ regular_event ev\<close>
by (meson isRecvMarker_def isSnapshot_def nonregular_event)
from f have "occurs_on ev \<noteq> ?q"
proof (elim disjE)
assume snapshot: "ev = Snapshot r"
show ?thesis
proof (rule ccontr)
assume occurs_on_q: "~ occurs_on ev \<noteq> ?q"
then have "?q = r" using snapshot by auto
then have q_snapshotted: "has_snapshotted (S t (?k+1)) ?q"
using snapshot step_ev by auto
then show False
proof -
have "l + (i + 1) < j"
by (meson less_diff_conv range_l)
then show ?thesis
by (metis (no_types) Suc_eq_plus1 Suc_le_eq computation.snapshot_stable_ver_2 computation_axioms pre_ej prerecording_event q_snapshotted valid)
qed
qed
next
assume RecvMarker: "ev = RecvMarker cid s r"
show ?thesis
proof (rule ccontr)
assume occurs_on_q: "~ occurs_on ev \<noteq> ?q"
then have "s = ?q" using RecvMarker by auto
then have q_snapshotted: "has_snapshotted (S t (?k+1)) ?q"
proof (cases "has_snapshotted (S t ?k) ?q")
case True
then show ?thesis using snapshot_stable_ver_2 step_Suc step_ev valid by auto
next
case False
then show "has_snapshotted (S t (?k+1)) ?q"
using \<open>s = ?q\<close> next_recv_marker RecvMarker step_ev by auto
qed
then show False
proof -
have "l + (i + 1) < j"
using less_diff_conv range_l by blast
then show ?thesis
by (metis (no_types) Suc_eq_plus1 Suc_le_eq computation.snapshot_stable_ver_2 computation_axioms pre_ej prerecording_event q_snapshotted valid)
qed
qed
qed
then show "\<not> regular_event ev \<and> occurs_on ev \<noteq> ?q"
- using `~ regular_event ev` by simp
+ using \<open>~ regular_event ev\<close> by simp
qed
qed
lemma can_swap_neighboring_pre_and_postrecording_events:
assumes
"i < j" and
"j < length t" and
"occurs_on (t ! i) = p" and
"occurs_on (t ! j) = q" and
neighboring: "\<forall>k. (i < k \<and> k < j)
\<longrightarrow> ~ regular_event (t ! k)" and
post_ei: "postrecording_event t i" and
pre_ej: "prerecording_event t j" and
valid: "trace init t final"
shows
"can_occur (t ! j) (S t i)"
proof -
have "p \<noteq> q" using post_before_pre_different_processes assms by auto
have sp: "has_snapshotted (S t i) p"
using assms(3) post_ei postrecording_event prerecording_event by blast
have nsq: "~ has_snapshotted (S t j) q"
using assms(4) pre_ej prerecording_event by auto
let ?nr = "take (j - (Suc i)) (drop (Suc i) t)"
have valid_subtrace: "trace (S t (Suc i)) ?nr (S t j)"
using assms(1) exists_trace_for_any_i_j valid by fastforce
have "Ball (set ?nr) (%ev. ~ occurs_on ev = q \<and> ~ regular_event ev)"
proof -
have "?nr = take (j - (i+1)) (drop (i+1) t)" by auto
then show ?thesis
by (metis assms(1) assms(2) assms(4) neighboring post_ei pre_ej valid post_before_pre_neighbors)
qed
then have la: "list_all (%ev. ~ occurs_on ev = q) ?nr"
by (meson list_all_length nth_mem)
have tj_to_tSi: "can_occur (t ! j) (S t (Suc i))"
proof -
have "list_all (%ev. ~ isSend ev) ?nr"
proof -
have "list_all (%ev. ~ regular_event ev) ?nr"
using \<open>\<forall>ev\<in>set (take (j - (Suc i)) (drop (Suc i) t)). occurs_on ev \<noteq> q \<and> \<not> regular_event ev\<close> \<open>list_all (\<lambda>ev. occurs_on ev \<noteq> q) (take (j - (Suc i)) (drop (Suc i) t))\<close> list.pred_mono_strong by fastforce
then show ?thesis
by (simp add: list.pred_mono_strong)
qed
moreover have "~ isRecvMarker (t ! j)" using prerecording_event assms by auto
moreover have "can_occur (t ! j) (S t j)"
proof -
have "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (Suc j))"
using assms(2) step_Suc valid by auto
then show ?thesis
using happen_implies_can_occur by blast
qed
ultimately show "can_occur (t ! j) (S t (Suc i))"
using assms(4) event_can_go_back_if_no_sender_trace valid_subtrace la by blast
qed
show "can_occur (t ! j) (S t i)"
proof (cases "isSend (t ! i)")
case False
have "~ isRecvMarker (t ! j)" using assms prerecording_event by auto
moreover have "~ isSend (t ! i)" using False by simp
ultimately show ?thesis
by (metis \<open>p \<noteq> q\<close> assms(3) assms(4) event_can_go_back_if_no_sender post_ei postrecording_event step_Suc tj_to_tSi valid)
next
case True
obtain cid s u u' m where Send: "t ! i = Send cid p s u u' m"
by (metis True isSend_def assms(3) event.sel(2))
have chan: "channel cid = Some (p, s)"
proof -
have "can_occur (t ! i) (S t i)"
by (meson computation.postrecording_event computation_axioms happen_implies_can_occur post_ei step_Suc valid)
then show ?thesis using can_occur_def Send by simp
qed
have n: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i))"
using assms(1) assms(2) step_Suc valid True by auto
have st: "states (S t i) q = states (S t (Suc i)) q"
using Send \<open>p \<noteq> q\<close> n by auto
have "isTrans (t ! j) \<or> isSend (t ! j) \<or> isRecv (t ! j)"
using assms(7) computation.prerecording_event computation_axioms regular_event by blast
then show ?thesis
proof (elim disjE)
assume "isTrans (t ! j)"
then show ?thesis
by (metis (no_types, lifting) tj_to_tSi st can_occur_def assms(4) event.case(1) event.collapse(1))
next
assume "isSend (t ! j)"
then obtain cid' s' u'' u''' m' where Send: "t ! j = Send cid' q s' u'' u''' m'"
by (metis (no_types, lifting) assms(4) event.sel(2) isSend_def)
have co_tSi: "can_occur (Send cid' q s' u'' u''' m') (S t (Suc i))"
using Send tj_to_tSi by auto
then have "channel cid' = Some (q, s') \<and> send cid' q s' u'' u''' m'"
using Send can_occur_def by simp
then show ?thesis using can_occur_def st Send assms co_tSi by auto
next
assume "isRecv (t ! j)"
then obtain cid' s' u'' u''' m' where Recv: "t ! j = Recv cid' q s' u'' u''' m'"
by (metis assms(4) event.sel(3) isRecv_def)
have co_tSi: "can_occur (Recv cid' q s' u'' u''' m') (S t (Suc i))"
using Recv tj_to_tSi by auto
then have a: "channel cid' = Some (s', q) \<and> length (msgs (S t (Suc i)) cid') > 0
\<and> hd (msgs (S t (Suc i)) cid') = Msg m'"
using can_occur_def co_tSi by fastforce
show "can_occur (t ! j) (S t i)"
proof (cases "cid = cid'")
case False
with Send n have "msgs (S t (Suc i)) cid' = msgs (S t i) cid'" by auto
then have b: "length (msgs (S t i) cid') > 0 \<and> hd (msgs (S t i) cid') = Msg m'"
using a by simp
with can_occur_Recv co_tSi st a Recv show ?thesis
unfolding can_occur_def by auto
next
case True (* This is the interesting case *)
have stu: "states (S t i) q = u''"
using can_occur_Recv co_tSi st by blast
show ?thesis
proof (rule ccontr)
have marker_in_set: "Marker \<in> set (msgs (S t i) cid)"
proof -
have "(s', q) = (p, q)"
using True a chan by auto
then show ?thesis
by (metis (no_types, lifting) True \<open>p \<noteq> q\<close> a assms(3) marker_must_stay_if_no_snapshot n no_state_change_if_no_event nsq snapshot_stable_2 sp valid valid_subtrace)
qed
assume asm: "~ can_occur (t ! j) (S t i)"
then show False
proof (unfold can_occur_def, (auto simp add: marker_in_set True Recv stu))
assume "msgs (S t i) cid' = []"
then show False using marker_in_set
by (simp add: True)
next
assume "hd (msgs (S t i) cid') \<noteq> Msg m'"
have "msgs (S t i) cid \<noteq> []" using marker_in_set by auto
then have "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Msg m]"
using Send True n chan by auto
then have "hd (msgs (S t (Suc i)) cid) \<noteq> Msg m'"
using True \<open>hd (msgs (S t i) cid') \<noteq> Msg m'\<close> \<open>msgs (S t i) cid \<noteq> []\<close> by auto
then have "~ can_occur (t ! j) (S t (Suc i))"
using True a by blast
then show False
using tj_to_tSi by blast
next
assume "~ recv cid' q s' u'' u''' m'"
then show False
using can_occur_Recv co_tSi by blast
next
assume "channel cid' \<noteq> Some (s', q)"
then show False using can_occur_def tj_to_tSi Recv by simp
qed
qed
qed
qed
qed
qed
subsubsection \<open>Event swapping\<close>
lemma swap_events:
shows "\<lbrakk> i < j; j < length t;
\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k);
postrecording_event t i; prerecording_event t j;
trace init t final \<rbrakk>
\<Longrightarrow> trace init (swap_events i j t) final
\<and> (\<forall>k. k \<ge> j + 1 \<longrightarrow> S (swap_events i j t) k = S t k)
\<and> (\<forall>k. k \<le> i \<longrightarrow> S (swap_events i j t) k = S t k)
\<and> prerecording_event (swap_events i j t) i
\<and> postrecording_event (swap_events i j t) (i+1)
\<and> (\<forall>k. k > i+1 \<and> k < j+1
\<longrightarrow> ~ regular_event ((swap_events i j t) ! k))"
proof (induct "j - (i+1)" arbitrary: j t)
case 0
let ?p = "occurs_on (t ! i)"
let ?q = "occurs_on (t ! j)"
have "j = (i+1)"
using "0.prems" "0.hyps" by linarith
let ?subt = "take (j - (i+1)) (drop (i+1) t)"
have "t = take i t @ [t ! i] @ ?subt @ [t ! j] @ drop (j+1) t"
proof -
have "take (Suc i) t = take i t @ [t ! i]"
using "0.prems"(2) \<open>j = i + 1\<close> add_lessD1 take_Suc_conv_app_nth by blast
then show ?thesis
by (metis (no_types) "0.hyps" "0.prems"(2) Suc_eq_plus1 \<open>j = i + 1\<close> append_assoc append_take_drop_id self_append_conv2 take_Suc_conv_app_nth take_eq_Nil)
qed
have sp: "has_snapshotted (S t i) ?p"
using "0.prems" postrecording_event prerecording_event by blast
have nsq: "~ has_snapshotted (S t j) ?q"
using "0.prems" postrecording_event prerecording_event by blast
have "?p \<noteq> ?q"
using "0.prems" computation.post_before_pre_different_processes computation_axioms by blast
have "?subt = Nil"
by (simp add: \<open>j = i + 1\<close>)
have reg_step_1: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t j)"
by (metis "0.prems"(2) "0.prems"(6) Suc_eq_plus1 \<open>j = i + 1\<close> add_lessD1 step_Suc)
have reg_step_2: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
using "0.prems"(2) "0.prems"(6) step_Suc by auto
have "can_occur (t ! j) (S t i)"
using "0.prems" can_swap_neighboring_pre_and_postrecording_events by blast
then obtain d' where new_step1: "(S t i) \<turnstile> (t ! j) \<mapsto> d'"
using exists_next_if_can_occur by blast
have st: "states d' ?p = states (S t i) ?p"
using \<open>(S t i) \<turnstile> t ! j \<mapsto> d'\<close> \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> no_state_change_if_no_event by auto
then have "can_occur (t ! i) d'"
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> event_stays_valid_if_no_occurrence happen_implies_can_occur new_step1 reg_step_1 by auto
then obtain e where new_step2: "d' \<turnstile> (t ! i) \<mapsto> e"
using exists_next_if_can_occur by blast
have "states e = states (S t (j+1))"
proof (rule ext)
fix p
show "states e p = states (S t (j+1)) p"
proof (cases "p = ?p \<or> p = ?q")
case True
then show ?thesis
proof (elim disjE)
assume "p = ?p"
then have "states d' p = states (S t i) p"
by (simp add: st)
thm same_state_implies_same_result_state
then have "states e p = states (S t j) p"
using "0.prems"(2) "0.prems"(6) new_step2 reg_step_1 by (blast intro:same_state_implies_same_result_state[symmetric])
moreover have "states (S t j) p = states (S t (j+1)) p"
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> \<open>p = occurs_on (t ! i)\<close> no_state_change_if_no_event reg_step_2 by auto
ultimately show ?thesis by simp
next
assume "p = ?q"
then have "states (S t j) p = states (S t i) p"
using reg_step_1 \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> no_state_change_if_no_event by auto
then have "states d' p = states (S t (j+1)) p"
using "0.prems"(5) prerecording_event computation_axioms new_step1 reg_step_2 same_state_implies_same_result_state by blast
moreover have "states e p = states (S t (j+1)) p"
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> \<open>p = occurs_on (t ! j)\<close> calculation new_step2 no_state_change_if_no_event by auto
ultimately show ?thesis by simp
qed
next
case False
then have "states (S t i) p = states (S t j) p"
using no_state_change_if_no_event reg_step_1 by auto
moreover have "... = states (S t (j+1)) p"
using False no_state_change_if_no_event reg_step_2 by auto
moreover have "... = states d' p"
using False calculation new_step1 no_state_change_if_no_event by auto
moreover have "... = states e p"
using False new_step2 no_state_change_if_no_event by auto
ultimately show ?thesis by simp
qed
qed
moreover have "msgs e = msgs (S t (j+1))"
proof (rule ext)
fix cid
have "isTrans (t ! i) \<or> isSend (t ! i) \<or> isRecv (t ! i)"
using "0.prems"(4) computation.postrecording_event computation_axioms regular_event by blast
moreover have "isTrans (t ! j) \<or> isSend (t ! j) \<or> isRecv (t ! j)"
using "0.prems"(5) computation.prerecording_event computation_axioms regular_event by blast
ultimately show "msgs e cid = msgs (S t (j+1)) cid"
proof (elim disjE, goal_cases)
case 1
then have "msgs d' cid = msgs (S t j) cid"
by (metis Trans_msg new_step1 reg_step_1)
then show ?thesis
using Trans_msg \<open>isTrans (t ! i)\<close> \<open>isTrans (t ! j)\<close> new_step2 reg_step_2 by auto
next
case 2
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Trans_Send by auto
next
case 3
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Trans_Recv by auto
next
case 4
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Trans by auto
next
case 5
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Recv_Trans by auto
next
case 6
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 by (blast intro:swap_msgs_Send_Send[symmetric])
next
case 7
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Recv by auto
next
case 8
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 swap_msgs_Send_Recv by simp
next
case 9
then show ?thesis
using \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> new_step1 new_step2 reg_step_1 reg_step_2 by (blast intro:swap_msgs_Recv_Recv[symmetric])
qed
qed
moreover have "process_snapshot e = process_snapshot (S t (j+1))"
proof (rule ext)
fix p
have "process_snapshot d' p = process_snapshot (S t j) p"
by (metis "0.prems"(4) "0.prems"(5) computation.postrecording_event computation.prerecording_event computation_axioms new_step1 reg_step_1 regular_event_preserves_process_snapshots)
then show "process_snapshot e p = process_snapshot (S t (j+1)) p"
by (metis "0.prems"(4) "0.prems"(5) computation.postrecording_event computation.prerecording_event computation_axioms new_step2 reg_step_2 regular_event_preserves_process_snapshots)
qed
moreover have "channel_snapshot e = channel_snapshot (S t (j+1))"
proof (rule ext)
fix cid
show "cs e cid = cs (S t (j+1)) cid"
proof (cases "isRecv (t ! i)"; cases "isRecv (t ! j)", goal_cases)
case 1
then show ?thesis
using \<open>?p \<noteq> ?q\<close> new_step1 new_step2 reg_step_1 reg_step_2
by (blast intro:regular_event_implies_same_channel_snapshot_Recv_Recv[symmetric])
next
case 2
moreover have "regular_event (t ! j)" using prerecording_event 0 by simp
ultimately show ?thesis
using \<open>?p \<noteq> ?q\<close> new_step1 new_step2 reg_step_1 reg_step_2 regular_event_implies_same_channel_snapshot_Recv by auto
next
assume 3: "~ isRecv (t ! i)" "isRecv (t ! j)"
moreover have "regular_event (t ! i)" using postrecording_event 0 by simp
ultimately show ?thesis
using \<open>?p \<noteq> ?q\<close> new_step1 new_step2 reg_step_1 reg_step_2 regular_event_implies_same_channel_snapshot_Recv by auto
next
assume 4: "~ isRecv (t ! i)" "~ isRecv (t ! j)"
moreover have "regular_event (t ! j)" using prerecording_event 0 by simp
moreover have "regular_event (t ! i)" using postrecording_event 0 by simp
ultimately show ?thesis
using \<open>?p \<noteq> ?q\<close> new_step1 new_step2 reg_step_1 reg_step_2
by (metis no_cs_change_if_no_event)
qed
qed
ultimately have "e = S t (j+1)" by simp
then have "(S t i) \<turnstile> (t ! j) \<mapsto> d' \<and> d' \<turnstile> (t ! i) \<mapsto> (S t (j+1))"
using new_step1 new_step2 by blast
then have swap: "trace (S t i) [t ! j, t ! i] (S t (j+1))"
by (meson trace.simps)
have "take (j-1) t @ [t ! j, t ! i] = ((take (j+1) t)[i := t ! j])[j := t ! i]"
proof -
have "i = j - 1"
by (simp add: \<open>j = i + 1\<close>)
show ?thesis
- proof (subst (1 2 3) `i = j - 1`)
+ proof (subst (1 2 3) \<open>i = j - 1\<close>)
have "j < length t" using "0.prems" by auto
then have "take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t = t[j - 1 := t ! j, j := t ! (j - 1)]"
by (metis Suc_eq_plus1 \<open>i = j - 1\<close> \<open>j = i + 1\<close> add_Suc_right arith_special(3) swap_neighbors)
then show "take (j - 1) t @ [t ! j, t ! (j - 1)] = (take (j+1) t)[j - 1 := t ! j, j := t ! (j - 1)]"
proof -
assume a1: "take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t = t [j - 1 := t ! j, j := t ! (j - 1)]"
have f2: "t[j - 1 := t ! j, j := t ! (j - 1)] = take j (t[j - 1 := t ! j]) @ t ! (j - 1) # drop (Suc j) (t[j - 1 := t ! j])"
by (metis (no_types) "0.prems"(2) length_list_update upd_conv_take_nth_drop) (* 32 ms *)
have f3: "\<forall>n na. \<not> n < na \<or> Suc n \<le> na"
using Suc_leI by blast (* 0.0 ms *)
then have "min (length t) (j + 1) = j + 1"
by (metis (no_types) "0.prems"(2) Suc_eq_plus1 min.absorb2) (* 16 ms *)
then have f4: "length ((take (j + 1) t)[j - 1 := t ! j]) = j + 1"
by simp (* 4 ms *)
have f5: "j + 1 \<le> length (t[j - 1 := t ! j])"
using f3 by (metis (no_types) "0.prems"(2) Suc_eq_plus1 length_list_update) (* 8 ms *)
have "Suc j \<le> j + 1"
by linarith (* 0.0 ms *)
then have "(take (j + 1) (t[j - 1 := t ! j]))[j := t ! (j - 1)] = take j (t[j - 1 := t ! j]) @ t ! (j - 1) # [] @ []"
using f5 f4 by (metis (no_types) Suc_eq_plus1 add_diff_cancel_right' butlast_conv_take butlast_take drop_eq_Nil lessI self_append_conv2 take_update_swap upd_conv_take_nth_drop) (* 180 ms *)
then show ?thesis
using f2 a1 by (simp add: take_update_swap) (* 120 ms *)
qed
qed
qed
have s: "trace init (take i t) (S t i)"
using "0.prems"(6) exists_trace_for_any_i by blast
have e: "trace (S t (j+1)) (take (length t - (j+1)) (drop (j+1) t)) final"
proof -
have "trace init (take (length t) t) final"
by (simp add: "0.prems"(6))
then show ?thesis
by (metis "0.prems"(2) Suc_eq_plus1 Suc_leI exists_trace_for_any_i exists_trace_for_any_i_j nat_le_linear take_all trace_and_start_determines_end)
qed
have "trace init (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) final"
proof -
from s swap have "trace init (take i t @ [t ! j,t ! i]) (S t (j+1))" using trace_trans by blast
then have "trace init (take i t @ [t ! j, t ! i] @ (take (length t - (j+1)) (drop (j+1) t))) final"
using e trace_trans by fastforce
moreover have "take (length t - (j+1)) (drop (j+1) t) = drop (j+1) t" by simp
ultimately show ?thesis by simp
qed
moreover have "take i t @ [t ! j] @ [t ! i] @ drop (j+1) t = (t[i := t ! j])[j := t ! i]"
proof -
have "length (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) = length ((t[i := t ! j])[j := t ! i])"
by (metis (mono_tags, lifting) \<open>t = take i t @ [t ! i] @ take (j - (i + 1)) (drop (i + 1) t) @ [t ! j] @ drop (j + 1) t\<close> \<open>take (j - (i + 1)) (drop (i + 1) t) = []\<close> length_append length_list_update list.size(4) self_append_conv2)
moreover have "\<And>k. k < length ((t[i := t ! j])[j := t ! i]) \<Longrightarrow> (take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = ((t[i := t ! j])[j := t ! i]) ! k"
proof -
fix k
assume "k < length ((t[i := t ! j])[j := t ! i])"
show "(take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = ((t[i := t ! j])[j := t ! i]) ! k"
proof (cases "k = i \<or> k = j")
case True
then show ?thesis
proof (elim disjE)
assume "k = i"
then show ?thesis
by (metis (no_types, lifting) \<open>k < length (t[i := t ! j, j := t ! i])\<close> append_Cons le_eq_less_or_eq length_list_update length_take min.absorb2 nth_append_length nth_list_update_eq nth_list_update_neq)
next
assume "k = j"
then show ?thesis
by (metis (no_types, lifting) "0.prems"(4) Suc_eq_plus1 \<open>j = i + 1\<close> \<open>k < length (t[i := t ! j, j := t ! i])\<close> append.assoc append_Cons le_eq_less_or_eq length_append_singleton length_list_update length_take min.absorb2 nth_append_length nth_list_update postrecording_event)
qed
next
case knij: False
then show ?thesis
proof (cases "k < i")
case True
then show ?thesis
by (metis (no_types, lifting) "0.prems"(2) \<open>j = i + 1\<close> add_lessD1 length_take less_imp_le_nat min.absorb2 not_less nth_append nth_list_update_neq nth_take)
next
case False
then have "k > j"
using \<open>j = i + 1\<close> knij by linarith
then have "(take i t @ [t ! j] @ [t ! i] @ drop (j+1) t) ! k = drop (j+1) t ! (k-(j+1))"
proof -
assume a1: "j < k"
have f2: "\<forall>n na. ((n::nat) < na) = (n \<le> na \<and> n \<noteq> na)"
using nat_less_le by blast (* 0.0 ms *)
have f3: "i + 0 = min (length t) i + (0 + 0)"
using "0.prems"(2) \<open>j = i + 1\<close> by linarith (* 8 ms *)
have f4: "min (length t) i + Suc (0 + 0) = length (take i t) + length [t ! j]"
by force (* 4 ms *)
have f5: "take i t @ [t ! j] @ [] = take i t @ [t ! j]"
by auto (* 0.0 ms *)
have "j = length (take i t @ [t ! j] @ [])"
using f3 by (simp add: \<open>j = i + 1\<close>) (* 4 ms *)
then have "j + 1 = length (take i t @ [t ! j] @ [t ! i])"
by fastforce (* 4 ms *)
then show ?thesis
using f5 f4 f3 f2 a1 by (metis (no_types) One_nat_def \<open>j = i + 1\<close> add_Suc_right append.assoc length_append less_antisym list.size(3) not_less nth_append) (* 284 ms *)
qed
moreover have "(t[i := t ! j])[j := t ! i] ! k = drop (j+1) ((t[i := t ! j])[j := t ! i]) ! (k-(j+1))"
using "0.prems"(2) \<open>j < k\<close> by auto
moreover have "drop (j+1) ((t[i := t ! j])[j := t ! i]) = drop (j+1) t"
using "0.prems"(1) by auto
ultimately show ?thesis by simp
qed
qed
qed
ultimately show ?thesis by (simp add: list_eq_iff_nth_eq)
qed
moreover have "\<forall>k. k \<ge> j + 1 \<longrightarrow> S t k = S ((t[i := t ! j])[j := t ! i]) k"
proof (rule allI, rule impI)
fix k
assume "k \<ge> j + 1"
let ?newt = "((t[i := t ! j])[j := t ! i])"
have "trace init (take k ?newt) (S ?newt k)"
using calculation(1) calculation(2) exists_trace_for_any_i by auto
have "take k ?newt = take (j+1) ?newt @ take (k - (j+1)) (drop (j+1) ?newt)"
by (metis \<open>j + 1 \<le> k\<close> le_add_diff_inverse take_add)
have same_traces: "drop (j+1) t = drop (j+1) ?newt"
by (metis "0.prems"(1) Suc_eq_plus1 \<open>j = i + 1\<close> drop_update_cancel less_SucI less_add_same_cancel1)
have "trace init (take (j+1) ((t[i := t ! j])[j := t ! i])) (S t (j+1))"
by (metis (no_types, lifting) \<open>j = i + 1\<close> \<open>take (j - 1) t @ [t ! j, t ! i] = (take (j + 1) t)[i := t ! j, j := t ! i]\<close> add_diff_cancel_right' local.swap s take_update_swap trace_trans)
moreover have "trace init (take (j+1) ?newt) (S ?newt (j+1))"
using \<open>take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\<close> \<open>trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\<close> exists_trace_for_any_i by auto
ultimately have "S ?newt (j+1) = S t (j+1)"
using trace_and_start_determines_end by blast
have "trace (S t (j+1)) (take (k - (j+1)) (drop (j+1) t)) (S t k)"
using "0.prems"(6) \<open>j + 1 \<le> k\<close> exists_trace_for_any_i_j by blast
moreover have "trace (S ?newt (j+1)) (take (k - (j+1)) (drop (j+1) ?newt)) (S ?newt k)"
using \<open>j + 1 \<le> k\<close> \<open>take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\<close> \<open>trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\<close> exists_trace_for_any_i_j by fastforce
ultimately show "S t k = S ?newt k"
using \<open>S (t[i := t ! j, j := t ! i]) (j + 1) = S t (j + 1)\<close> same_traces trace_and_start_determines_end by auto
qed
moreover have "\<forall>k. k \<le> i \<longrightarrow> S t k = S ((t[i := t ! j])[j := t ! i]) k"
proof (rule allI, rule impI)
fix k
assume "k \<le> i"
let ?newt = "((t[i := t ! j])[j := t ! i])"
have "trace init (take k t) (S t k)"
using "0.prems"(6) exists_trace_for_any_i by blast
moreover have "trace init (take k ?newt) (S ?newt k)"
using \<open>take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\<close> \<open>trace init (take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t) final\<close> exists_trace_for_any_i by auto
moreover have "take k t = take k ?newt"
using "0.prems"(1) \<open>k \<le> i\<close> by auto
ultimately show "S t k = S ?newt k"
by (simp add: trace_and_start_determines_end)
qed
moreover have "prerecording_event (swap_events i j t) i"
proof -
have "~ has_snapshotted (S ((t[i := t ! j])[j := t ! i]) i) ?q"
by (metis "0.prems"(6) \<open>j = i + 1\<close> add.right_neutral calculation(4) le_add1 nsq snapshot_stable_ver_3)
moreover have "regular_event ((t[i := t ! j])[j := t ! i] ! i)"
by (metis "0.prems"(4) "0.prems"(5) \<open>occurs_on (t ! i) \<noteq> occurs_on (t ! j)\<close> nth_list_update_eq nth_list_update_neq postrecording_event prerecording_event)
moreover have "i < length ((t[i := t ! j])[j := t ! i])"
using "0.prems"(1) "0.prems"(2) by auto
ultimately show ?thesis unfolding prerecording_event
by (metis (no_types, hide_lams) "0.prems"(1) \<open>take (j - (i + 1)) (drop (i + 1) t) = []\<close> \<open>take i t @ [t ! j] @ [t ! i] @ drop (j + 1) t = t[i := t ! j, j := t ! i]\<close> append_Cons length_list_update nat_less_le nth_list_update_eq nth_list_update_neq self_append_conv2)
qed
moreover have "postrecording_event (swap_events i j t) (i+1)"
proof -
have "has_snapshotted (S ((t[i := t ! j])[j := t ! i]) (i+1)) ?p"
by (metis "0.prems"(4) add.right_neutral calculation(1) calculation(2) calculation(4) le_add1 postrecording_event snapshot_stable_ver_3)
moreover have "regular_event ((t[i := t ! j])[j := t ! i] ! j)"
using "0.prems"(2) "0.prems"(4) length_list_update postrecording_event by auto
moreover have "j < length t" using "0.prems" by auto
ultimately show ?thesis unfolding postrecording_event
by (metis \<open>j = i + 1\<close> length_list_update nth_list_update_eq swap_neighbors_2)
qed
moreover have "\<forall>k. k > i+1 \<and> k < j+1 \<longrightarrow> ~ regular_event ((swap_events i j t) ! k)" using "0" by force
- ultimately show ?case using `j = i + 1` by force
+ ultimately show ?case using \<open>j = i + 1\<close> by force
next
case (Suc n)
let ?p = "occurs_on (t ! i)"
let ?q = "occurs_on (t ! j)"
let ?t = "take ((j+1) - i) (drop i t)"
let ?subt = "take (j - (i+1)) (drop (i+1) t)"
let ?subt' = "take ((j-1) - (i+1)) (drop (i+1) t)"
have sp: "has_snapshotted (S t i) ?p"
using Suc.prems postrecording_event prerecording_event by blast
have nsq: "~ has_snapshotted (S t j) ?q"
using Suc.prems postrecording_event prerecording_event by blast
have "?p \<noteq> ?q"
using Suc.prems computation.post_before_pre_different_processes computation_axioms by blast
have "?subt \<noteq> Nil"
using Suc.hyps(2) Suc.prems(1) Suc.prems(2) by auto
have "?subt' = butlast ?subt"
by (metis Suc.prems(2) Suc_eq_plus1 butlast_drop butlast_take drop_take less_imp_le_nat)
have "?t = t ! i # ?subt @ [t ! j]"
proof -
have f1: "Suc j - i = Suc (j - i)"
using Suc.prems(1) Suc_diff_le le_simps(1) by presburger
have f2: "t ! i # drop (Suc i) t = drop i t"
by (meson Cons_nth_drop_Suc Suc.prems(1) Suc.prems(2) less_trans)
have f3: "t ! j # drop (Suc j) t = drop j t"
using Cons_nth_drop_Suc Suc.prems(2) by blast
have f4: "j - (i + 1) + (i + 1) = j"
using Suc.prems(1) by force
have "j - (i + 1) + Suc 0 = j - i"
using Suc.prems(1) Suc_diff_Suc by presburger
then show ?thesis
using f4 f3 f2 f1 by (metis One_nat_def Suc.hyps(2) Suc_eq_plus1 drop_drop take_Suc_Cons take_add take_eq_Nil)
qed
then have "trace (S t i) ?t (S t (j+1))"
by (metis Suc.prems(1) Suc.prems(6) Suc_eq_plus1 exists_trace_for_any_i_j less_SucI nat_less_le)
then have reg_tr_1: "trace (S t i) (t ! i # ?subt) (S t j)"
by (metis (no_types, hide_lams) Suc.hyps(2) Suc.prems(1) Suc.prems(4) Suc.prems(6) Suc_eq_plus1 discrete exists_trace_for_any_i_j postrecording_event step_Suc tr_step)
have reg_st_2: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
using Suc.prems(2) Suc.prems(6) step_Suc by auto
have "?subt = ?subt' @ [t ! (j-1)]"
proof -
have f1: "\<forall>n es. \<not> n < length es \<or> take n es @ [hd (drop n es)::('a, 'b, 'c) event] = take (Suc n) es"
by (meson take_hd_drop) (* 0.0 ms *)
have f2: "j - 1 - (i + 1) = n"
by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 diff_Suc_1 diff_diff_left plus_1_eq_Suc) (* 28 ms *)
have f3: "\<forall>n na. \<not> n < na \<or> Suc n \<le> na"
using Suc_leI by blast (* 0.0 ms *)
then have f4: "Suc i \<le> j - 1"
by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 diff_diff_left plus_1_eq_Suc zero_less_Suc zero_less_diff) (* 12 ms *)
have f5: "i + 1 < j"
by (metis Suc.hyps(2) zero_less_Suc zero_less_diff) (* 4 ms *)
then have f6: "t ! (j - 1) = hd (drop n (drop (i + 1) t))"
using f4 f3 by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_eq_plus1 Suc_lessD add_Suc_right diff_Suc_1 drop_drop hd_drop_conv_nth le_add_diff_inverse2 plus_1_eq_Suc) (* 140 ms *)
have "n < length (drop (i + 1) t)"
using f5 f3 by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_eq_plus1 Suc_lessD drop_drop le_add_diff_inverse2 length_drop zero_less_diff) (* 144 ms *)
then show ?thesis
using f6 f2 f1 Suc.hyps(2) by presburger (* 4 ms *)
qed
then have reg_tr: "trace (S t i) (t ! i # ?subt') (S t (j-1))"
proof -
have f1: "j - Suc i = Suc n"
using Suc.hyps(2) by presburger
have f2: "length (take j t) = j"
by (metis (no_types) Suc.prems(2) length_take min.absorb2 nat_le_linear not_less)
have f3: "(t ! i # drop (Suc i) (take j t)) @ [t ! j] = drop i (take (Suc j) t)"
by (metis (no_types) Suc_eq_plus1 \<open>take (j + 1 - i) (drop i t) = t ! i # take (j - (i + 1)) (drop (i + 1) t) @ [t ! j]\<close> append_Cons drop_take)
have f4: "Suc (i + n) = j - 1"
using f1 by (metis (no_types) Suc.prems(1) Suc_diff_Suc add_Suc_right diff_Suc_1 le_add_diff_inverse nat_le_linear not_less)
have "Suc (j - 1) = j"
using f1 by simp
then have f5: "butlast (take (Suc j) t) = take j t"
using f4 f3 f2 f1 by (metis (no_types) Groups.add_ac(2) One_nat_def append_eq_conv_conj append_take_drop_id butlast_take diff_Suc_1 drop_drop length_append length_drop list.size(3) list.size(4) order_refl plus_1_eq_Suc plus_nat.simps(2) take_add take_all)
have f6: "butlast (take j t) = take (j - 1) t"
by (meson Suc.prems(2) butlast_take nat_le_linear not_less)
have "drop (Suc i) (take j t) \<noteq> []"
by (metis (no_types) Nil_is_append_conv Suc_eq_plus1 \<open>take (j - (i + 1)) (drop (i + 1) t) = take (j - 1 - (i + 1)) (drop (i + 1) t) @ [t ! (j - 1)]\<close> drop_take list.distinct(1))
then show ?thesis
using f6 f5 f4 f3 by (metis (no_types) Suc.prems(6) Suc_eq_plus1 butlast.simps(2) butlast_drop butlast_snoc drop_take exists_trace_for_any_i_j less_add_Suc1 nat_le_linear not_less)
qed
have reg_st_1: "(S t (j-1)) \<turnstile> (t ! (j-1)) \<mapsto> (S t j)"
by (metis Suc.prems(1) Suc.prems(2) Suc.prems(6) Suc_lessD diff_Suc_1 less_imp_Suc_add step_Suc)
have "~ regular_event (t ! (j-1))"
using Suc.prems(3) \<open>take (j - (i + 1)) (drop (i + 1) t) \<noteq> []\<close> less_diff_conv by auto
moreover have "regular_event (t ! j)"
using Suc.prems(5) computation.prerecording_event computation_axioms by blast
moreover have "can_occur (t ! j) (S t j)"
using happen_implies_can_occur reg_tr_1 reg_st_2 by blast
moreover have njmiq: "occurs_on (t ! (j-1)) \<noteq> ?q"
proof (rule ccontr)
assume "~ occurs_on (t ! (j-1)) \<noteq> ?q"
then have "occurs_on (t ! (j-1)) = ?q" by simp
then have "has_snapshotted (S t j) ?q"
using Suc.prems(6) calculation(1) diff_le_self nonregular_event_induces_snapshot reg_st_1 snapshot_stable_ver_2 by blast
then show False using nsq by simp
qed
ultimately have "can_occur (t ! j) (S t (j-1))"
using reg_tr reg_st_1 event_can_go_back_if_no_sender by auto
then obtain d where new_st_1: "(S t (j-1)) \<turnstile> (t ! j) \<mapsto> d"
using exists_next_if_can_occur by blast
then have "trace (S t i) (t ! i # ?subt' @ [t ! j]) d" using reg_tr trace_snoc by fastforce
moreover have "can_occur (t ! (j-1)) d"
using \<open>(S t (j-1)) \<turnstile> t ! j \<mapsto> d\<close> \<open>occurs_on (t ! (j - 1)) \<noteq> occurs_on (t ! j)\<close> event_stays_valid_if_no_occurrence happen_implies_can_occur reg_st_1 by auto
moreover obtain e where new_st_2: "d \<turnstile> (t ! (j-1)) \<mapsto> e"
using calculation(2) exists_next_if_can_occur by blast
have pre_swap: "e = (S t (j+1))"
proof -
have "states e = states (S t (j+1))"
proof (rule ext)
fix p
have "states (S t (j-1)) p = states (S t j) p"
- using no_state_change_if_nonregular_event`~ regular_event (t ! (j-1))` reg_st_1 by auto
+ using no_state_change_if_nonregular_event\<open>~ regular_event (t ! (j-1))\<close> reg_st_1 by auto
moreover have "states d p = states e p"
- using no_state_change_if_nonregular_event`~ regular_event (t ! (j-1))` new_st_2 by auto
+ using no_state_change_if_nonregular_event\<open>~ regular_event (t ! (j-1))\<close> new_st_2 by auto
moreover have "states d p = states (S t (j+1)) p"
proof -
have "\<forall>a. states (S t (j + 1)) a = states d a"
by (meson \<open>\<not> regular_event (t ! (j - 1))\<close> new_st_1 no_state_change_if_nonregular_event reg_st_1 reg_st_2 same_state_implies_same_result_state)
then show ?thesis
by presburger
qed
ultimately show "states e p = states (S t (j+1)) p" by simp
qed
moreover have "msgs e = msgs (S t (j+1))"
proof (rule ext)
fix cid
have "isTrans (t ! j) \<or> isSend (t ! j) \<or> isRecv (t ! j)"
using \<open>regular_event (t ! j)\<close> by auto
moreover have "isSnapshot (t ! (j-1)) \<or> isRecvMarker (t ! (j-1))"
- using nonregular_event `~ regular_event (t ! (j-1))` by auto
+ using nonregular_event \<open>~ regular_event (t ! (j-1))\<close> by auto
ultimately show "msgs e cid = msgs (S t (j+1)) cid"
proof (elim disjE, goal_cases)
case 1
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Trans_Snapshot by auto
next
case 2
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Trans_RecvMarker by auto
next
case 3
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Send_Snapshot by auto
next
case 4
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_Recv_Snapshot by auto
next
case 5
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Send_RecvMarker by auto
next
case 6
then show ?case
using new_st_1 new_st_2 njmiq reg_st_1 reg_st_2 swap_msgs_Recv_RecvMarker by auto
qed
qed
moreover have "process_snapshot e = process_snapshot (S t (j+1))"
proof (rule ext)
fix p
have "process_snapshot (S t j) p = process_snapshot (S t (j+1)) p"
using \<open>regular_event (t ! j)\<close> reg_st_2 regular_event_preserves_process_snapshots by blast
moreover have "process_snapshot (S t (j-1)) p = process_snapshot d p"
using \<open>regular_event (t ! j)\<close> new_st_1 regular_event_preserves_process_snapshots by blast
moreover have "process_snapshot e p = process_snapshot (S t j) p"
proof -
have "occurs_on (t ! j) = p \<longrightarrow> ps e p = ps (S t j) p"
using calculation(2) new_st_2 njmiq no_state_change_if_no_event reg_st_1 by force
then show ?thesis
by (meson new_st_1 new_st_2 no_state_change_if_no_event reg_st_1 same_snapshot_state_implies_same_result_snapshot_state)
qed
ultimately show "process_snapshot e p = process_snapshot (S t (j+1)) p" by simp
qed
moreover have "cs e = cs (S t (j+1))"
proof (rule ext)
fix cid
have "isTrans (t ! j) \<or> isSend (t ! j) \<or> isRecv (t ! j)"
using \<open>regular_event (t ! j)\<close> by auto
moreover have "isSnapshot (t ! (j-1)) \<or> isRecvMarker (t ! (j-1))"
- using nonregular_event `~ regular_event (t ! (j-1))` by auto
+ using nonregular_event \<open>~ regular_event (t ! (j-1))\<close> by auto
ultimately show "cs e cid = cs (S t (j+1)) cid"
proof (elim disjE, goal_cases)
case 1
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Trans_Snapshot by auto
next
case 2
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Trans_RecvMarker by auto
next
case 3
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Send_Snapshot by auto
next
case 4
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Recv_Snapshot njmiq by auto
next
case 5
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Send_RecvMarker by auto
next
case 6
then show ?case
using new_st_1 new_st_2 reg_st_1 reg_st_2 swap_cs_Recv_RecvMarker njmiq by auto
qed
qed
ultimately show ?thesis by auto
qed
let ?it = "(t[j-1 := t ! j])[j := t ! (j-1)]"
have same_prefix: "take (j-1) ?it = take (j-1) t" by simp
have same_suffix: "drop (j+1) ?it = drop (j+1) t" by simp
have trace_prefix: "trace init (take (j-1) ?it) (S t (j-1))"
using Suc.prems(6) exists_trace_for_any_i by auto
have "?it = take (j-1) t @ [t ! j, t ! (j-1)] @ drop (j+1) t"
proof -
have "1 < j"
by (metis (no_types) Suc.hyps(2) Suc_eq_plus1 add_lessD1 plus_1_eq_Suc zero_less_Suc zero_less_diff) (* 12 ms *)
then have "j - 1 + 1 = j"
by (metis (no_types) le_add_diff_inverse2 nat_less_le) (* 4 ms *)
then show ?thesis
by (metis (no_types) Suc.prems(2) Suc_eq_plus1 add_Suc_right one_add_one swap_neighbors) (* 76 ms *)
qed
have "trace (S t (j-1)) [t ! j, t ! (j-1)] (S t (j+1))"
by (metis new_st_1 new_st_2 pre_swap trace.simps)
have "trace init (take (j+1) t @ drop (j+1) t) final"
by (simp add: Suc.prems(6))
then have "trace init (take (j+1) t) (S t (j+1)) \<and> trace (S t (j+1)) (drop (j+1) t) final"
using Suc.prems(6) exists_trace_for_any_i split_trace trace_and_start_determines_end by blast
then have trace_suffix: "trace (S t (j+1)) (drop (j+1) ?it) final" using same_suffix by simp
have "trace init ?it final"
by (metis (no_types, lifting) \<open>t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\<close> \<open>trace (S t (j + 1)) (drop (j + 1) (t[j - 1 := t ! j, j := t ! (j - 1)])) final\<close> \<open>trace (S t (j - 1)) [t ! j, t ! (j - 1)] (S t (j + 1))\<close> \<open>trace init (take (j - 1) (t[j - 1 := t ! j, j := t ! (j - 1)])) (S t (j - 1))\<close> same_prefix same_suffix trace_trans)
have suffix_same_states: "\<forall>k. k > j \<longrightarrow> S t k = S ?it k"
proof (rule allI, rule impI)
fix k
assume "k > j"
have eq_trace: "drop (j+1) t = drop (j+1) ?it" by simp
have "trace init (take (j+1) ?it) (S ?it (j+1))"
using \<open>trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\<close> exists_trace_for_any_i by blast
moreover have "trace init (take (j+1) ?it) (S t (j+1))"
proof -
have f1: "\<forall>es esa esb esc. (esb::('a, 'b, 'c) event list) @ es \<noteq> esa @ esc @ es \<or> esa @ esc = esb"
by auto
have f2: "take (j + 1) (t[j - 1 := t ! j, j := t ! (j - 1)]) @ drop (j + 1) t = t [j - 1 := t ! j, j := t ! (j - 1)]"
by (metis append_take_drop_id same_suffix)
have "trace init (take (j - 1) t @ [t ! j, t ! (j - 1)]) (S t (j + 1))"
using \<open>trace (S t (j - 1)) [t ! j, t ! (j - 1)] (S t (j + 1))\<close> same_prefix trace_prefix trace_trans by presburger
then show ?thesis
using f2 f1 by (metis (no_types) \<open>t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\<close>)
qed
ultimately have eq_start: "S ?it (j+1) = S t (j+1)"
using trace_and_start_determines_end by blast
then have "take k ?it = take (j+1) ?it @ take (k - (j+1)) (drop (j+1) ?it)"
by (metis Suc_eq_plus1 Suc_leI \<open>j < k\<close> le_add_diff_inverse take_add)
have "trace (S ?it (j+1)) (take (k - (j+1)) (drop (j+1) ?it)) (S ?it k)"
by (metis Suc_eq_plus1 Suc_leI \<open>j < k\<close> \<open>trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\<close> exists_trace_for_any_i_j)
moreover have "trace (S t (j+1)) (take (k - (j+1)) (drop (j+1) t)) (S t k)"
using Suc.prems(6) \<open>j < k\<close> exists_trace_for_any_i_j by fastforce
ultimately show "S t k = S ?it k"
using eq_start trace_and_start_determines_end by auto
qed
have prefix_same_states: "\<forall>k. k < j \<longrightarrow> S t k = S ?it k"
proof (rule allI, rule impI)
fix k
assume "k < j"
have "trace init (take k t) (S t k)"
using Suc.prems(6) exists_trace_for_any_i by blast
moreover have "trace init (take k ?it) (S ?it k)"
by (meson \<open>trace init (t[j - 1 := t ! j, j := t ! (j - 1)]) final\<close> exists_trace_for_any_i)
ultimately show "S t k = S ?it k"
using \<open>k < j\<close> s_def by auto
qed
moreover have "j - 1 < length ?it"
using Suc.prems(2) by auto
moreover have "prerecording_event ?it (j-1)"
proof -
have f1: "t[j - 1 := t ! j, j := t ! (j - 1)] ! (j - 1) = t[j - 1 := t ! j] ! (j - 1)"
by (metis (no_types) njmiq nth_list_update_neq) (* 28 ms *)
have "j \<noteq> 0"
by (metis (no_types) Suc.prems(1) not_less_zero) (* 0.0 ms *)
then have "\<not> j < 1"
by blast (* 0.0 ms *)
then have "S t (j - 1) = S (t[j - 1 := t ! j, j := t ! (j - 1)]) (j - 1)"
by (simp add: prefix_same_states) (* 8 ms *)
then show ?thesis
using f1 by (metis \<open>regular_event (t ! j)\<close> calculation(4) computation.prerecording_event computation_axioms length_list_update njmiq no_state_change_if_no_event nsq nth_list_update_eq reg_st_1) (* 456 ms *)
qed
moreover have "postrecording_event ?it i"
proof -
have "i < length ?it"
using Suc.prems(4) postrecording_event by auto
then show ?thesis
proof -
assume "i < length (t[j - 1 := t ! j, j := t ! (j - 1)])"
have "i < j - 1"
by (metis (no_types) Suc.hyps(2) cancel_ab_semigroup_add_class.diff_right_commute diff_diff_left zero_less_Suc zero_less_diff)
then show ?thesis
using Suc.prems(1) Suc.prems(4) postrecording_event prefix_same_states by auto
qed
qed
moreover have "i < j - 1"
using Suc.hyps(2) by auto
moreover have "\<forall>k. i < k \<and> k < (j-1) \<longrightarrow> ~ regular_event (?it ! k)"
proof (rule allI, rule impI)
fix k
assume "i < k \<and> k < (j-1)"
show "~ regular_event (?it ! k)"
using Suc.prems(3) \<open>i < k \<and> k < j - 1\<close> by force
qed
moreover have "(j-1) - (i+1) = n" using Suc.prems Suc.hyps by auto
ultimately have ind: "trace init (swap_events i (j-1) ?it) final
\<and> (\<forall>k. k \<ge> (j-1)+1 \<longrightarrow> S (swap_events i (j-1) ?it) k = S ?it k)
\<and> (\<forall>k. k \<le> i \<longrightarrow> S (swap_events i (j-1) ?it) k = S ?it k)
\<and> prerecording_event (swap_events i (j-1) ?it) i
\<and> postrecording_event (swap_events i (j-1) ?it) (i+1)
\<and> (\<forall>k. k > i+1 \<and> k < (j-1)+1 \<longrightarrow> ~ regular_event ((swap_events i (j-1) ?it) ! k))"
using Suc.hyps \<open>trace init ?it final\<close> by blast
then have new_trace: "trace init (swap_events i (j-1) ?it) final" by blast
have equal_suffix_states: "\<forall>k. k \<ge> j \<longrightarrow> S (swap_events i (j-1) ?it) k = S ?it k"
using Suc.prems(1) ind by simp
have equal_prefix_states: "\<forall>k. k \<le> i \<longrightarrow> S (swap_events i (j-1) ?it) k = S ?it k"
using ind by blast
have neighboring_events_shifted: "\<forall>k. k > i+1 \<and> k < j \<longrightarrow> ~ regular_event ((swap_events i (j-1) ?it) ! k)"
using ind by force
let ?itn = "swap_events i (j-1) ?it"
have "?itn = swap_events i j t"
proof -
have f1: "i \<le> j - 1"
using \<open>i < j - 1\<close> less_imp_le_nat by blast
have "t ! j # [t ! (j - 1)] @ drop (j + 1) t = drop (j - 1) (take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t)"
using \<open>t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\<close> same_prefix by force
then have f2: "t[j - 1 := t ! j, j := t ! (j - 1)] ! (j - 1) = t ! j \<and> drop (j - 1 + 1) (t[j - 1 := t ! j, j := t ! (j - 1)]) = t ! (j - 1) # [] @ drop (j + 1) t"
by (metis (no_types) Cons_nth_drop_Suc Suc_eq_plus1 \<open>j - 1 < length (t[j - 1 := t ! j, j := t ! (j - 1)])\<close> \<open>t[j - 1 := t ! j, j := t ! (j - 1)] = take (j - 1) t @ [t ! j, t ! (j - 1)] @ drop (j + 1) t\<close> append_Cons list.inject)
have "t ! i = t[j - 1 := t ! j, j := t ! (j - 1)] ! i"
by (metis (no_types) Suc.prems(1) \<open>i < j - 1\<close> nat_neq_iff nth_list_update_neq)
then show ?thesis
using f2 f1 by (metis (no_types) Suc.prems(1) \<open>take (j - (i + 1)) (drop (i + 1) t) = take (j - 1 - (i + 1)) (drop (i + 1) t) @ [t ! (j - 1)]\<close> append.assoc append_Cons drop_take less_imp_le_nat same_prefix take_update_cancel)
qed
moreover have "\<forall>k. k \<le> i \<longrightarrow> S t k = S ?itn k"
using Suc.prems(1) equal_prefix_states prefix_same_states by auto
moreover have "\<forall>k. k \<ge> j + 1 \<longrightarrow> S t k = S ?itn k"
by (metis (no_types, lifting) Suc_eq_plus1 add_lessD1 equal_suffix_states lessI nat_less_le suffix_same_states)
moreover have "\<forall>k. k > i+1 \<and> k < j+1 \<longrightarrow> ~ regular_event (?itn ! k)"
proof -
have "~ regular_event (?itn ! j)"
proof -
have f1: "j - 1 < length t"
using \<open>j - 1 < length (t[j - 1 := t ! j, j := t ! (j - 1)])\<close> by force
have f2: "\<And>n na es. \<not> n < na \<or> \<not> na < length es \<or> drop (Suc na) (take n es @ [hd (drop na es), es ! n::('a, 'b, 'c) event] @ take (na - Suc n) (drop (Suc n) es) @ drop (Suc na) es) = drop (Suc na) es"
by (metis Suc_eq_plus1 hd_drop_conv_nth swap_identical_tails)
have f3: "t ! j = hd (drop j t)"
by (simp add: Suc.prems(2) hd_drop_conv_nth)
have "\<not> j < 1"
using Suc.prems(1) by blast
then have "\<not> regular_event (hd (drop j (take i (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)]) @ [hd (drop (j - 1) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)])), t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)] ! i] @ take (j - 1 - Suc i) (drop (Suc i) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)])) @ drop (Suc (j - 1)) (t[j - 1 := hd (drop j t), j := hd (drop (j - 1) t)]))))"
using f2 f1 by (metis (no_types) Suc.prems(2) \<open>\<not> regular_event (t ! (j - 1))\<close> \<open>i < j - 1\<close> add_diff_inverse_nat hd_drop_conv_nth length_list_update nth_list_update_eq plus_1_eq_Suc)
then show ?thesis
using f3 f1 by (metis Suc.prems(2) Suc_eq_plus1 \<open>i < j - 1\<close> hd_drop_conv_nth length_list_update swap_identical_length)
qed
then show ?thesis
by (metis Suc_eq_plus1 less_Suc_eq neighboring_events_shifted)
qed
ultimately show ?case using ind by presburger
qed
subsubsection \<open>Relating configurations and the computed snapshot\<close>
definition ps_equal_to_snapshot where
"ps_equal_to_snapshot c c' \<equiv>
\<forall>p. Some (states c p) = process_snapshot c' p"
definition cs_equal_to_snapshot where
"cs_equal_to_snapshot c c' \<equiv>
\<forall>cid. channel cid \<noteq> None
\<longrightarrow> filter ((\<noteq>) Marker) (msgs c cid)
= map Msg (fst (channel_snapshot c' cid))"
definition state_equal_to_snapshot where
"state_equal_to_snapshot c c' \<equiv>
ps_equal_to_snapshot c c' \<and> cs_equal_to_snapshot c c'"
lemma init_is_s_t_0:
assumes
"trace init t final"
shows
"init = (S t 0)"
by (metis assms exists_trace_for_any_i take_eq_Nil tr_init trace_and_start_determines_end)
lemma final_is_s_t_len_t:
assumes
"trace init t final"
shows
"final = S t (length t)"
by (metis assms exists_trace_for_any_i order_refl take_all trace_and_start_determines_end)
lemma snapshot_event:
assumes
"trace init t final" and
"~ has_snapshotted (S t i) p" and
"has_snapshotted (S t (i+1)) p"
shows
"isSnapshot (t ! i) \<or> isRecvMarker (t ! i)"
proof -
have "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis Suc_eq_plus1 assms(1) assms(2) assms(3) distributed_system.step_Suc computation_axioms computation_def nat_less_le not_less not_less_eq s_def take_all)
then show ?thesis
using assms(2) assms(3) nonregular_event regular_event_cannot_induce_snapshot by blast
qed
lemma snapshot_state:
assumes
"trace init t final" and
"states (S t i) p = u" and
"~ has_snapshotted (S t i) p" and
"has_snapshotted (S t (i+1)) p"
shows
"ps (S t (i+1)) p = Some u"
proof -
have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis add.commute assms(1) assms(3) assms(4) le_SucI le_eq_less_or_eq le_refl nat_neq_iff no_change_if_ge_length_t plus_1_eq_Suc step_Suc)
let ?q = "occurs_on (t ! i)"
have qp: "?q = p"
proof (rule ccontr)
assume "?q \<noteq> p"
then have "has_snapshotted (S t (i+1)) p = has_snapshotted (S t i) p"
using local.step no_state_change_if_no_event by auto
then show False using assms by simp
qed
have "isSnapshot (t ! i) \<or> isRecvMarker (t ! i)" using assms snapshot_event by auto
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have "t ! i = Snapshot p"
by (metis event.collapse(4) qp)
then show ?thesis
using assms(2) local.step by auto
next
case 2
then obtain cid' q where "t ! i = RecvMarker cid' p q"
by (metis event.collapse(5) qp)
then show ?thesis using assms step by auto
qed
qed
lemma snapshot_state_unchanged_trace_2:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
ps (S t i) p = Some u
\<rbrakk> \<Longrightarrow> ps (S t j) p = Some u"
proof (induct i j rule:S_induct)
case S_init
then show ?case by simp
next
case S_step
then show ?case using snapshot_state_unchanged by auto
qed
lemma no_recording_cs_if_not_snapshotted:
shows
"\<lbrakk> trace init t final; ~ has_snapshotted (S t i) p;
channel cid = Some (q, p) \<rbrakk> \<Longrightarrow> cs (S t i) cid = cs init cid"
proof (induct i)
case 0
then show ?case
by (metis exists_trace_for_any_i list.discI take_eq_Nil trace.simps)
next
case (Suc i)
have "Suc i < length t"
proof -
have "has_snapshotted final p"
using all_processes_snapshotted_in_final_state valid by blast
show ?thesis
proof (rule ccontr)
assume "~ Suc i < length t"
then have "Suc i \<ge> length t" by simp
then have "has_snapshotted (S t (Suc i)) p"
using Suc.prems(1) \<open>ps final p \<noteq> None\<close> final_is_s_t_len_t snapshot_stable_ver_3 by blast
then show False using Suc by simp
qed
qed
then have t_dec: "trace init (take i t) (S t i) \<and> (S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i))"
using Suc.prems(1) exists_trace_for_any_i step_Suc by auto
moreover have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i))" using calculation by simp
ultimately have IH: "cs (S t i) cid = cs init cid"
using Suc.hyps Suc.prems(1) Suc.prems(2) Suc.prems(3) snapshot_state_unchanged by fastforce
then show ?case
proof (cases "t ! i")
case (Snapshot r)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
then have "r = p" by simp
then have "has_snapshotted (S t (Suc i)) p"
using Snapshot step by auto
then show False using Suc by simp
qed
then have "cs (S t i) cid = cs (S t (Suc i)) cid"
using Snapshot Suc.prems(3) local.step by auto
then show ?thesis using IH by simp
next
case (RecvMarker cid' r s)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
then have "r = p" by simp
then have "has_snapshotted (S t (Suc i)) p"
using RecvMarker t_dec recv_marker_means_snapshotted_1 by blast
then show False using Suc by simp
qed
have "cid' \<noteq> cid"
proof (rule ccontr)
assume "~ cid' \<noteq> cid"
then have "channel cid' = Some (s, r)" using t_dec can_occur_def RecvMarker by simp
then show False
using Suc.prems(3) \<open>\<not> cid' \<noteq> cid\<close> \<open>r \<noteq> p\<close> by auto
qed
then have "cs (S t i) cid = cs (S t (Suc i)) cid"
proof -
- have "\<nexists>s. channel cid = Some (s, r)" using `r \<noteq> p` Suc by simp
- with RecvMarker t_dec `cid' \<noteq> cid` `r \<noteq> p` Suc.prems(3) show ?thesis
+ have "\<nexists>s. channel cid = Some (s, r)" using \<open>r \<noteq> p\<close> Suc by simp
+ with RecvMarker t_dec \<open>cid' \<noteq> cid\<close> \<open>r \<noteq> p\<close> Suc.prems(3) show ?thesis
by (cases "has_snapshotted (S t i) r", auto)
qed
then show ?thesis using IH by simp
next
case (Trans r u u')
then show ?thesis
using IH t_dec by auto
next
case (Send cid' r s u u' m)
then show ?thesis
using IH local.step by auto
next
case (Recv cid' r s u u' m)
then have "snd (cs (S t i) cid) = NotStarted"
by (simp add: IH no_initial_channel_snapshot)
with Recv step Suc show ?thesis by (cases "cid' = cid", auto)
qed
qed
lemma cs_done_implies_has_snapshotted:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = Done" and
"channel cid = Some (p, q)"
shows
"has_snapshotted (S t i) q"
proof -
show ?thesis
using assms no_initial_channel_snapshot no_recording_cs_if_not_snapshotted by fastforce
qed
lemma exactly_one_snapshot:
assumes
"trace init t final"
shows
"\<exists>!i. ~ has_snapshotted (S t i) p \<and> has_snapshotted (S t (i+1)) p" (is ?P)
proof -
have "~ has_snapshotted init p"
using no_initial_process_snapshot by auto
moreover have "has_snapshotted final p"
using all_processes_snapshotted_in_final_state valid by blast
moreover have "trace (S t 0) t (S t (length t))"
using assms final_is_s_t_len_t init_is_s_t_0 by auto
ultimately have ex_snap: "\<exists>i. ~ has_snapshotted (S t i) p \<and> has_snapshotted (S t (i+1)) p"
using assms exists_snapshot_for_all_p by auto
show ?thesis
proof (rule ccontr)
assume "~ ?P"
then have "\<exists>i j. (i \<noteq> j) \<and> ~ has_snapshotted (S t i) p \<and> has_snapshotted (S t (i+1)) p \<and>
~ has_snapshotted (S t j) p \<and> has_snapshotted (S t (j+1)) p"
using ex_snap by blast
then have "\<exists>i j. (i < j) \<and> ~ has_snapshotted (S t i) p \<and> has_snapshotted (S t (i+1)) p \<and>
~ has_snapshotted (S t j) p \<and> has_snapshotted (S t (j+1)) p"
by (meson linorder_neqE_nat)
then obtain i j where "i < j" "~ has_snapshotted (S t i) p" "has_snapshotted (S t (i+1)) p"
"~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p"
by blast
have "trace (S t (i+1)) (take (j - (i+1)) (drop (i+1) t)) (S t j)"
using \<open>i < j\<close> assms exists_trace_for_any_i_j by fastforce
then have "has_snapshotted (S t j) p"
using \<open>ps (S t (i + 1)) p \<noteq> None\<close> snapshot_stable by blast
- then show False using `~ has_snapshotted (S t j) p` by simp
+ then show False using \<open>~ has_snapshotted (S t j) p\<close> by simp
qed
qed
lemma initial_cs_changes_implies_nonregular_event:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = NotStarted" and
"snd (cs (S t (i+1)) cid) \<noteq> NotStarted" and
"channel cid = Some (p, q)"
shows
"~ regular_event (t ! i)"
proof -
have "i < length t"
proof (rule ccontr)
assume "~ i < length t"
then have "S t i = S t (i+1)"
using assms(1) no_change_if_ge_length_t by auto
then show False using assms by presburger
qed
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
using assms(1) step_Suc by auto
show ?thesis
proof (rule ccontr)
assume "~ ~ regular_event (t ! i)"
then have "regular_event (t ! i)" by simp
then have "cs (S t i) cid = cs (S t (i+1)) cid"
proof (cases "isRecv (t ! i)")
case False
then show ?thesis
using \<open>regular_event (t ! i)\<close> local.step no_cs_change_if_no_event by blast
next
case True
then obtain cid' r s u u' m where Recv: "t ! i = Recv cid' r s u u' m" by (meson isRecv_def)
with assms step show ?thesis
proof (cases "cid = cid'")
case True
then show ?thesis using assms step Recv by simp
next
case False
then show ?thesis using assms step Recv by simp
qed
qed
then show False using assms by simp
qed
qed
lemma cs_in_initial_state_implies_not_snapshotted:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = NotStarted" and
"channel cid = Some (p, q)"
shows
"~ has_snapshotted (S t i) q"
proof (rule ccontr)
assume "~ ~ has_snapshotted (S t i) q"
then obtain j where "j < i" "~ has_snapshotted (S t j) q" "has_snapshotted (S t (j+1)) q"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p computation.snapshot_stable_ver_3 computation_axioms nat_le_linear order_le_less)
have step_j: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis \<open>\<not> \<not> ps (S t i) q \<noteq> None\<close> \<open>\<not> ps (S t j) q \<noteq> None\<close> \<open>j < i\<close> add.commute assms(1) linorder_neqE_nat no_change_if_ge_length_t order_le_less order_refl plus_1_eq_Suc step_Suc)
have tr_j_i: "trace (S t (j+1)) (take (i - (j+1)) (drop (j+1) t)) (S t i)"
using \<open>j < i\<close> assms(1) exists_trace_for_any_i_j by fastforce
have "~ regular_event (t ! j)"
using step_j \<open>\<not> ps (S t j) q \<noteq> None\<close> \<open>ps (S t (j + 1)) q \<noteq> None\<close> regular_event_cannot_induce_snapshot by blast
then have "isSnapshot (t ! j) \<or> isRecvMarker (t ! j)"
using nonregular_event by auto
then have "snd (cs (S t (j+1)) cid) \<noteq> NotStarted"
proof (elim disjE, goal_cases)
case 1
have "occurs_on (t ! j) = q"
using \<open>\<not> ps (S t j) q \<noteq> None\<close> \<open>ps (S t (j + 1)) q \<noteq> None\<close> distributed_system.no_state_change_if_no_event distributed_system_axioms step_j by fastforce
with 1 have "t ! j = Snapshot q" using isSnapshot_def by auto
then show ?thesis using step_j assms by simp
next
case 2
have "occurs_on (t ! j) = q"
using \<open>\<not> ps (S t j) q \<noteq> None\<close> \<open>ps (S t (j + 1)) q \<noteq> None\<close> distributed_system.no_state_change_if_no_event distributed_system_axioms step_j by fastforce
with 2 obtain cid' s where RecvMarker: "t ! j = RecvMarker cid' q s"
by (metis event.collapse(5))
then show ?thesis
proof (cases "cid' = cid")
case True
then show ?thesis using RecvMarker step_j assms by simp
next
case False
have "~ has_snapshotted (S t j) q"
using \<open>\<not> ps (S t j) q \<noteq> None\<close> by auto
moreover have "\<exists>r. channel cid = Some (r, q)"
by (simp add: assms(3))
ultimately show ?thesis using RecvMarker step_j assms False by simp
qed
qed
then have "snd (cs (S t i) cid) \<noteq> NotStarted"
using tr_j_i cs_not_not_started_stable_trace assms by blast
then show False using assms by simp
qed
lemma nonregular_event_in_initial_state_implies_cs_changed:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = NotStarted" and
"~ regular_event (t ! i)" and
"occurs_on (t ! i) = q" and
"channel cid = Some (p, q)" and
"i < length t"
shows
"snd (cs (S t (i+1)) cid) \<noteq> NotStarted"
proof -
have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))" using step_Suc assms by auto
have "isSnapshot (t ! i) \<or> isRecvMarker (t ! i)"
using assms(3) nonregular_event by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then show ?thesis
using assms cs_in_initial_state_implies_not_snapshotted local.step nonregular_event_induces_snapshot by blast
next
case 2
then show ?thesis
by (metis assms(1) assms(2) assms(3) assms(4) assms(5) cs_in_initial_state_implies_not_snapshotted local.step nonregular_event_induces_snapshot)
qed
qed
lemma cs_recording_implies_snapshot:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = Recording" and
"channel cid = Some (p, q)"
shows
"has_snapshotted (S t i) q"
proof (rule ccontr)
assume "~ has_snapshotted (S t i) q"
have "\<lbrakk> trace init t final; ~ has_snapshotted (S t i) p; channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> snd (cs (S t i) cid) = NotStarted"
proof (induct i)
case 0
then show ?case
using init_is_s_t_0 no_initial_channel_snapshot by auto
next
case (Suc n)
have step: "(S t n) \<turnstile> (t ! n) \<mapsto> (S t (n+1))"
by (metis Suc.prems(2) Suc_eq_plus1 all_processes_snapshotted_in_final_state assms(1) distributed_system.step_Suc distributed_system_axioms final_is_s_t_len_t le_add1 not_less snapshot_stable_ver_3)
have "snd (cs (S t n) cid) = NotStarted"
using Suc.hyps Suc.prems(2) assms snapshot_state_unchanged computation_axioms local.step by fastforce
then show ?case
by (metis Suc.prems(1) \<open>\<not> ps (S t i) q \<noteq> None\<close> assms(2) assms(3) cs_not_not_started_stable_trace exists_trace_for_any_i no_recording_cs_if_not_snapshotted recording_state.simps(2))
qed
then show False
using \<open>\<not> ps (S t i) q \<noteq> None\<close> assms computation.no_initial_channel_snapshot computation_axioms no_recording_cs_if_not_snapshotted by fastforce
qed
lemma cs_done_implies_both_snapshotted:
assumes
"trace init t final" and
"snd (cs (S t i) cid) = Done" and
"i < length t" and
"channel cid = Some (p, q)"
shows
"has_snapshotted (S t i) p"
"has_snapshotted (S t i) q"
proof -
have "trace init (take i t) (S t i)"
using assms(1) exists_trace_for_any_i by blast
then have "RecvMarker cid q p : set (take i t)"
by (metis assms(1,2,4) cs_done_implies_has_snapshotted done_only_from_recv_marker_trace computation.no_initial_process_snapshot computation_axioms init_is_s_t_0 list.discI trace.simps)
then obtain k where "t ! k = RecvMarker cid q p" "0 \<le> k" "k < i"
by (metis add.right_neutral add_diff_cancel_right' append_Nil append_take_drop_id assms(1) exists_index take0)
then have "has_snapshotted (S t (k+1)) q"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_leI assms(1,2,4) computation.cs_done_implies_has_snapshotted computation.no_change_if_ge_length_t computation_axioms less_le not_less_eq recv_marker_means_cs_Done)
then show "has_snapshotted (S t i) q"
using assms cs_done_implies_has_snapshotted by blast
have step_k: "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (k+1))"
by (metis Suc_eq_plus1 \<open>k < i\<close> add_lessD1 assms(1) assms(3) distributed_system.step_Suc distributed_system_axioms less_imp_add_positive)
then have "Marker : set (msgs (S t k) cid)"
proof -
have "can_occur (t ! k) (S t k)" using happen_implies_can_occur step_k by blast
- then show ?thesis unfolding can_occur_def `t ! k = RecvMarker cid q p`
+ then show ?thesis unfolding can_occur_def \<open>t ! k = RecvMarker cid q p\<close>
using hd_in_set by fastforce
qed
then have "has_snapshotted (S t k) p"
using assms(1,4) no_marker_if_no_snapshot by blast
then show "has_snapshotted (S t i) p"
using \<open>k < i\<close> assms(1) less_imp_le_nat snapshot_stable_ver_3 by blast
qed
lemma cs_done_implies_same_snapshots:
assumes "trace init t final" "i \<le> j" "j \<le> length t"
shows "snd (cs (S t i) cid) = Done \<Longrightarrow> channel cid = Some (p, q) \<Longrightarrow> cs (S t i) cid = cs (S t j) cid"
using assms proof (induct i j rule: S_induct)
case (S_init i)
then show ?case by auto
next
case (S_step i j)
have snap_p: "has_snapshotted (S t i) p"
using S_step.hyps(1) S_step.hyps(2) S_step.prems(1,2) assms(1) cs_done_implies_both_snapshotted(1) by auto
have snap_q: "has_snapshotted (S t i) q"
using S_step.prems(1,2) assms(1) cs_done_implies_has_snapshotted by blast
from S_step have "cs (S t i) cid = cs (S t (Suc i)) cid"
proof (cases "t ! i")
case (Snapshot r)
from Snapshot S_step.hyps(3) snap_p have False if "r = p" using that by (auto simp: can_occur_def)
moreover
from Snapshot S_step.hyps(3) snap_q have False if "r = q" using that by (auto simp: can_occur_def)
ultimately show ?thesis using Snapshot S_step by force
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "has_snapshotted (S t i) r")
case True
with RecvMarker S_step show ?thesis
proof (cases "cid = cid'")
case True
then have "cs (S t (Suc i)) cid = (fst (cs (S t i) cid), Done)"
using RecvMarker S_step by simp
then show ?thesis
by (metis S_step.prems(1) prod.collapse)
qed auto
next
case no_snap: False
then show ?thesis
proof (cases "cid = cid'")
case True
then have "cs (S t (Suc i)) cid = (fst (cs (S t i) cid), Done)"
using RecvMarker S_step by simp
then show ?thesis
by (metis S_step.prems(1) prod.collapse)
next
case False
then have "r \<noteq> p" using no_snap snap_p by auto
moreover have "\<nexists>s. channel cid = Some (s, r)"
using S_step(5) assms(1) cs_done_implies_has_snapshotted no_snap by blast
ultimately show ?thesis using RecvMarker S_step False no_snap by simp
qed
qed
next
case (Recv cid' r s u u' m)
with S_step show ?thesis by (cases "cid = cid'", auto)
qed auto
with S_step show ?case by auto
qed
lemma snapshotted_and_not_done_implies_marker_in_channel:
assumes
"trace init t final" and
"has_snapshotted (S t i) p" and
"snd (cs (S t i) cid) \<noteq> Done" and
"i \<le> length t" and
"channel cid = Some (p, q)"
shows
"Marker : set (msgs (S t i) cid)"
proof -
obtain j where jj: "j < i" "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 assms(1) assms(2) exists_snapshot_for_all_p computation.snapshot_stable_ver_2 computation_axioms le_eq_less_or_eq nat_neq_iff)
have step: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis \<open>\<not> ps (S t j) p \<noteq> None\<close> \<open>j < i\<close> add.commute assms(1) assms(2) linorder_neqE_nat no_change_if_ge_length_t order_le_less order_refl plus_1_eq_Suc step_Suc)
then have "Marker : set (msgs (S t (j+1)) cid)"
proof -
have "~ regular_event (t ! j)"
by (meson \<open>\<not> ps (S t j) p \<noteq> None\<close> \<open>ps (S t (j + 1)) p \<noteq> None\<close> distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms local.step)
then have "isSnapshot (t ! j) \<or> isRecvMarker (t ! j)" using nonregular_event by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then obtain r where Snapshot: "t ! j = Snapshot r" by (meson isSnapshot_def)
then have "r = p"
using jj(2) jj(3) local.step by auto
then show ?thesis using Snapshot assms step by simp
next
case 2
then obtain cid' s where RecvMarker: "t ! j = RecvMarker cid' p s"
by (metis jj(2,3) distributed_system.no_state_change_if_no_event distributed_system_axioms event.sel(5) isRecvMarker_def local.step)
moreover have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "snd (cs (S t (j+1)) cid) = Done" using RecvMarker step by simp
then have "snd (cs (S t i) cid) = Done"
proof -
assume a1: "snd (cs (S t (j + 1)) cid) = Done"
have f2: "ps (S t j) p = None"
using jj(2) by blast
have "j < length t"
using assms(4) jj(1) by linarith
then have "t ! j = RecvMarker cid q p"
using f2 a1 assms(1) assms(5) cs_done_implies_both_snapshotted(1) done_only_from_recv_marker local.step by blast
then show ?thesis
using f2 by (metis (no_types) Suc_eq_plus1 assms(1) local.step recv_marker_means_snapshotted)
qed
then show False using assms by simp
qed
ultimately show ?thesis using jj assms step by auto
qed
qed
show ?thesis
proof (rule ccontr)
let ?t = "take (i - (j+1)) (drop (j+1) t)"
have tr_j: "trace (S t (j+1)) ?t (S t i)"
by (metis \<open>j < i\<close> assms(1) discrete exists_trace_for_any_i_j)
assume "~ Marker : set (msgs (S t i) cid)"
then obtain ev where "ev \<in> set ?t" "\<exists>p q. ev = RecvMarker cid p q"
using \<open>Marker \<in> set (msgs (S t (j + 1)) cid)\<close> marker_must_be_delivered_2_trace tr_j assms by blast
obtain k where "t ! k = ev" "j < k" "k < i"
using \<open>ev \<in> set (take (i - (j + 1)) (drop (j + 1) t))\<close> assms(1) exists_index by fastforce
have step_k: "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (k+1))"
proof -
have "k < length t"
using \<open>k < i\<close> assms(4) by auto
then show ?thesis using step_Suc assms by simp
qed
have "ev = RecvMarker cid q p" using assms step_k can_occur_def
using \<open>\<exists>p q. ev = RecvMarker cid p q\<close> \<open>t ! k = ev\<close> by auto
then have "snd (cs (S t (k+1)) cid) = Done"
using \<open>k < i\<close> \<open>t ! k = ev\<close> assms(1) assms(4) recv_marker_means_cs_Done by auto
moreover have "trace (S t (k+1)) (take (i - (k+1)) (drop (k+1) t)) (S t i)"
by (meson \<open>k < i\<close> assms(1) discrete exists_trace_for_any_i_j)
ultimately have "snd (cs (S t i) cid) = Done"
by (metis \<open>k < i\<close> assms(1) assms(4) assms(5) cs_done_implies_same_snapshots discrete)
then show False using assms by simp
qed
qed
lemma no_marker_left_in_final_state:
assumes
"trace init t final"
shows
"Marker \<notin> set (msgs final cid)" (is ?P)
proof (rule ccontr)
assume "~ ?P"
then obtain i where "i > length t" "Marker \<notin> set (msgs (S t i) cid)" using assms l1
by (metis final_is_s_t_len_t le_neq_implies_less)
then have "S t (length t) \<noteq> S t i"
proof -
have "msgs (S t i) cid \<noteq> msgs final cid"
- using \<open>Marker \<notin> set (msgs (S t i) cid)\<close> `~ ?P` by auto
+ using \<open>Marker \<notin> set (msgs (S t i) cid)\<close> \<open>~ ?P\<close> by auto
then show ?thesis using final_is_s_t_len_t assms by auto
qed
moreover have "S t (length t) = S t i"
- using assms `i > length t` less_imp_le no_change_if_ge_length_t by simp
+ using assms \<open>i > length t\<close> less_imp_le no_change_if_ge_length_t by simp
ultimately show False by simp
qed
lemma all_channels_done_in_final_state:
assumes
"trace init t final" and
"channel cid = Some (p, q)"
shows
"snd (cs final cid) = Done"
proof (rule ccontr)
assume cs_not_done: "~ snd (cs final cid) = Done"
obtain i where snap_p: "~ has_snapshotted (S t i) p" "has_snapshotted (S t (i+1)) p"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have "i < length t"
proof -
have "S t i \<noteq> S t (i+1)" using snap_p by auto
then show ?thesis
by (meson assms(1) computation.no_change_if_ge_length_t computation_axioms le_add1 not_less)
qed
let ?t = "take (length t - (i+1)) (drop (i+1) t)"
have tr: "trace (S t (i+1)) ?t (S t (length t))"
by (meson \<open>i < length t\<close> assms(1) discrete exists_trace_for_any_i_j)
have "Marker \<in> set (msgs (S t (i+1)) cid)"
proof -
have n_done: "snd (cs (S t (i+1)) cid) \<noteq> Done"
proof (rule ccontr)
assume "~ snd (cs (S t (i+1)) cid) \<noteq> Done"
then have "snd (cs final cid) = Done"
by (metis Suc_eq_plus1 Suc_leI \<open>i < length t\<close> assms final_is_s_t_len_t computation.cs_done_implies_same_snapshots computation_axioms order_refl)
then show False using cs_not_done by simp
qed
then show ?thesis using snapshotted_and_not_done_implies_marker_in_channel snap_p assms
proof -
- have "i+1 \<le> length t" using `i < length t` by auto
+ have "i+1 \<le> length t" using \<open>i < length t\<close> by auto
then show ?thesis
using snapshotted_and_not_done_implies_marker_in_channel snap_p assms n_done by simp
qed
qed
moreover have "Marker \<notin> set (msgs (S t (length t)) cid)" using final_is_s_t_len_t no_marker_left_in_final_state assms by blast
ultimately have rm_prov: "\<exists>ev \<in> set ?t. (\<exists>q p. ev = RecvMarker cid q p)" using tr message_must_be_delivered_2_trace assms
by (simp add: marker_must_be_delivered_2_trace)
then obtain k where "\<exists>q p. t ! k = RecvMarker cid q p" "i+1 \<le> k" "k < length t"
by (metis assms(1) exists_index)
then have step: "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (k+1))"
by (metis Suc_eq_plus1_left add.commute assms(1) step_Suc)
then have RecvMarker: "t ! k = RecvMarker cid q p"
by (metis RecvMarker_given_channel \<open>\<exists>q p. t ! k = RecvMarker cid q p\<close> assms(2) event.disc(25) event.sel(10) happen_implies_can_occur)
then have "snd (cs (S t (k+1)) cid) = Done"
using step \<open>k < length t\<close> assms(1) recv_marker_means_cs_Done by blast
then have "snd (cs final cid) = Done"
using \<open>Marker \<notin> set (msgs (S t (length t)) cid)\<close> all_processes_snapshotted_in_final_state assms(1) assms(2) final_is_s_t_len_t snapshotted_and_not_done_implies_marker_in_channel by fastforce
then show False using cs_not_done by simp
qed
lemma cs_NotStarted_implies_empty_cs:
shows
"\<lbrakk> trace init t final; channel cid = Some (p, q); i < length t; ~ has_snapshotted (S t i) q \<rbrakk>
\<Longrightarrow> cs (S t i) cid = ([], NotStarted)"
by (simp add: no_initial_channel_snapshot no_recording_cs_if_not_snapshotted)
lemma fst_changed_by_recv_recording_trace:
assumes
"i < j" and
"j \<le> length t" and
"trace init t final" and
"fst (cs (S t i) cid) \<noteq> fst (cs (S t j) cid)" and
"channel cid = Some (p, q)"
shows
"\<exists>k. i \<le> k \<and> k < j \<and> (\<exists>p q u u' m. t ! k = Recv cid q p u u' m) \<and> (snd (cs (S t k) cid) = Recording)" (is ?P)
proof (rule ccontr)
assume "~ ?P"
have "\<lbrakk> i < j; j \<le> length t; ~ ?P; trace init t final; channel cid = Some (p, q) \<rbrakk> \<Longrightarrow> fst (cs (S t i) cid) = fst (cs (S t j) cid)"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by linarith
next
case (Suc n)
then have step: "(S t i) \<turnstile> t ! i \<mapsto> (S t (Suc i))"
using step_Suc by auto
then have "fst (cs (S t (Suc i)) cid) = fst (cs (S t i) cid)"
by (metis Suc.prems(1) Suc.prems(3) assms(5) fst_cs_changed_by_recv_recording le_eq_less_or_eq)
also have "fst (cs (S t (Suc i)) cid) = fst (cs (S t j) cid)"
proof -
have "j - Suc i = n" using Suc by simp
moreover have "~ (\<exists>k. (Suc i) \<le> k \<and> k < j \<and> (\<exists>p q u u' m. t ! k = Recv cid q p u u' m) \<and> (snd (cs (S t k) cid) = Recording))"
- using `~ ?P` Suc.prems(3) Suc_leD by blast
+ using \<open>~ ?P\<close> Suc.prems(3) Suc_leD by blast
ultimately show ?thesis using Suc by (metis Suc_lessI)
qed
finally show ?case by simp
qed
- then show False using assms `~ ?P` by blast
+ then show False using assms \<open>~ ?P\<close> by blast
qed
lemma cs_not_nil_implies_postrecording_event:
assumes
"trace init t final" and
"fst (cs (S t i) cid) \<noteq> []" and
"i \<le> length t" and
"channel cid = Some (p, q)"
shows
"\<exists>j. j < i \<and> postrecording_event t j"
proof -
have "fst (cs init cid) = []" using no_initial_channel_snapshot by auto
then have diff_cs: "fst (cs (S t 0) cid) \<noteq> fst (cs (S t i) cid)"
using assms(1) assms(2) init_is_s_t_0 by auto
moreover have "0 < i"
proof (rule ccontr)
assume "~ 0 < i"
then have "0 = i" by auto
then have "fst (cs (S t 0) cid) = fst (cs (S t i) cid)"
by blast
then show False using diff_cs by simp
qed
ultimately obtain j where "j < i" and Recv: "\<exists>p q u u' m. t ! j = Recv cid q p u u' m" "snd (cs (S t j) cid) = Recording"
using assms(1) assms(3) assms(4) fst_changed_by_recv_recording_trace by blast
then have "has_snapshotted (S t j) q"
using assms(1) assms(4) cs_recording_implies_snapshot by blast
moreover have "regular_event (t ! j)" using Recv by auto
moreover have "occurs_on (t ! j) = q"
proof -
have "can_occur (t ! j) (S t j)"
by (meson Suc_le_eq \<open>j < i\<close> assms(1) assms(3) happen_implies_can_occur le_trans step_Suc)
then show ?thesis using Recv Recv_given_channel assms(4) by force
qed
- ultimately have "postrecording_event t j" unfolding postrecording_event using `j < i` assms(3) by simp
- then show ?thesis using `j < i` by auto
+ ultimately have "postrecording_event t j" unfolding postrecording_event using \<open>j < i\<close> assms(3) by simp
+ then show ?thesis using \<open>j < i\<close> by auto
qed
subsubsection \<open>Relating process states\<close>
lemma snapshot_state_must_have_been_reached:
assumes
"trace init t final" and
"ps final p = Some u" and
"~ has_snapshotted (S t i) p" and
"has_snapshotted (S t (i+1)) p" and
"i < length t"
shows
"states (S t i) p = u"
proof (rule ccontr)
assume "states (S t i) p \<noteq> u"
then have "ps (S t (i+1)) p \<noteq> Some u"
using assms(1) assms(3) snapshot_state by force
then have "ps final p \<noteq> Some u"
by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right assms(1) assms(3) assms(4) assms(5) final_is_s_t_len_t order_refl snapshot_state snapshot_state_unchanged_trace_2)
then show False using assms by simp
qed
lemma ps_after_all_prerecording_events:
assumes
"trace init t final" and
"\<forall>i'. i' \<ge> i \<longrightarrow> ~ prerecording_event t i'" and
"\<forall>j'. j' < i \<longrightarrow> ~ postrecording_event t j'"
shows
"ps_equal_to_snapshot (S t i) final"
proof (unfold ps_equal_to_snapshot_def, rule allI)
fix p
show "Some (states (S t i) p) = ps final p"
proof (rule ccontr)
obtain s where "ps final p = Some s \<or> ps final p = None" by auto
moreover assume "Some (states (S t i) p) \<noteq> ps final p"
ultimately have "ps final p = None \<or> states (S t i) p \<noteq> s" by auto
then show False
proof (elim disjE)
assume "ps final p = None"
then show False
using assms all_processes_snapshotted_in_final_state by blast
next
assume st: "states (S t i) p \<noteq> s"
then obtain j where "~ has_snapshotted (S t j) p \<and> has_snapshotted (S t (j+1)) p"
using Suc_eq_plus1 assms(1) exists_snapshot_for_all_p by presburger
then show False
proof (cases "has_snapshotted (S t i) p")
case False
then have "j \<ge> i"
by (metis Suc_eq_plus1 \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> assms(1) not_less_eq_eq snapshot_stable_ver_3)
let ?t = "take (j-i) (drop i t)"
have "\<exists>ev. ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p"
proof (rule ccontr)
assume "~ (\<exists>ev. ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p)"
moreover have "trace (S t i) ?t (S t j)"
using \<open>i \<le> j\<close> assms(1) exists_trace_for_any_i_j by blast
ultimately have "states (S t j) p = states (S t i) p"
using no_state_change_if_only_nonregular_events st by blast
then show False
by (metis \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> \<open>ps final p = Some s \<or> ps final p = None\<close> assms(1) final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation.snapshot_stable_ver_3 computation_axioms linorder_not_le snapshot_state_must_have_been_reached st)
qed
then obtain ev where "ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p"
by blast
then obtain k where t_ind: "0 \<le> k \<and> k < length ?t \<and> ev = ?t ! k"
by (metis in_set_conv_nth not_le not_less_zero)
moreover have "length ?t \<le> j - i" by simp
ultimately have "?t ! k = (drop i t) ! k"
using less_le_trans nth_take by blast
also have "... = t ! (k+i)"
by (metis \<open>ev \<in> set (take (j - i) (drop i t)) \<and> regular_event ev \<and> occurs_on ev = p\<close> add.commute drop_eq_Nil length_greater_0_conv length_pos_if_in_set nat_le_linear nth_drop take_eq_Nil)
finally have "?t ! k = t ! (k+i)" by simp
have "prerecording_event t (k+i)"
proof -
have "regular_event (?t ! k)"
using \<open>ev \<in> set (take (j - i) (drop i t)) \<and> regular_event ev \<and> occurs_on ev = p\<close> t_ind by blast
moreover have "occurs_on (?t ! k) = p"
using \<open>ev \<in> set (take (j - i) (drop i t)) \<and> regular_event ev \<and> occurs_on ev = p\<close> t_ind by blast
moreover have "~ has_snapshotted (S t (k+i)) p"
proof -
have "k+i \<le> j"
using \<open>length (take (j - i) (drop i t)) \<le> j - i\<close> t_ind by linarith
show ?thesis
using \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> \<open>k+i \<le> j\<close> assms(1) snapshot_stable_ver_3 by blast
qed
ultimately show ?thesis
using \<open>take (j - i) (drop i t) ! k = t ! (k + i)\<close> prerecording_event t_ind by auto
qed
then show False using assms by auto
next
case True
have "j < i"
proof (rule ccontr)
assume "~ j < i"
then have "j \<ge> i" by simp
moreover have "~ has_snapshotted (S t j) p"
using \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> by blast
moreover have "trace (S t i) (take (j - i) (drop i t)) (S t j)"
using assms(1) calculation(1) exists_trace_for_any_i_j by blast
ultimately have "~ has_snapshotted (S t i) p"
using snapshot_stable by blast
then show False using True by simp
qed
let ?t = "take (i-j) (drop j t)"
have "\<exists>ev. ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p"
proof (rule ccontr)
assume "~ (\<exists>ev. ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p)"
moreover have "trace (S t j) ?t (S t i)"
using \<open>j < i\<close> assms(1) exists_trace_for_any_i_j less_imp_le by blast
ultimately have "states (S t j) p = states (S t i) p"
using no_state_change_if_only_nonregular_events by auto
moreover have "states (S t j) p = s"
by (metis \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> \<open>ps final p = Some s \<or> ps final p = None\<close> assms(1) final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation.snapshot_stable_ver_3 computation_axioms linorder_not_le snapshot_state_must_have_been_reached)
- ultimately show False using `states (S t i) p \<noteq> s` by simp
+ ultimately show False using \<open>states (S t i) p \<noteq> s\<close> by simp
qed
then obtain ev where ev: "ev \<in> set ?t \<and> regular_event ev \<and> occurs_on ev = p" by blast
then obtain k where t_ind: "0 \<le> k \<and> k < length ?t \<and> ev = ?t ! k"
by (metis in_set_conv_nth le0)
have "length ?t \<le> i - j" by simp
have "?t ! k = (drop j t) ! k"
using t_ind by auto
also have "... = t ! (k + j)"
by (metis \<open>ev \<in> set (take (i - j) (drop j t)) \<and> regular_event ev \<and> occurs_on ev = p\<close> add.commute drop_eq_Nil length_greater_0_conv length_pos_if_in_set nat_le_linear nth_drop take_eq_Nil)
finally have "?t ! k = t ! (k+j)" by simp
have "postrecording_event t (k+j)"
proof -
have "trace (S t j) (take k (drop j t)) (S t (k+j))"
by (metis add_diff_cancel_right' assms(1) exists_trace_for_any_i_j le_add_same_cancel2 t_ind)
then have "has_snapshotted (S t (k+j)) p"
by (metis Suc_eq_plus1 Suc_leI \<open>\<not> ps (S t j) p \<noteq> None \<and> ps (S t (j + 1)) p \<noteq> None\<close> \<open>take (i - j) (drop j t) ! k = t ! (k + j)\<close> assms(1) drop_eq_Nil ev computation.snapshot_stable_ver_3 computation_axioms le_add_same_cancel2 length_greater_0_conv length_pos_if_in_set linorder_not_le order_le_less regular_event_preserves_process_snapshots step_Suc t_ind take_eq_Nil)
then show ?thesis
using \<open>take (i - j) (drop j t) ! k = t ! (k + j)\<close> ev postrecording_event t_ind by auto
qed
moreover have "k + j < i"
using \<open>length (take (i - j) (drop j t)) \<le> i - j\<close> t_ind by linarith
ultimately show False using assms(3) by simp
qed
qed
qed
qed
subsubsection \<open>Relating channel states\<close>
lemma cs_when_recording:
shows
"\<lbrakk> i < j; j \<le> length t; trace init t final;
has_snapshotted (S t i) p;
snd (cs (S t i) cid) = Recording;
snd (cs (S t j) cid) = Done;
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> map Msg (fst (cs (S t j) cid))
= map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
proof (induct "j - (i+1)" arbitrary: i)
case 0
then have "j = i+1" by simp
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t j)" using "0.prems" step_Suc by simp
then have rm: "\<exists>q p. t ! i = RecvMarker cid q p" using done_only_from_recv_marker "0.prems" by force
then have RecvMarker: "t ! i = RecvMarker cid q p"
by (metis "0.prems"(7) RecvMarker_given_channel event.collapse(5) event.disc(25) event.inject(5) happen_implies_can_occur local.step)
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = []"
proof -
have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by simp
then show ?thesis
proof -
have "\<And>p ms. takeWhile p ms = [] \<or> p (hd ms::'c message)"
by (metis (no_types) hd_append2 hd_in_set set_takeWhileD takeWhile_dropWhile_id)
then show ?thesis
using \<open>can_occur (t ! i) (S t i)\<close> can_occur_def rm by fastforce
qed
qed
then show ?case
using local.step rm by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis Suc_eq_plus1 less_SucI nat_induct_at_least step_Suc)
have ib: "i+1 < j \<and> j \<le> length t \<and> has_snapshotted (S t (i+1)) p \<and> snd (cs (S t j) cid) = Done"
using Suc.hyps(2) Suc.prems(2) Suc.prems(4) Suc.prems(6) local.step snapshot_state_unchanged by auto
have snap_q: "has_snapshotted (S t i) q"
using Suc(7) Suc.prems(3) Suc cs_recording_implies_snapshot by blast
then show ?case
proof (cases "t ! i")
case (Snapshot r)
then have "r \<noteq> p"
using Suc.prems(4) can_occur_def local.step by auto
then have "msgs (S t (i+1)) cid = msgs (S t i) cid"
using Snapshot local.step Suc.prems(7) by auto
moreover have "cs (S t (i+1)) cid = cs (S t i) cid"
proof -
have "r \<noteq> q" using Snapshot can_occur_def snap_q step by auto
then show ?thesis using Snapshot local.step Suc.prems(7) by auto
qed
ultimately show ?thesis using Suc ib by force
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = []"
proof -
have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by simp
then show ?thesis
proof -
have "\<And>p ms. takeWhile p ms = [] \<or> p (hd ms::'c message)"
by (metis (no_types) hd_append2 hd_in_set set_takeWhileD takeWhile_dropWhile_id)
then show ?thesis
using RecvMarker True \<open>can_occur (t ! i) (S t i)\<close> can_occur_def by fastforce
qed
qed
moreover have "snd (cs (S t (i+1)) cid) = Done"
using RecvMarker Suc.prems(1) Suc.prems(2) Suc.prems(3) True recv_marker_means_cs_Done by auto
moreover have "fst (cs (S t i) cid) = fst (cs (S t (i+1)) cid)"
using RecvMarker True local.step by auto
ultimately show ?thesis
by (metis Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(7) Suc_eq_plus1 Suc_leI append_Nil2 cs_done_implies_same_snapshots)
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step Suc False by auto
next
case False
- with RecvMarker step Suc `cid \<noteq> cid'` show ?thesis by (cases "s = p", auto)
+ with RecvMarker step Suc \<open>cid \<noteq> cid'\<close> show ?thesis by (cases "s = p", auto)
qed
moreover have "cs (S t i) cid = cs (S t (i+1)) cid"
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step Suc False by auto
next
case no_snap: False
then show ?thesis
proof (cases "r = q")
case True
- then show ?thesis using snap_q no_snap `r = q` by simp
+ then show ?thesis using snap_q no_snap \<open>r = q\<close> by simp
next
case False
- then show ?thesis using RecvMarker step Suc no_snap False `cid \<noteq> cid'` by simp
+ then show ?thesis using RecvMarker step Suc no_snap False \<open>cid \<noteq> cid'\<close> by simp
qed
qed
ultimately show ?thesis using Suc ib by simp
qed
next
case (Trans r u u')
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Trans by auto
ultimately show ?thesis using Suc ib by simp
next
case (Send cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
have marker_in_msgs: "Marker \<in> set (msgs (S t i) cid)"
proof -
have "has_snapshotted (S t i) p" using Suc by simp
moreover have "i < length t"
using Suc.prems(1) Suc.prems(2) less_le_trans by blast
moreover have "snd (cs (S t i) cid) \<noteq> Done" using Suc by simp
ultimately show ?thesis using snapshotted_and_not_done_implies_marker_in_channel less_imp_le using Suc by blast
qed
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
proof -
have "butlast (msgs (S t (i+1)) cid) = msgs (S t i) cid" using step True Send by auto
then show ?thesis
proof -
have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid @ [last (msgs (S t (i + 1)) cid)]) = takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
using marker_in_msgs by force
then show ?thesis
by (metis (no_types) \<open>butlast (msgs (S t (i + 1)) cid) = msgs (S t i) cid\<close> append_butlast_last_id in_set_butlastD length_greater_0_conv length_pos_if_in_set marker_in_msgs)
qed
qed
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto
ultimately show ?thesis using ib Suc by simp
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send by auto
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto
ultimately show ?thesis using Suc ib by simp
qed
next
case (Recv cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have msgs_ip1: "Msg m # msgs (S t (i+1)) cid = msgs (S t i) cid"
using Suc Recv step by auto
moreover have cs_ip1: "cs (S t (i+1)) cid = (fst (cs (S t i) cid) @ [m], Recording)"
using True Suc Recv step by auto
ultimately show ?thesis
proof -
have "map Msg (fst (cs (S t j) cid))
= map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
using Suc ib cs_ip1 by force
moreover have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)
= map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
proof -
have "takeWhile ((\<noteq>) Marker) (Msg m # msgs (S t (i+1)) cid) = Msg m # takeWhile ((\<noteq>) Marker) (msgs (S t (i + 1)) cid)"
by auto
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = Msg m # takeWhile ((\<noteq>) Marker) (msgs (S t (i + 1)) cid)"
by (metis msgs_ip1)
then show ?thesis
using cs_ip1 by auto
qed
ultimately show ?thesis by simp
qed
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv by auto
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Recv False by auto
ultimately show ?thesis using Suc ib by simp
qed
qed
qed
lemma cs_when_recording_2:
shows
"\<lbrakk> i \<le> j; trace init t final;
~ has_snapshotted (S t i) p;
\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = p;
snd (cs (S t i) cid) = Recording;
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)
\<and> snd (cs (S t j) cid) = Recording"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis Suc_eq_plus1 all_processes_snapshotted_in_final_state distributed_system.step_Suc distributed_system_axioms computation.final_is_s_t_len_t computation_axioms linorder_not_le snapshot_stable_ver_3)
have ib: "i+1 \<le> j \<and> ~ has_snapshotted (S t (i+1)) p
\<and> (\<forall>k. (i+1) \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = p) \<and> j - (i+1) = n"
by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(4) diff_Suc_1 diff_diff_left Suc_eq_plus1 Suc_leD Suc_le_eq Suc_neq_Zero cancel_comm_monoid_add_class.diff_cancel le_neq_implies_less le_refl local.step no_state_change_if_no_event)
have snap_q: "has_snapshotted (S t i) q"
using Suc.prems(5,6) Suc.prems(2) cs_recording_implies_snapshot by blast
then show ?case
proof (cases "t ! i")
case (Snapshot r)
then have "r \<noteq> p" using Suc by auto
then have "msgs (S t (i+1)) cid = msgs (S t i) cid"
using Snapshot local.step Suc.prems(6) by auto
moreover have "cs (S t (i+1)) cid = cs (S t i) cid"
proof -
have "r \<noteq> q" using step can_occur_def Snapshot snap_q by auto
then show ?thesis using Snapshot step Suc by simp
qed
ultimately show ?thesis using Suc ib by auto
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "Marker \<in> set (msgs (S t i) cid)"
using RecvMarker RecvMarker_implies_Marker_in_set local.step by blast
then have "has_snapshotted (S t i) p"
using Suc.prems(2) no_marker_if_no_snapshot Suc by blast
then show ?thesis using Suc.prems by simp
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step Suc False by auto
next
case False
- with RecvMarker step Suc `cid \<noteq> cid'` show ?thesis by (cases "s = p", auto)
+ with RecvMarker step Suc \<open>cid \<noteq> cid'\<close> show ?thesis by (cases "s = p", auto)
qed
moreover have "cs (S t i) cid = cs (S t (i+1)) cid"
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step Suc False by auto
next
case no_snap: False
then show ?thesis
proof (cases "r = q")
case True
- then show ?thesis using snap_q no_snap `r = q` by simp
+ then show ?thesis using snap_q no_snap \<open>r = q\<close> by simp
next
case False
- then show ?thesis using RecvMarker step Suc no_snap False `cid \<noteq> cid'` by simp
+ then show ?thesis using RecvMarker step Suc no_snap False \<open>cid \<noteq> cid'\<close> by simp
qed
qed
ultimately show ?thesis using Suc ib by auto
qed
next
case (Trans r u u')
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Trans by auto
ultimately show ?thesis using Suc ib by auto
next
case (Send cid' r s u u' m)
then have "r \<noteq> p"
using Suc.hyps(2) Suc.prems(4) Suc by auto
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "channel cid = channel cid'" by auto
then have "(p, q) = (r, s)" using can_occur_def step Send Suc by auto
- then show False using `r \<noteq> p` by simp
+ then show False using \<open>r \<noteq> p\<close> by simp
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Send by simp
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Send by auto
ultimately show ?thesis using Suc ib by auto
next
case (Recv cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have msgs_ip1: "Msg m # msgs (S t (i+1)) cid = msgs (S t i) cid"
using Suc Recv step by auto
moreover have cs_ip1: "cs (S t (i+1)) cid = (fst (cs (S t i) cid) @ [m], Recording)"
using True Suc Recv step by auto
ultimately show ?thesis
proof -
have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)
\<and> snd (cs (S t j) cid) = Recording"
using Suc ib cs_ip1 by auto
moreover have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)
= map Msg (fst (cs (S t (i+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
proof -
have "takeWhile ((\<noteq>) Marker) (Msg m # msgs (S t (i + 1)) cid) = Msg m # takeWhile ((\<noteq>) Marker) (msgs (S t (i + 1)) cid)"
by fastforce
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = Msg m # takeWhile ((\<noteq>) Marker) (msgs (S t (i + 1)) cid)"
by (metis msgs_ip1)
then show ?thesis
using cs_ip1 by force
qed
ultimately show ?thesis using cs_ip1 by simp
qed
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv by auto
moreover have "cs (S t i) cid = cs (S t (i+1)) cid" using step Recv False by auto
ultimately show ?thesis using Suc ib by auto
qed
qed
qed
lemma cs_when_recording_3:
shows
"\<lbrakk> i \<le> j; trace init t final;
~ has_snapshotted (S t i) q;
\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = q;
snd (cs (S t i) cid) = NotStarted;
has_snapshotted (S t i) p;
Marker : set (msgs (S t i) cid);
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)
\<and> snd (cs (S t j) cid) = NotStarted"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis Suc_eq_plus1 all_processes_snapshotted_in_final_state distributed_system.step_Suc distributed_system_axioms computation.final_is_s_t_len_t computation_axioms linorder_not_le snapshot_stable_ver_3)
have ib: "i+1 \<le> j \<and> ~ has_snapshotted (S t (i+1)) q \<and> has_snapshotted (S t (i+1)) p
\<and> (\<forall>k. (i+1) \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = q) \<and> j - (i+1) = n
\<and> Marker : set (msgs (S t (i+1)) cid) \<and> cs (S t i) cid = cs (S t (i+1)) cid"
proof -
have "i+1 \<le> j \<and> ~ has_snapshotted (S t (i+1)) q
\<and> (\<forall>k. (i+1) \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = q) \<and> j - (i+1) = n"
by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(4) diff_Suc_1 diff_diff_left Suc_eq_plus1 Suc_leD Suc_le_eq Suc_neq_Zero cancel_comm_monoid_add_class.diff_cancel le_neq_implies_less le_refl local.step no_state_change_if_no_event)
moreover have "has_snapshotted (S t (i+1)) p"
using Suc.prems(6) local.step snapshot_state_unchanged by auto
moreover have "Marker : set (msgs (S t (i+1)) cid)"
using Suc calculation(1) local.step recv_marker_means_snapshotted_2 by blast
moreover have "cs (S t i) cid = cs (S t (i+1)) cid"
using Suc calculation(1) no_recording_cs_if_not_snapshotted by auto
ultimately show ?thesis by simp
qed
then show ?case
proof (cases "t ! i")
case (Snapshot r)
then have "r \<noteq> q" using Suc by auto
then have "takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid) = takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
proof (cases "occurs_on (t ! i) = p")
case True
then show ?thesis
by (metis (mono_tags, lifting) Snapshot Suc.prems(6) distributed_system.can_occur_def event.sel(4) event.simps(29) computation_axioms computation_def happen_implies_can_occur local.step)
next
case False
then have "msgs (S t (i+1)) cid = msgs (S t i) cid"
using Snapshot local.step Suc by auto
then show ?thesis by simp
qed
then show ?thesis using Suc ib by metis
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "snd (cs (S t i) cid) = Done"
by (metis RecvMarker Suc.prems(2) Suc_eq_plus1 Suc_le_eq exactly_one_snapshot computation.no_change_if_ge_length_t computation.recv_marker_means_cs_Done computation.snapshot_stable_ver_2 computation_axioms ib nat_le_linear)
then show ?thesis using Suc.prems by simp
next
case False
then have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
proof (cases "has_snapshotted (S t i) r")
case True
with RecvMarker False step show ?thesis by auto
next
case no_snap: False
then have "r \<noteq> p"
using Suc.prems(6) by auto
then show ?thesis using no_snap RecvMarker step Suc.prems False by auto
qed
then show ?thesis using Suc ib by metis
qed
next
case (Trans r u u')
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step by auto
then show ?thesis using Suc ib by auto
next
case (Send cid' r s u u' m)
then have "r \<noteq> q"
using Suc.hyps(2) Suc.prems(4) by auto
have marker: "Marker \<in> set (msgs (S t i) cid)" using Suc by simp
with step Send marker have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = takeWhile ((\<noteq>) Marker) (msgs (S t (i+1)) cid)"
by (cases "cid = cid'", auto)
then show ?thesis using Suc ib by auto
next
case (Recv cid' r s u u' m)
then have "cid' \<noteq> cid"
by (metis Suc.hyps(2) Suc.prems(4) Suc.prems(8) distributed_system.can_occur_Recv distributed_system_axioms event.sel(3) happen_implies_can_occur local.step option.inject order_refl prod.inject zero_less_Suc zero_less_diff)
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using step Recv Suc by simp
then show ?thesis using Suc ib by auto
qed
qed
lemma at_most_one_marker:
shows
"\<lbrakk> trace init t final; channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> Marker \<notin> set (msgs (S t i) cid)
\<or> (\<exists>!j. j < length (msgs (S t i) cid) \<and> msgs (S t i) cid ! j = Marker)"
proof (induct i)
case 0
then show ?case using no_initial_Marker init_is_s_t_0 by auto
next
case (Suc i)
then show ?case
proof (cases "i < length t")
case False
then show ?thesis
by (metis Suc.prems(1) final_is_s_t_len_t computation.no_change_if_ge_length_t computation_axioms le_refl less_imp_le_nat no_marker_left_in_final_state not_less_eq)
next
case True
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (Suc i))" using step_Suc Suc.prems by blast
moreover have "Marker \<notin> set (msgs (S t i) cid)
\<or> (\<exists>!j. j < length (msgs (S t i) cid) \<and> msgs (S t i) cid ! j = Marker)"
using Suc.hyps Suc.prems(1) Suc.prems(2) by linarith
moreover have "Marker \<notin> set (msgs (S t (Suc i)) cid)
\<or> (\<exists>!j. j < length (msgs (S t (Suc i)) cid) \<and> msgs (S t (Suc i)) cid ! j = Marker)"
proof (cases "Marker \<notin> set (msgs (S t i) cid)")
case no_marker: True
then show ?thesis
proof (cases "t ! i")
case (Snapshot r)
then show ?thesis
proof (cases "r = p")
case True
then have new_msgs: "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Marker]"
using step Snapshot Suc by auto
then show ?thesis using util_exactly_one_element no_marker by fastforce
next
case False
then show ?thesis
using Snapshot local.step no_marker Suc by auto
qed
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then show ?thesis
using RecvMarker RecvMarker_implies_Marker_in_set local.step no_marker by blast
next
case False
then show ?thesis
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step Suc False by simp
next
case no_snap: False
then show ?thesis
proof (cases "r = p")
case True
- then have "msgs (S t (i+1)) cid = msgs (S t i) cid @ [Marker]" using RecvMarker step Suc.prems no_snap `cid \<noteq> cid'` by simp
+ then have "msgs (S t (i+1)) cid = msgs (S t i) cid @ [Marker]" using RecvMarker step Suc.prems no_snap \<open>cid \<noteq> cid'\<close> by simp
then show ?thesis
proof -
assume a1: "msgs (S t (i + 1)) cid = msgs (S t i) cid @ [Marker]"
{ fix nn :: "nat \<Rightarrow> nat"
have "\<forall>ms m. \<exists>n. \<forall>na. ((m::'c message) \<in> set ms \<or> n < length (ms @ [m])) \<and> (m \<in> set ms \<or> (ms @ [m]) ! n = m) \<and> (\<not> na < length (ms @ [m]) \<or> (ms @ [m]) ! na \<noteq> m \<or> m \<in> set ms \<or> na = n)"
by (metis (no_types) util_exactly_one_element)
then have "\<exists>n. n < length (msgs (S t (Suc i)) cid) \<and> nn n = n \<and> msgs (S t (Suc i)) cid ! n = Marker \<or> n < length (msgs (S t (Suc i)) cid) \<and> msgs (S t (Suc i)) cid ! n = Marker \<and> \<not> nn n < length (msgs (S t (Suc i)) cid) \<or> n < length (msgs (S t (Suc i)) cid) \<and> msgs (S t (Suc i)) cid ! n = Marker \<and> msgs (S t (Suc i)) cid ! nn n \<noteq> Marker"
using a1 by (metis Suc_eq_plus1 no_marker) }
then show ?thesis
by (metis (no_types))
qed
next
case False
- then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using RecvMarker step Suc.prems `cid \<noteq> cid'` no_snap by simp
+ then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using RecvMarker step Suc.prems \<open>cid \<noteq> cid'\<close> no_snap by simp
then show ?thesis using Suc by simp
qed
qed
qed
next
case (Trans r u u')
then show ?thesis using no_marker step by auto
next
case (Send cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "Marker \<notin> set (msgs (S t (Suc i)) cid)" using step no_marker Send by auto
then show ?thesis by simp
next
case False
then have "Marker \<notin> set (msgs (S t (Suc i)) cid)" using step no_marker Send by auto
then show ?thesis by simp
qed
next
case (Recv cid' r s u u' m)
with step no_marker Recv show ?thesis by (cases "cid = cid'", auto)
qed
next
case False
then have asm: "\<exists>!j. j < length (msgs (S t i) cid) \<and> msgs (S t i) cid ! j = Marker"
using Suc by simp
have len_filter: "length (filter ((=) Marker) (msgs (S t i) cid)) = 1"
by (metis False \<open>Marker \<notin> set (msgs (S t i) cid) \<or> (\<exists>!j. j < length (msgs (S t i) cid) \<and> msgs (S t i) cid ! j = Marker)\<close> exists_one_iff_filter_one)
have snap_p: "has_snapshotted (S t i) p"
using False Suc.prems no_marker_if_no_snapshot by blast
show ?thesis
proof (cases "t ! i")
case (Snapshot r)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
moreover have "can_occur (t ! i) (S t i)" using happen_implies_can_occur step by blast
ultimately show False using snap_p can_occur_def Snapshot by auto
qed
then have "msgs (S t (Suc i)) cid = msgs (S t i) cid" using step Snapshot Suc by auto
then show ?thesis using asm by simp
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid"
using RecvMarker step by auto
then have "Marker \<notin> set (msgs (S t (Suc i)) cid)"
proof -
have "\<forall>j. j \<noteq> 0 \<and> j < length (msgs (S t i) cid) \<longrightarrow> msgs (S t i) cid ! j \<noteq> Marker"
by (metis False \<open>Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\<close> asm length_pos_if_in_set nth_Cons_0)
then show ?thesis
proof -
assume a1: "\<forall>j. j \<noteq> 0 \<and> j < length (msgs (S t i) cid) \<longrightarrow> msgs (S t i) cid ! j \<noteq> Marker"
have "\<And>ms n. ms \<noteq> msgs (S t i) cid \<or> length (msgs (S t (Suc i)) cid) \<noteq> n \<or> length ms = Suc n"
by (metis \<open>Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\<close> length_Suc_conv)
then show ?thesis
using a1 by (metis (no_types) Suc_mono Zero_not_Suc \<open>Marker # msgs (S t (Suc i)) cid = msgs (S t i) cid\<close> in_set_conv_nth nth_Cons_Suc)
qed
qed
then show ?thesis by simp
next
case cid_neq_cid': False
then show ?thesis
proof (cases "has_snapshotted (S t i) r")
case True
then have "msgs (S t (Suc i)) cid = msgs (S t i) cid"
using cid_neq_cid' RecvMarker local.step snap_p by auto
then show ?thesis using asm by simp
next
case False
then have "r \<noteq> p"
using snap_p by blast
then have "msgs (S t (Suc i)) cid = msgs (S t i) cid" using cid_neq_cid' RecvMarker step False Suc by auto
then show ?thesis using asm by simp
qed
qed
next
case (Trans r u u')
then show ?thesis using step asm by auto
next
case (Send cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have new_messages: "msgs (S t (Suc i)) cid = msgs (S t i) cid @ [Msg m]"
using step Send by auto
then have "\<exists>!j. j < length (msgs (S t (Suc i)) cid) \<and> msgs (S t (Suc i)) cid ! j = Marker"
proof -
have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid))
= length (filter ((=) Marker) (msgs (S t i) cid))
+ length (filter ((=) Marker) [Msg m])"
by (simp add: new_messages)
then have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid)) = 1"
using len_filter by simp
then show ?thesis using exists_one_iff_filter_one by metis
qed
then show ?thesis by simp
next
case False
then show ?thesis using step Send asm by auto
qed
next
case (Recv cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have new_msgs: "Msg m # msgs (S t (Suc i)) cid = msgs (S t i) cid" using step Recv by auto
then show ?thesis
proof -
have "length (filter ((=) Marker) (msgs (S t i) cid))
= length (filter ((=) Marker) [Msg m])
+ length (filter ((=) Marker) (msgs (S t (Suc i)) cid))"
by (metis append_Cons append_Nil filter_append len_filter length_append new_msgs)
then have "length (filter ((=) Marker) (msgs (S t (Suc i)) cid)) = 1"
using len_filter by simp
then show ?thesis using exists_one_iff_filter_one by metis
qed
next
case False
then show ?thesis using step Recv asm by auto
qed
qed
qed
then show ?thesis by simp
qed
qed
lemma last_changes_implies_send_when_msgs_nonempty:
assumes
"trace init t final" and
"msgs (S t i) cid \<noteq> []" and
"msgs (S t (i+1)) cid \<noteq> []" and
"last (msgs (S t i) cid) = Marker" and
"last (msgs (S t (i+1)) cid) \<noteq> Marker" and
"channel cid = Some (p, q)"
shows
"(\<exists>u u' m. t ! i = Send cid p q u u' m)"
proof -
have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis Suc_eq_plus1_left add.commute assms(1) assms(4) assms(5) le_Suc_eq nat_le_linear nat_less_le no_change_if_ge_length_t step_Suc)
then show ?thesis
proof (cases "t ! i")
case (Snapshot r)
then show ?thesis
by (metis assms(4) assms(5) last_snoc local.step next_snapshot)
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)"
proof -
have "Marker # msgs (S t (i + 1)) cid = msgs (S t i) cid"
using RecvMarker local.step True by auto
then show ?thesis
by (metis assms(3) last_ConsR)
qed
then show ?thesis using assms by simp
next
case no_snap: False
then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)"
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis using RecvMarker step no_snap by simp
next
case False
- with RecvMarker step no_snap `cid \<noteq> cid'` assms show ?thesis by (cases "p = r", auto)
+ with RecvMarker step no_snap \<open>cid \<noteq> cid'\<close> assms show ?thesis by (cases "p = r", auto)
qed
then show ?thesis using assms by simp
qed
next
case (Trans r u u')
then show ?thesis
using assms(4) assms(5) local.step by auto
next
case (Send cid' r s u u' m)
then have "cid = cid'"
by (metis (no_types, hide_lams) assms(4) assms(5) local.step next_send)
moreover have "(p, q) = (r, s)"
proof -
- have "channel cid = channel cid'" using `cid = cid'` by simp
+ have "channel cid = channel cid'" using \<open>cid = cid'\<close> by simp
moreover have "channel cid = Some (p, q)" using assms by simp
moreover have "channel cid' = Some (r, s)" using Send step can_occur_def by auto
ultimately show ?thesis by simp
qed
ultimately show ?thesis using Send by auto
next
case (Recv cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "last (msgs (S t i) cid) = last (msgs (S t (i+1)) cid)"
by (metis (no_types, lifting) Recv assms(3) assms(4) last_ConsR local.step next_recv)
then show ?thesis using assms by simp
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto
then show ?thesis using assms by simp
qed
qed
qed
lemma no_marker_after_RecvMarker:
assumes
"trace init t final" and
"(S t i) \<turnstile> RecvMarker cid p q \<mapsto> (S t (i+1))" and
"channel cid = Some (q, p)"
shows
"Marker \<notin> set (msgs (S t (i+1)) cid)"
proof -
have new_msgs: "msgs (S t i) cid = Marker # msgs (S t (i+1)) cid"
using assms(2) by auto
have one_marker: "\<exists>!j. j < length (msgs (S t i) cid) \<and> msgs (S t i) cid ! j = Marker"
by (metis assms(1,3) at_most_one_marker list.set_intros(1) new_msgs)
then obtain j where "j < length (msgs (S t i) cid)" "msgs (S t i) cid ! j = Marker" by blast
then have "j = 0" using one_marker new_msgs by auto
then have "\<forall>j. j \<noteq> 0 \<and> j < length (msgs (S t i) cid) \<longrightarrow> msgs (S t i) cid ! j \<noteq> Marker"
using one_marker
using \<open>j < length (msgs (S t i) cid)\<close> \<open>msgs (S t i) cid ! j = Marker\<close> by blast
then have "\<forall>j. j < length (msgs (S t (i+1)) cid) \<longrightarrow> msgs (S t (i+1)) cid ! j \<noteq> Marker"
by (metis One_nat_def Suc_eq_plus1 Suc_le_eq Suc_mono le_zero_eq list.size(4) new_msgs not_less0 nth_Cons_Suc)
then show ?thesis
by (simp add: in_set_conv_nth)
qed
lemma no_marker_and_snapshotted_implies_no_more_markers_trace:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
has_snapshotted (S t i) p;
Marker \<notin> set (msgs (S t i) cid);
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> Marker \<notin> set (msgs (S t j) cid)"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis (no_types, hide_lams) Suc_eq_plus1 cancel_comm_monoid_add_class.diff_cancel distributed_system.step_Suc distributed_system_axioms less_le_trans linorder_not_less old.nat.distinct(2) order_eq_iff)
then have "Marker \<notin> set (msgs (S t (i+1)) cid)"
using no_marker_and_snapshotted_implies_no_more_markers Suc step by blast
moreover have "has_snapshotted (S t (i+1)) p"
using Suc.prems(4) local.step snapshot_state_unchanged by auto
ultimately show ?case
proof -
assume a1: "ps (S t (i + 1)) p \<noteq> None"
assume a2: "Marker \<notin> set (msgs (S t (i + 1)) cid)"
have f3: "j \<noteq> i"
using Suc.hyps(2) by force
have "j - Suc i = n"
by (metis (no_types) Suc.hyps(2) Suc.prems(2) add.commute add_Suc_right add_diff_cancel_left' le_add_diff_inverse)
then show ?thesis
using f3 a2 a1 by (metis Suc.hyps(1) Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(6) Suc_eq_plus1_left add.commute less_Suc_eq linorder_not_less)
qed
qed
lemma marker_not_vanishing_means_always_present:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
Marker : set (msgs (S t i) cid);
Marker : set (msgs (S t j) cid);
channel cid = Some (p, q)
\<rbrakk> \<Longrightarrow> \<forall>k. i \<le> k \<and> k \<le> j \<longrightarrow> Marker : set (msgs (S t k) cid)"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis (no_types, lifting) Suc_eq_plus1 add_lessD1 distributed_system.step_Suc distributed_system_axioms le_add_diff_inverse order_le_less zero_less_Suc zero_less_diff)
have "Marker : set (msgs (S t (i+1)) cid)"
proof (rule ccontr)
assume asm: "~ Marker : set (msgs (S t (i+1)) cid)"
have snap_p: "has_snapshotted (S t i) p"
using Suc.prems(1) Suc.prems(4,6) no_marker_if_no_snapshot by blast
then have "has_snapshotted (S t (i+1)) p"
using local.step snapshot_state_unchanged by auto
then have "Marker \<notin> set (msgs (S t j) cid)"
by (metis Suc.hyps(2) Suc.prems(1) Suc.prems(3) Suc.prems(6) asm discrete no_marker_and_snapshotted_implies_no_more_markers_trace zero_less_Suc zero_less_diff)
then show False using Suc.prems by simp
qed
then show ?case
by (meson Suc.prems(1) Suc.prems(3) Suc.prems(4) Suc.prems(5) Suc.prems(6) computation.snapshot_stable_ver_3 computation_axioms no_marker_and_snapshotted_implies_no_more_markers_trace no_marker_if_no_snapshot)
qed
lemma last_stays_if_no_recv_marker_and_no_send:
shows "\<lbrakk> trace init t final; i < j; j \<le> length t;
last (msgs (S t i) cid) = Marker;
Marker : set (msgs (S t i) cid);
Marker : set (msgs (S t j) cid);
\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ (\<exists>u u' m. t ! k = Send cid p q u u' m);
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> last (msgs (S t j) cid) = Marker"
proof (induct "j - (i+1)" arbitrary: i)
case 0
then have "j = i+1" by simp
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis "0"(2) "0.prems"(2) "0.prems"(3) Suc_eq_plus1 distributed_system.step_Suc distributed_system_axioms less_le_trans)
have "Marker = last (msgs (S t (i+1)) cid)"
proof (rule ccontr)
assume "~ Marker = last (msgs (S t (i+1)) cid)"
then have "\<exists>u u' m. t ! i = Send cid p q u u' m"
proof -
- have "msgs (S t (i+1)) cid \<noteq> []" using "0" `j = i+1` by auto
+ have "msgs (S t (i+1)) cid \<noteq> []" using "0" \<open>j = i+1\<close> by auto
moreover have "msgs (S t i) cid \<noteq> []" using "0" by auto
ultimately show ?thesis
using "0.prems"(1) "0.prems"(4) "0.prems"(8) \<open>Marker \<noteq> last (msgs (S t (i + 1)) cid)\<close> last_changes_implies_send_when_msgs_nonempty by auto
qed
then show False using 0 by auto
qed
- then show ?case using `j = i+1` by simp
+ then show ?case using \<open>j = i+1\<close> by simp
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis (no_types, hide_lams) Suc_eq_plus1 distributed_system.step_Suc distributed_system_axioms less_le_trans)
have marker_present: "Marker : set (msgs (S t (i+1)) cid)"
by (meson Suc.prems(1) Suc.prems(2) Suc.prems(3) Suc.prems(5) Suc.prems(6) Suc.prems(8) discrete le_add1 less_imp_le_nat marker_not_vanishing_means_always_present)
moreover have "Marker = last (msgs (S t (i+1)) cid)"
proof (rule ccontr)
assume asm: "~ Marker = last (msgs (S t (i+1)) cid)"
then have "\<exists>u u' m. t ! i = Send cid p q u u' m"
proof -
have "msgs (S t (i+1)) cid \<noteq> []" using marker_present by auto
moreover have "msgs (S t i) cid \<noteq> []" using Suc by auto
ultimately show ?thesis
using Suc.prems(1) Suc.prems(4) Suc.prems(8) asm last_changes_implies_send_when_msgs_nonempty by auto
qed
then show False using Suc by auto
qed
moreover have "\<forall>k. i+1 \<le> k \<and> k < j \<longrightarrow> ~ (\<exists>u u' m. t ! k = Send cid p q u u' m)"
using Suc.prems by force
moreover have "i+1 < j" using Suc by auto
moreover have "j \<le> length t" using Suc by auto
moreover have "trace init t final" using Suc by auto
moreover have "Marker : set (msgs (S t j) cid)" using Suc by auto
ultimately show ?case using Suc
by (metis diff_Suc_1 diff_diff_left)
qed
lemma last_changes_implies_send_when_msgs_nonempty_trace:
assumes
"trace init t final"
"i < j"
"j \<le> length t"
"Marker : set (msgs (S t i) cid)"
"Marker : set (msgs (S t j) cid)"
"last (msgs (S t i) cid) = Marker"
"last (msgs (S t j) cid) \<noteq> Marker"
"channel cid = Some (p, q)"
shows
"\<exists>k u u' m. i \<le> k \<and> k < j \<and> t ! k = Send cid p q u u' m"
proof (rule ccontr)
assume "~ (\<exists>k u u' m. i \<le> k \<and> k < j \<and> t ! k = Send cid p q u u' m)"
then have "\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ (\<exists>u u' m. t ! k = Send cid p q u u' m)" by blast
then have "last (msgs (S t j) cid) = Marker" using assms last_stays_if_no_recv_marker_and_no_send by blast
then show False using assms by simp
qed
lemma msg_after_marker_and_nonempty_implies_postrecording_event:
assumes
"trace init t final" and
"Marker : set (msgs (S t i) cid)" and
"Marker \<noteq> last (msgs (S t i) cid)" and
"i \<le> length t" and
"channel cid = Some (p, q)"
shows
"\<exists>j. j < i \<and> postrecording_event t j" (is ?P)
proof -
let ?len = "length (msgs (S t i) cid)"
have "?len \<noteq> 0" using assms(2) by auto
have snap_p_i: "has_snapshotted (S t i) p"
using assms no_marker_if_no_snapshot by blast
obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have "j < i"
by (meson assms(1) computation.snapshot_stable_ver_2 computation_axioms not_less snap_p(1) snap_p_i)
have step_snap: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis Suc_eq_plus1 assms(1) l2 nat_le_linear nat_less_le snap_p(1) snapshot_stable_ver_2 step_Suc)
have re: "~ regular_event (t ! j)"
by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_snap)
have op: "occurs_on (t ! j) = p"
using no_state_change_if_no_event snap_p(1) snap_p(2) step_snap by force
have marker_last: "Marker = last (msgs (S t (j+1)) cid) \<and> msgs (S t (j+1)) cid \<noteq> []"
proof -
have "isSnapshot (t ! j) \<or> isRecvMarker (t ! j)" using re nonregular_event by auto
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have "t ! j = Snapshot p"
using op by auto
then show ?thesis using step_snap assms by auto
next
case 2
then obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' p r"
by (metis event.collapse(5) op)
then have "cid \<noteq> cid'"
using RecvMarker_implies_Marker_in_set assms(1) assms(5) no_marker_if_no_snapshot snap_p(1) step_snap by blast
then show ?thesis
using assms snap_p(1) step_snap RecvMarker by auto
qed
qed
then have "\<exists>k u u' m. j+1 \<le> k \<and> k < i \<and> t ! k = Send cid p q u u' m"
proof -
have "j+1 < i"
proof -
have "(S t (j+1)) \<noteq> (S t i)"
using assms(3) marker_last by auto
then have "j+1 \<noteq> i" by auto
- moreover have "j+1 \<le> i" using `j < i` by simp
+ moreover have "j+1 \<le> i" using \<open>j < i\<close> by simp
ultimately show ?thesis by simp
qed
moreover have "trace init t final" using assms by simp
moreover have "Marker = last (msgs (S t (j+1)) cid)" using marker_last by simp
moreover have "Marker : set (msgs (S t (j+1)) cid)" using marker_last by (simp add: marker_last)
ultimately show ?thesis using assms last_changes_implies_send_when_msgs_nonempty_trace by simp
qed
then obtain k where Send: "\<exists>u u' m. j+1 \<le> k \<and> k < i \<and> t ! k = Send cid p q u u' m" by blast
then have "postrecording_event t k"
proof -
have "k < length t" using Send assms by simp
moreover have "isSend (t ! k)" using Send by auto
moreover have "has_snapshotted (S t k) p" using Send snap_p
using assms(1) snapshot_stable_ver_3 by blast
moreover have "occurs_on (t ! k) = p" using Send by auto
ultimately show ?thesis unfolding postrecording_event by simp
qed
then show ?thesis using Send by auto
qed
lemma same_messages_if_no_occurrence_trace:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
(\<forall>k. i \<le> k \<and> k < j \<longrightarrow> occurs_on (t ! k) \<noteq> p \<and> occurs_on (t ! k) \<noteq> q);
channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> msgs (S t i) cid = msgs (S t j) cid \<and> cs (S t i) cid = cs (S t j) cid"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
by (metis (no_types, hide_lams) Suc_eq_plus1 Suc_n_not_le_n diff_self_eq_0 distributed_system.step_Suc distributed_system_axioms le0 le_eq_less_or_eq less_le_trans)
then have "msgs (S t i) cid = msgs (S t (i+1)) cid \<and> cs (S t i) cid = cs (S t (i+1)) cid"
proof -
have "~ occurs_on (t ! i) = p" using Suc by simp
moreover have "~ occurs_on (t ! i) = q" using Suc by simp
ultimately show ?thesis using step Suc same_messages_if_no_occurrence by blast
qed
moreover have "msgs (S t (i+1)) cid = msgs (S t j) cid \<and> cs (S t (i+1)) cid = cs (S t j) cid"
proof -
have "i+1 \<le> j" using Suc by linarith
moreover have "\<forall>k. i+1 \<le> k \<and> k < j \<longrightarrow> occurs_on (t ! k) \<noteq> p \<and> occurs_on (t ! k) \<noteq> q" using Suc by force
ultimately show ?thesis using Suc by auto
qed
ultimately show ?case by simp
qed
lemma snapshot_step_cs_preservation_p:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"~ regular_event ev" and
"occurs_on ev = p" and
"channel cid = Some (p, q)"
shows
"map Msg (fst (cs c cid)) @ takeWhile ((\<noteq>) Marker) (msgs c cid)
= map Msg (fst (cs c' cid)) @ takeWhile ((\<noteq>) Marker) (msgs c' cid)
\<and> snd (cs c cid) = snd (cs c' cid)"
proof -
have "isSnapshot ev \<or> isRecvMarker ev" using assms nonregular_event by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have Snap: "ev = Snapshot p" by (metis event.collapse(4) assms(3))
then have "fst (cs c cid) = fst (cs c' cid)"
using assms(1) assms(2) regular_event same_cs_if_not_recv by blast
moreover have "takeWhile ((\<noteq>) Marker) (msgs c cid)
= takeWhile ((\<noteq>) Marker) (msgs c' cid)"
proof -
have "msgs c' cid = msgs c cid @ [Marker]" using assms Snap by auto
then show ?thesis
by (simp add: takeWhile_tail)
qed
moreover have "snd (cs c cid) = snd (cs c' cid)"
using Snap assms no_self_channel by fastforce
ultimately show ?thesis by simp
next
case 2
then obtain cid' r where RecvMarker: "ev = RecvMarker cid' p r" by (metis event.collapse(5) assms(3))
have "cid \<noteq> cid'"
by (metis "2" RecvMarker assms(1) assms(4) distributed_system.RecvMarker_given_channel distributed_system.happen_implies_can_occur distributed_system_axioms event.sel(5,10) no_self_channel)
then have "fst (cs c cid) = fst (cs c' cid)"
using RecvMarker assms(1) assms(2) regular_event same_cs_if_not_recv by blast
moreover have "takeWhile ((\<noteq>) Marker) (msgs c cid)
= takeWhile ((\<noteq>) Marker) (msgs c' cid)"
proof (cases "has_snapshotted c p")
case True
- then have "msgs c cid = msgs c' cid" using RecvMarker `cid \<noteq> cid'` assms by auto
+ then have "msgs c cid = msgs c' cid" using RecvMarker \<open>cid \<noteq> cid'\<close> assms by auto
then show ?thesis by simp
next
case False
- then have "msgs c' cid = msgs c cid @ [Marker]" using RecvMarker `cid \<noteq> cid'` assms by auto
+ then have "msgs c' cid = msgs c cid @ [Marker]" using RecvMarker \<open>cid \<noteq> cid'\<close> assms by auto
then show ?thesis
by (simp add: takeWhile_tail)
qed
moreover have "snd (cs c cid) = snd (cs c' cid)"
proof (cases "has_snapshotted c p")
case True
- then have "cs c cid = cs c' cid" using RecvMarker `cid \<noteq> cid'` assms by simp
+ then have "cs c cid = cs c' cid" using RecvMarker \<open>cid \<noteq> cid'\<close> assms by simp
then show ?thesis by simp
next
case False
then show ?thesis
using RecvMarker \<open>cid \<noteq> cid'\<close> assms(1) assms(4) no_self_channel by auto
qed
ultimately show ?thesis by simp
qed
qed
lemma snapshot_step_cs_preservation_q:
assumes
"c \<turnstile> ev \<mapsto> c'" and
"~ regular_event ev" and
"occurs_on ev = q" and
"channel cid = Some (p, q)" and
"Marker \<notin> set (msgs c cid)" and
"~ has_snapshotted c q"
shows
"map Msg (fst (cs c cid)) @ takeWhile ((\<noteq>) Marker) (msgs c cid)
= map Msg (fst (cs c' cid)) @ takeWhile ((\<noteq>) Marker) (msgs c' cid)
\<and> snd (cs c' cid) = Recording"
proof -
have "isSnapshot ev \<or> isRecvMarker ev" using assms nonregular_event by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have Snapshot: "ev = Snapshot q" by (metis event.collapse(4) assms(3))
then have "fst (cs c cid) = fst (cs c' cid)"
using assms(1) assms(2) regular_event same_cs_if_not_recv by blast
moreover have "takeWhile ((\<noteq>) Marker) (msgs c cid)
= takeWhile ((\<noteq>) Marker) (msgs c' cid)"
proof -
have "msgs c' cid = msgs c cid" using assms Snapshot
by (metis distributed_system.next_snapshot distributed_system_axioms eq_fst_iff no_self_channel option.inject)
then show ?thesis by simp
qed
moreover have "snd (cs c' cid) = Recording" using assms Snapshot by auto
ultimately show ?thesis by simp
next
case 2
then obtain cid' r where RecvMarker: "ev = RecvMarker cid' q r" by (metis event.collapse(5) assms(3))
have "cid \<noteq> cid'"
using RecvMarker RecvMarker_implies_Marker_in_set assms(1) assms(5) by blast
have "fst (cs c cid) = fst (cs c' cid)"
using assms(1) assms(2) regular_event same_cs_if_not_recv by blast
moreover have "takeWhile ((\<noteq>) Marker) (msgs c cid)
= takeWhile ((\<noteq>) Marker) (msgs c' cid)"
proof -
have "\<nexists>r. channel cid = Some (q, r)"
using assms(4) no_self_channel by auto
- with RecvMarker assms `cid \<noteq> cid'` have "msgs c cid = msgs c' cid" by (cases "has_snapshotted c r", auto)
+ with RecvMarker assms \<open>cid \<noteq> cid'\<close> have "msgs c cid = msgs c' cid" by (cases "has_snapshotted c r", auto)
then show ?thesis by simp
qed
- moreover have "snd (cs c' cid) = Recording" using assms RecvMarker `cid \<noteq> cid'` by simp
+ moreover have "snd (cs c' cid) = Recording" using assms RecvMarker \<open>cid \<noteq> cid'\<close> by simp
ultimately show ?thesis by simp
qed
qed
lemma Marker_in_channel_implies_not_done:
assumes
"trace init t final" and
"Marker : set (msgs (S t i) cid)" and
"channel cid = Some (p, q)" and
"i \<le> length t"
shows
"snd (cs (S t i) cid) \<noteq> Done"
proof (rule ccontr)
assume is_done: "~ snd (cs (S t i) cid) \<noteq> Done"
let ?t = "take i t"
have tr: "trace init ?t (S t i)"
using assms(1) exists_trace_for_any_i by blast
have "\<exists>q p. RecvMarker cid q p \<in> set ?t"
by (metis (mono_tags, lifting) assms(3) distributed_system.trace.simps distributed_system_axioms done_only_from_recv_marker_trace computation.no_initial_channel_snapshot computation_axioms is_done list.discI recording_state.simps(4) snd_conv tr)
then obtain j where RecvMarker: "\<exists>q p. t ! j = RecvMarker cid q p" "j < i"
by (metis (no_types, lifting) assms(4) in_set_conv_nth length_take min.absorb2 nth_take)
then have step_j: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis Suc_eq_plus1 assms(1) distributed_system.step_Suc distributed_system_axioms assms(4) less_le_trans)
then have "t ! j = RecvMarker cid q p"
by (metis RecvMarker(1) RecvMarker_given_channel assms(3) event.disc(25) event.sel(10) happen_implies_can_occur)
then have "Marker \<notin> set (msgs (S t (j+1)) cid)"
using assms(1) assms(3) no_marker_after_RecvMarker step_j by presburger
moreover have "has_snapshotted (S t (j+1)) p"
using Suc_eq_plus1 \<open>t ! j = RecvMarker cid q p\<close> assms(1) recv_marker_means_snapshotted snapshot_state_unchanged step_j by presburger
ultimately have "Marker \<notin> set (msgs (S t i) cid)"
by (metis RecvMarker(2) Suc_eq_plus1 Suc_leI assms(1,3) assms(4) no_marker_and_snapshotted_implies_no_more_markers_trace)
then show False using assms by simp
qed
lemma keep_empty_if_no_events:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
msgs (S t i) cid = [];
has_snapshotted (S t i) p;
channel cid = Some (p, q);
\<forall>k. i \<le> k \<and> k < j \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p \<rbrakk>
\<Longrightarrow> msgs (S t j) cid = []"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
proof -
have "i < length t"
using Suc.hyps(2) Suc.prems(3) by linarith
then show ?thesis
by (metis (full_types) Suc.prems(1) Suc_eq_plus1 step_Suc)
qed
have "msgs (S t (i+1)) cid = []"
proof (cases "t ! i")
case (Snapshot r)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
moreover have "can_occur (t ! i) (S t i)"
using happen_implies_can_occur local.step by blast
ultimately show False using can_occur_def Snapshot Suc by simp
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
using Snapshot local.step Suc by auto
then show ?thesis using Suc by simp
next
case (RecvMarker cid' r s)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "msgs (S t i) cid \<noteq> []"
by (metis RecvMarker length_greater_0_conv linorder_not_less list.size(3) local.step nat_less_le recv_marker_other_channels_not_shrinking)
then show False using Suc by simp
qed
then show ?thesis
proof (cases "has_snapshotted (S t i) r")
case True
- then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using RecvMarker Suc step `cid \<noteq> cid'` by auto
+ then have "msgs (S t (i+1)) cid = msgs (S t i) cid" using RecvMarker Suc step \<open>cid \<noteq> cid'\<close> by auto
then show ?thesis using Suc by simp
next
case False
have "r \<noteq> p"
using False Suc.prems(5) by blast
- then show ?thesis using RecvMarker Suc step `cid \<noteq> cid'` False by simp
+ then show ?thesis using RecvMarker Suc step \<open>cid \<noteq> cid'\<close> False by simp
qed
next
case (Trans r u u')
then show ?thesis using Suc step by simp
next
case (Send cid' r s u u' m)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
then have "occurs_on (t ! i) = p \<and> regular_event (t ! i)" using Send by simp
moreover have "i \<le> i \<and> i < j" using Suc by simp
ultimately show False using Suc.prems by blast
qed
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "channel cid = channel cid'" by auto
then have "channel cid' = Some (r, s)" using Send step can_occur_def by simp
- then show False using Suc `r \<noteq> p` `~ cid \<noteq> cid'` by auto
+ then show False using Suc \<open>r \<noteq> p\<close> \<open>~ cid \<noteq> cid'\<close> by auto
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
using step Send Suc by simp
then show ?thesis using Suc by simp
next
case (Recv cid' r s u u' m)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "msgs (S t i) cid \<noteq> []"
using Recv local.step by auto
then show False using Suc by simp
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto
then show ?thesis using Suc by simp
qed
moreover have "\<forall>k. i+1 \<le> k \<and> k < j \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p"
using Suc by simp
moreover have "has_snapshotted (S t (i+1)) p"
by (meson Suc.prems(1) Suc.prems(5) discrete less_not_refl nat_le_linear snapshot_stable_ver_3)
moreover have "i+1 \<le> j" using Suc by simp
moreover have "j \<le> length t" using Suc by simp
moreover have "j - (i+1) = n" using Suc by linarith
ultimately show ?case using Suc by blast
qed
lemma last_unchanged_or_empty_if_no_events:
shows
"\<lbrakk> trace init t final; i \<le> j; j \<le> length t;
msgs (S t i) cid \<noteq> [];
last (msgs (S t i) cid) = Marker;
has_snapshotted (S t i) p;
channel cid = Some (p, q);
\<forall>k. i \<le> k \<and> k < j \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p \<rbrakk>
\<Longrightarrow> msgs (S t j) cid = [] \<or> (msgs (S t j) cid \<noteq> [] \<and> last (msgs (S t j) cid) = Marker)"
proof (induct "j - i" arbitrary: i)
case 0
then show ?case
by auto
next
case (Suc n)
then have step: "(S t i) \<turnstile> (t ! i) \<mapsto> (S t (i+1))"
proof -
have "i < length t"
using Suc.hyps(2) Suc.prems(3) by linarith
then show ?thesis
by (metis (full_types) Suc.prems(1) Suc_eq_plus1 step_Suc)
qed
have msgs_s_ip1: "msgs (S t (i+1)) cid = [] \<or> (msgs (S t (i+1)) cid \<noteq> [] \<and> last (msgs (S t (i+1)) cid) = Marker)"
proof (cases "t ! i")
case (Snapshot r)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
moreover have "can_occur (t ! i) (S t i)"
using happen_implies_can_occur local.step by blast
ultimately show False using can_occur_def Snapshot Suc by simp
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
using Snapshot local.step Suc by auto
then show ?thesis using Suc by simp
next
case (RecvMarker cid' r s)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "msgs (S t (i+1)) cid = []"
proof -
have "Marker # msgs (S t (i+1)) cid = msgs (S t i) cid"
using RecvMarker True local.step by auto
then show ?thesis
proof -
assume a1: "Marker # msgs (S t (i + 1)) cid = msgs (S t i) cid"
have "i < j"
by (metis (no_types) Suc.hyps(2) Suc.prems(2) Suc_neq_Zero diff_is_0_eq le_neq_implies_less)
then have "i < length t"
using Suc.prems(3) less_le_trans by blast
then show ?thesis
using a1 by (metis (no_types) Marker_in_channel_implies_not_done RecvMarker Suc.prems(1) Suc.prems(5) Suc.prems(7) Suc_eq_plus1 Suc_le_eq True last_ConsR last_in_set recv_marker_means_cs_Done)
qed
qed
then show ?thesis by simp
next
case False
then show ?thesis
proof (cases "has_snapshotted (S t i) r")
case True
then show ?thesis
using False RecvMarker Suc.prems(5) local.step by auto
next
case False
then have "r \<noteq> p"
using Suc.prems(6) by blast
- with RecvMarker False Suc.prems step `cid \<noteq> cid'` show ?thesis by auto
+ with RecvMarker False Suc.prems step \<open>cid \<noteq> cid'\<close> show ?thesis by auto
qed
qed
next
case (Trans r u u')
then show ?thesis using Suc step by simp
next
case (Send cid' r s u u' m)
have "r \<noteq> p"
proof (rule ccontr)
assume "~ r \<noteq> p"
then have "occurs_on (t ! i) = p \<and> regular_event (t ! i)" using Send by simp
moreover have "i \<le> i \<and> i < j" using Suc by simp
ultimately show False using Suc.prems by blast
qed
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "channel cid = channel cid'" by auto
then have "channel cid' = Some (r, s)" using Send step can_occur_def by simp
- then show False using Suc `r \<noteq> p` `~ cid \<noteq> cid'` by auto
+ then show False using Suc \<open>r \<noteq> p\<close> \<open>~ cid \<noteq> cid'\<close> by auto
qed
then have "msgs (S t i) cid = msgs (S t (i+1)) cid"
using step Send by simp
then show ?thesis using Suc by simp
next
case (Recv cid' r s u u' m)
then show ?thesis
proof (cases "cid = cid'")
case True
then have "msgs (S t i) cid = Msg m # msgs (S t (i+1)) cid"
using Recv local.step by auto
then have "last (msgs (S t (i+1)) cid) = Marker"
by (metis Suc.prems(5) last.simps message.simps(3))
then show ?thesis by blast
next
case False
then have "msgs (S t i) cid = msgs (S t (i+1)) cid" using Recv step by auto
then show ?thesis using Suc by simp
qed
qed
then show ?case
proof (elim disjE, goal_cases)
case 1
moreover have "trace init t final" using Suc by simp
moreover have "i+1 \<le> j" using Suc by simp
moreover have "j \<le> length t" using Suc by simp
moreover have "has_snapshotted (S t (i+1)) p"
using Suc.prems(6) local.step snapshot_state_unchanged by auto
moreover have "j - (i+1) = n" using Suc by linarith
moreover have "\<forall>k. i+1 \<le> k \<and> k < j \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p"
using Suc by auto
ultimately have "msgs (S t j) cid = []" using keep_empty_if_no_events Suc.prems(7) by blast
then show ?thesis by simp
next
case 2
moreover have "trace init t final" using Suc by simp
moreover have "i+1 \<le> j" using Suc by simp
moreover have "j \<le> length t" using Suc by simp
moreover have "has_snapshotted (S t (i+1)) p"
using Suc.prems(6) local.step snapshot_state_unchanged by auto
moreover have "j - (i+1) = n" using Suc by linarith
moreover have "\<forall>k. i+1 \<le> k \<and> k < j \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p"
using Suc by auto
ultimately show ?thesis using Suc.prems(7) Suc.hyps by blast
qed
qed
lemma cs_after_all_prerecording_events:
assumes
"trace init t final" and
"\<forall>i'. i' \<ge> i \<longrightarrow> ~ prerecording_event t i'" and
"\<forall>j'. j' < i \<longrightarrow> ~ postrecording_event t j'" and
"i \<le> length t"
shows
"cs_equal_to_snapshot (S t i) final"
proof (unfold cs_equal_to_snapshot_def, rule allI, rule impI)
fix cid
assume "channel cid \<noteq> None"
then obtain p q where chan: "channel cid = Some (p, q)" by auto
have cs_done: "snd (cs (S t (length t)) cid) = Done"
using chan all_channels_done_in_final_state assms(1) final_is_s_t_len_t by blast
have "filter ((\<noteq>) Marker) (msgs (S t i) cid) = takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)" (is ?B)
proof (rule ccontr)
let ?m = "msgs (S t i) cid"
assume "~ ?B"
then obtain j k where range: "j < k" "k < length ?m" and "?m ! j = Marker" "?m ! k \<noteq> Marker"
using filter_neq_takeWhile by metis
then have "Marker \<in> set ?m"
by (metis less_trans nth_mem)
moreover have "last ?m \<noteq> Marker"
proof -
have "\<forall>l. l < length ?m \<and> l \<noteq> j \<longrightarrow> ?m ! l \<noteq> Marker"
using chan \<open>j < k\<close> \<open>k < length (msgs (S t i) cid)\<close> \<open>msgs (S t i) cid ! j = Marker\<close> assms(1) at_most_one_marker calculation less_trans by blast
moreover have "last ?m = ?m ! (length ?m - 1)"
by (metis \<open>Marker \<in> set (msgs (S t i) cid)\<close> empty_iff last_conv_nth list.set(1))
moreover have "length ?m - 1 \<noteq> j" using range by auto
ultimately show ?thesis using range by auto
qed
moreover have "i \<le> length t"
using chan assms(1) calculation(1) computation.exists_next_marker_free_state computation.no_change_if_ge_length_t computation_axioms nat_le_linear by fastforce
ultimately have "\<exists>j. j < i \<and> postrecording_event t j"
using chan assms(1) msg_after_marker_and_nonempty_implies_postrecording_event by auto
then show False using assms by auto
qed
moreover have "takeWhile ((\<noteq>) Marker) (msgs (S t i) cid) = map Msg (fst (cs final cid))"
proof (cases "snd (cs (S t i) cid)")
case NotStarted
text \<open>show that q and p have to snapshot, and then reduce it to the case below depending on
the order they snapshotted in\<close>
have nsq: "~ has_snapshotted (S t i) q"
using NotStarted chan assms(1) cs_in_initial_state_implies_not_snapshotted by auto
obtain j where snap_q: "~ has_snapshotted (S t j) q" "has_snapshotted (S t (j+1)) q"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have step_q: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis \<open>\<not> ps (S t j) q \<noteq> None\<close> add.commute assms(1) le_SucI le_eq_less_or_eq le_refl linorder_neqE_nat no_change_if_ge_length_t plus_1_eq_Suc snap_q step_Suc)
obtain k where snap_p: "~ has_snapshotted (S t k) p" "has_snapshotted (S t (k+1)) p"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have bound: "i \<le> j"
proof (rule ccontr)
assume "~ i \<le> j"
then have "i \<ge> j+1" by simp
then have "has_snapshotted (S t i) q"
by (meson assms(1) computation.snapshot_stable_ver_3 computation_axioms snap_q(2))
then show False using nsq by simp
qed
have step_p: "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (k+1))"
by (metis \<open>\<not> ps (S t k) p \<noteq> None\<close> add.commute assms(1) le_SucI le_eq_less_or_eq le_refl linorder_neqE_nat no_change_if_ge_length_t plus_1_eq_Suc snap_p step_Suc)
have oq: "occurs_on (t ! j) = q"
proof (rule ccontr)
assume "~ occurs_on (t ! j) = q"
then have "has_snapshotted (S t j) q = has_snapshotted (S t (j+1)) q"
using no_state_change_if_no_event step_q by auto
then show False using snap_q by blast
qed
have op: "occurs_on (t ! k) = p"
proof (rule ccontr)
assume "~ occurs_on (t ! k) = p"
then have "has_snapshotted (S t k) p = has_snapshotted (S t (k+1)) p"
using no_state_change_if_no_event step_p by auto
then show False using snap_p by blast
qed
have "p \<noteq> q" using chan no_self_channel by blast
then have "j \<noteq> k" using oq op event_occurs_on_unique by blast
show ?thesis
proof (cases "j < k")
case True
then have "msgs (S t i) cid = msgs (S t j) cid \<and> cs (S t i) cid = cs (S t j) cid"
proof -
have "\<forall>k. i \<le> k \<and> k < j \<longrightarrow> occurs_on (t ! k) \<noteq> p \<and> occurs_on (t ! k) \<noteq> q" (is ?Q)
proof (rule ccontr)
assume "~ ?Q"
then obtain l where range: "i \<le> l" "l < j" and "occurs_on (t ! l) = p \<or> occurs_on (t ! l) = q" by blast
then show False
proof (elim disjE, goal_cases)
case 1
then show ?thesis
proof (cases "regular_event (t ! l)")
case True
- have "l < k" using range `j < k` by simp
- have "~ has_snapshotted (S t l) p" using snap_p(1) range `j < k` snapshot_stable_ver_3 assms(1) by simp
+ have "l < k" using range \<open>j < k\<close> by simp
+ have "~ has_snapshotted (S t l) p" using snap_p(1) range \<open>j < k\<close> snapshot_stable_ver_3 assms(1) by simp
then have "prerecording_event t l" using True "1" prerecording_event
using s_def snap_q(1) snap_q(2) by fastforce
then show False using assms range by blast
next
case False
then have step_l: "(S t l) \<turnstile> t ! l \<mapsto> (S t (l+1))"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD True assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all)
then have "has_snapshotted (S t (l+1)) p" using False nonregular_event_induces_snapshot
by (metis "1"(3) snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI True assms(1) less_imp_le_nat local.range(2) snap_p(1) snapshot_stable_ver_3)
qed
next
case 2
then show ?thesis
proof (cases "regular_event (t ! l)")
case True
- have "~ has_snapshotted (S t l) q" using snap_q(1) range `j < k` snapshot_stable_ver_3 assms(1) by simp
+ have "~ has_snapshotted (S t l) q" using snap_q(1) range \<open>j < k\<close> snapshot_stable_ver_3 assms(1) by simp
then have "prerecording_event t l" using True "2" prerecording_event
using s_def snap_q(2) by fastforce
then show False using assms range by blast
next
case False
then have step_l: "(S t l) \<turnstile> t ! l \<mapsto> (S t (l+1))"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD True assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all)
then have "has_snapshotted (S t (l+1)) q" using False nonregular_event_induces_snapshot
by (metis "2"(3) snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI assms(1) range(2) snap_q(1) snapshot_stable_ver_3)
qed
qed
qed
moreover have "j \<le> length t"
proof (rule ccontr)
assume "~ j \<le> length t"
then have "S t j = S t (j+1)" using no_change_if_ge_length_t assms by simp
then show False using snap_q by auto
qed
ultimately show ?thesis using chan same_messages_if_no_occurrence_trace assms less_imp_le bound by blast
qed
moreover have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)
\<and> snd (cs (S t (j+1)) cid) = Recording"
proof -
have "~ regular_event (t ! j)" using snap_q
using regular_event_cannot_induce_snapshot step_q by blast
moreover have "Marker \<notin> set (msgs (S t j) cid)"
by (meson chan True assms(1) computation.no_marker_if_no_snapshot computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_p(1))
ultimately show ?thesis using oq snapshot_step_cs_preservation_q step_q chan snap_q(1) by blast
qed
moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid)
= map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)"
proof -
have "snd (cs (S t (j+1)) cid) = Recording" using calculation by simp
moreover have "\<forall>a. j+1 \<le> a \<and> a < k \<longrightarrow> ~ occurs_on (t ! a) = p" (is ?R)
proof (rule ccontr)
assume "~ ?R"
then obtain a where "j+1 \<le> a" "a < k" and ocp: "occurs_on (t ! a) = p" by blast
have "a < length t"
proof -
have "k < length t"
proof (rule ccontr)
assume "~ k < length t"
then have "S t k = S t (k+1)"
using assms(1) no_change_if_ge_length_t by auto
then show False using snap_p by auto
qed
- then show ?thesis using `a < k` by simp
+ then show ?thesis using \<open>a < k\<close> by simp
qed
then show False
proof (cases "regular_event (t ! a)")
case True
have "~ has_snapshotted (S t a) p"
by (meson \<open>a < k\<close> assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_p(1))
- then have "prerecording_event t a" using `a < length t` ocp True prerecording_event by simp
- then show False using `j+1 \<le> a` `j \<ge> i` assms by auto
+ then have "prerecording_event t a" using \<open>a < length t\<close> ocp True prerecording_event by simp
+ then show False using \<open>j+1 \<le> a\<close> \<open>j \<ge> i\<close> assms by auto
next
case False
then have "(S t a) \<turnstile> (t ! a) \<mapsto> (S t (a+1))"
using \<open>a < length t\<close> assms(1) step_Suc by auto
then have "has_snapshotted (S t (a+1)) p"
by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI \<open>a < k\<close> assms(1) snap_p(1) snapshot_stable_ver_3)
qed
qed
moreover have "~ has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 Suc_le_eq True assms(1) computation.snapshot_stable_ver_2 computation_axioms snap_p(1))
ultimately show ?thesis using chan cs_when_recording_2 True assms(1) by auto
qed
moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid)
= map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k+1)) cid)"
proof -
have "\<not> regular_event (t ! k)"
using regular_event_preserves_process_snapshots snap_p(1) snap_p(2) step_p by force
then show ?thesis
using chan computation.snapshot_step_cs_preservation_p computation_axioms op step_p by fastforce
qed
moreover have "map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k+1)) cid)
= map Msg (fst (cs final cid))"
proof -
have f1: "\<forall>f p pa pb c ca es n a na. \<not> computation f p pa pb (c::('a, 'b, 'c) configuration) ca \<or> \<not> distributed_system.trace f p pa pb c es ca \<or> ps (distributed_system.s f p pa pb c es n) a = None \<or> \<not> n \<le> na \<or> ps (distributed_system.s f p pa pb c es na) a \<noteq> None"
by (meson computation.snapshot_stable_ver_2)
have f2: "computation channel trans send recv init (S t (length t))"
using assms(1) final_is_s_t_len_t computation_axioms by blast
have f3: "trace init t (S t (length t))"
using assms(1) final_is_s_t_len_t by blast
have f4: "ps (S t k) p = None"
by (meson snap_p(1))
then have f5: "k < length t"
using f3 f2 f1 by (metis le_eq_less_or_eq not_le s_def snap_p(2) take_all)
have "\<not> regular_event (t ! k)"
using f4 by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(2) step_p)
then have f6: "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k + 1)) cid) \<and> snd (cs (S t k) cid) = snd (cs (S t (k + 1)) cid)"
using chan computation.snapshot_step_cs_preservation_p computation_axioms op step_p by fastforce
then have f7: "snd (cs (S t (k + 1)) cid) \<noteq> Done"
using f5 f4 by (metis (no_types) assms(1) chan cs_done_implies_both_snapshotted(1))
have "j + 1 \<le> k + 1"
using True by linarith
then have "snd (cs (S t (k + 1)) cid) = Recording"
using f7 f3 f2 f1 by (meson chan computation.cs_in_initial_state_implies_not_snapshotted recording_state.exhaust snap_q(2))
then show ?thesis
using f6 f5 by (metis (no_types) Suc_eq_plus1 Suc_leI assms(1) chan cs_done cs_done_implies_both_snapshotted(1) cs_when_recording final_is_s_t_len_t le_eq_less_or_eq snap_p(2))
qed
ultimately show ?thesis
by (metis (no_types, lifting) chan Nil_is_map_conv assms(1) computation.no_initial_channel_snapshot computation_axioms fst_conv no_recording_cs_if_not_snapshotted self_append_conv2 snap_q(1))
next
case False
- then have "k < j" using `j \<noteq> k` False by simp
+ then have "k < j" using \<open>j \<noteq> k\<close> False by simp
then have "map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)
= map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)"
proof (cases "i \<le> k")
case True
then have "msgs (S t i) cid = msgs (S t k) cid \<and> cs (S t i) cid = cs (S t k) cid"
proof -
have "\<forall>j. i \<le> j \<and> j < k \<longrightarrow> occurs_on (t ! j) \<noteq> p \<and> occurs_on (t ! j) \<noteq> q" (is ?Q)
proof (rule ccontr)
assume "~ ?Q"
then obtain l where range: "i \<le> l" "l < k" and "occurs_on (t ! l) = p \<or> occurs_on (t ! l) = q" by blast
then show False
proof (elim disjE, goal_cases)
case 1
then show ?thesis
proof (cases "regular_event (t ! l)")
case True
- have "l < k" using range `k < j` by simp
- have "~ has_snapshotted (S t l) p" using snap_p(1) range `k < j` snapshot_stable_ver_3 assms(1) by simp
+ have "l < k" using range \<open>k < j\<close> by simp
+ have "~ has_snapshotted (S t l) p" using snap_p(1) range \<open>k < j\<close> snapshot_stable_ver_3 assms(1) by simp
then have "prerecording_event t l" using True "1" prerecording_event
using s_def snap_p by fastforce
then show False using assms range by blast
next
case False
then have step_l: "(S t l) \<turnstile> t ! l \<mapsto> (S t (l+1))"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all)
then have "has_snapshotted (S t (l+1)) p" using False nonregular_event_induces_snapshot
by (metis "1"(3) snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI assms(1) local.range(2) snap_p(1) snapshot_stable_ver_3)
qed
next
case 2
then show ?thesis
proof (cases "regular_event (t ! l)")
case True
- have "~ has_snapshotted (S t l) p" using snap_p(1) range `k < j` snapshot_stable_ver_3 assms(1) by simp
+ have "~ has_snapshotted (S t l) p" using snap_p(1) range \<open>k < j\<close> snapshot_stable_ver_3 assms(1) by simp
moreover have "l < length t"
using \<open>k < j\<close> local.range(2) s_def snap_q(1) snap_q(2) by force
ultimately have "prerecording_event t l" using True "2" prerecording_event
proof -
have "l \<le> j"
by (meson False \<open>l < k\<close> less_trans not_less)
then show ?thesis
by (metis (no_types) True \<open>l < length t\<close> \<open>occurs_on (t ! l) = q\<close> assms(1) computation.prerecording_event computation.snapshot_stable_ver_2 computation_axioms snap_q(1))
qed
then show False using assms range by blast
next
case False
then have step_l: "(S t l) \<turnstile> t ! l \<mapsto> (S t (l+1))"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_lessD assms(1) distributed_system.step_Suc distributed_system_axioms less_trans_Suc linorder_not_le local.range(2) s_def snap_p(1) snap_p(2) take_all)
then have "has_snapshotted (S t (l+1)) q" using False nonregular_event_induces_snapshot
by (metis "2"(3) snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI \<open>k < j\<close> assms(1) less_imp_le_nat local.range(2) snap_q(1) snapshot_stable_ver_3)
qed
qed
qed
moreover have "k \<le> length t"
proof (rule ccontr)
assume "~ k \<le> length t"
then have "S t k = S t (k+1)" using no_change_if_ge_length_t assms by simp
then show False using snap_p by auto
qed
ultimately show ?thesis using chan same_messages_if_no_occurrence_trace assms True less_imp_le by auto
qed
moreover have "map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid)
= map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k+1)) cid)
\<and> snd (cs (S t (k+1)) cid) = NotStarted"
proof -
have "~ regular_event (t ! k)" using snap_p
using regular_event_cannot_induce_snapshot step_p by blast
then show ?thesis
using calculation op snapshot_step_cs_preservation_p step_p chan NotStarted by auto
qed
moreover have "map Msg (fst (cs (S t (k+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k+1)) cid)
= map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)"
proof -
have "\<forall>a. k+1 \<le> a \<and> a < j \<longrightarrow> ~ occurs_on (t ! a) = q" (is ?R)
proof (rule ccontr)
assume "~ ?R"
then obtain a where "k+1 \<le> a" "a < j" and ocp: "occurs_on (t ! a) = q" by blast
have "a < length t"
proof -
have "j < length t"
proof (rule ccontr)
assume "~ j < length t"
then have "S t j = S t (j+1)"
using assms(1) no_change_if_ge_length_t by auto
then show False using snap_q by auto
qed
- then show ?thesis using `a < j` by simp
+ then show ?thesis using \<open>a < j\<close> by simp
qed
then show False
proof (cases "regular_event (t ! a)")
case True
have "~ has_snapshotted (S t a) q"
by (meson \<open>a < j\<close> assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_q(1))
- then have "prerecording_event t a" using `a < length t` ocp True prerecording_event by simp
- then show False using `k+1 \<le> a` `k \<ge> i` assms by auto
+ then have "prerecording_event t a" using \<open>a < length t\<close> ocp True prerecording_event by simp
+ then show False using \<open>k+1 \<le> a\<close> \<open>k \<ge> i\<close> assms by auto
next
case False
then have "(S t a) \<turnstile> (t ! a) \<mapsto> (S t (a+1))"
using \<open>a < length t\<close> assms(1) step_Suc by auto
then have "has_snapshotted (S t (a+1)) q"
by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI \<open>a < j\<close> assms(1) snap_q(1) snapshot_stable_ver_3)
qed
qed
moreover have "Marker : set (msgs (S t (k+1)) cid)"
using chan \<open>map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k + 1)) cid) \<and> snd (cs (S t (k + 1)) cid) = NotStarted\<close> assms(1) cs_in_initial_state_implies_not_snapshotted marker_must_stay_if_no_snapshot snap_p(2) by blast
moreover have "has_snapshotted (S t (k+1)) p"
using snap_p(2) by blast
moreover have "~ has_snapshotted (S t (k+1)) q"
using chan \<open>map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k + 1)) cid) \<and> snd (cs (S t (k + 1)) cid) = NotStarted\<close> assms(1) cs_in_initial_state_implies_not_snapshotted by blast
moreover have "k+1 \<le> j"
using \<open>k < j\<close> by auto
moreover have "trace init t final" using assms by simp
moreover have "snd (cs (S t (k+1)) cid) = NotStarted"
using \<open>map Msg (fst (cs (S t k) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t k) cid) = map Msg (fst (cs (S t (k + 1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (k + 1)) cid) \<and> snd (cs (S t (k + 1)) cid) = NotStarted\<close> by blast
ultimately show ?thesis using cs_when_recording_3 chan by simp
qed
ultimately show ?thesis by simp
next
case False
show ?thesis
proof -
have "has_snapshotted (S t i) p"
by (metis False Suc_eq_plus1 assms(1) not_less_eq_eq snap_p(2) snapshot_stable_ver_3)
moreover have "~ has_snapshotted (S t i) q"
using nsq by auto
moreover have "Marker : set (msgs (S t i) cid)"
using chan assms(1) calculation(1) marker_must_stay_if_no_snapshot nsq by blast
moreover have "\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = q" (is ?R)
proof (rule ccontr)
assume "~ ?R"
then obtain k where "i \<le> k" "k < j" and ocp: "occurs_on (t ! k) = q" by blast
have "k < length t"
proof -
have "j < length t"
proof (rule ccontr)
assume "~ j < length t"
then have "S t j = S t (j+1)"
using assms(1) no_change_if_ge_length_t by auto
then show False using snap_q by auto
qed
- then show ?thesis using `k < j` by simp
+ then show ?thesis using \<open>k < j\<close> by simp
qed
then show False
proof (cases "regular_event (t ! k)")
case True
have "~ has_snapshotted (S t k) q"
- by (meson `k < j` assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_q(1))
- then have "prerecording_event t k" using `k < length t` ocp True prerecording_event by simp
- then show False using `i \<le> j` `k \<ge> i` assms by auto
+ by (meson \<open>k < j\<close> assms(1) computation.snapshot_stable_ver_2 computation_axioms less_imp_le_nat snap_q(1))
+ then have "prerecording_event t k" using \<open>k < length t\<close> ocp True prerecording_event by simp
+ then show False using \<open>i \<le> j\<close> \<open>k \<ge> i\<close> assms by auto
next
case False
then have "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (k+1))"
using \<open>k < length t\<close> assms(1) step_Suc by auto
then have "has_snapshotted (S t (k+1)) q"
by (metis False ocp nonregular_event_induces_snapshot snapshot_state_unchanged)
then show False
by (metis Suc_eq_plus1 Suc_leI \<open>k < j\<close> assms(1) snap_q(1) snapshot_stable_ver_3)
qed
qed
ultimately show ?thesis using cs_when_recording_3
using NotStarted assms(1) bound chan by auto
qed
qed
moreover have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs final cid))"
proof (cases "\<exists>q p. t ! j = RecvMarker cid q p")
case True
then have "fst (cs (S t j) cid) = fst (cs (S t (j+1)) cid)"
using step_q by auto
moreover have RecvMarker: "t ! j = RecvMarker cid q p"
proof -
have "can_occur (t ! j) (S t j)" using happen_implies_can_occur step_q by simp
then show ?thesis
using RecvMarker_given_channel True chan by force
qed
moreover have "takeWhile ((\<noteq>) Marker) (msgs (S t j) cid) = []"
proof -
have "can_occur (t ! j) (S t j)" using happen_implies_can_occur step_q by simp
then have "length (msgs (S t j) cid) > 0 \<and> hd (msgs (S t j) cid) = Marker"
using RecvMarker can_occur_def by auto
then show ?thesis
by (metis (mono_tags, lifting) hd_conv_nth length_greater_0_conv nth_mem set_takeWhileD takeWhile_nth)
qed
moreover have "snd (cs (S t (j+1)) cid) = Done" using step_q True by auto
moreover have "cs (S t (j+1)) cid = cs final cid" using chan calculation cs_done_implies_same_snapshots assms(1)
by (metis final_is_s_t_len_t nat_le_linear no_change_if_ge_length_t)
ultimately show ?thesis
by simp
next
case False
have "~ regular_event (t ! j)"
using regular_event_preserves_process_snapshots snap_q(1) snap_q(2) step_q by auto
then have "isSnapshot (t ! j) \<or> isRecvMarker (t ! j)" using nonregular_event by auto
then have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)
\<and> snd (cs (S t (j+1)) cid) = Recording"
proof (elim disjE, goal_cases)
case 1
have Snapshot: "t ! j = Snapshot q"
using "1" oq by auto
then have "msgs (S t j) cid = msgs (S t (j+1)) cid"
using \<open>p \<noteq> q\<close> step_q chan by auto
moreover have "cs (S t (j+1)) cid = (fst (cs (S t j) cid), Recording)"
using step_q Snapshot chan by simp
ultimately show ?thesis by simp
next
case 2
obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' q r"
by (metis "2" event.collapse(5) oq)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "channel cid = channel cid'" by simp
then have "channel cid' = Some (r, q)"
using False RecvMarker \<open>\<not> cid \<noteq> cid'\<close> by blast
then show False
using False RecvMarker \<open>\<not> cid \<noteq> cid'\<close> by blast
qed
then have "msgs (S t j) cid = msgs (S t (j+1)) cid"
- using `cid \<noteq> cid'` step_q snap_q RecvMarker chan `p \<noteq> q` by simp
+ using \<open>cid \<noteq> cid'\<close> step_q snap_q RecvMarker chan \<open>p \<noteq> q\<close> by simp
moreover have "cs (S t (j+1)) cid = (fst (cs (S t j) cid), Recording)"
- using \<open>p \<noteq> q\<close> `cid \<noteq> cid'`step_q snap_q RecvMarker chan by auto
+ using \<open>p \<noteq> q\<close> \<open>cid \<noteq> cid'\<close>step_q snap_q RecvMarker chan by auto
ultimately show ?thesis by simp
qed
moreover have "map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)
= map Msg (fst (cs final cid))"
proof -
have "snd (cs (S t (j+1)) cid) = Recording"
using calculation by blast
moreover have "has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 Suc_leI \<open>k < j\<close> assms(1) le_add1 snap_p(2) snapshot_stable_ver_3)
moreover have "has_snapshotted (S t (j+1)) q" using snap_q by auto
moreover have "j < length t"
by (metis (no_types, lifting) chan Suc_eq_plus1 assms(1) cs_done cs_done_implies_both_snapshotted(2) computation.no_change_if_ge_length_t computation.snapshot_stable_ver_3 computation_axioms leI le_Suc_eq snap_q(1) snap_q(2))
ultimately show ?thesis using cs_when_recording assms(1) cs_done final_is_s_t_len_t
proof -
assume a1: "j < length t"
assume a2: "trace init t final"
assume a3: "snd (cs (S t (length t)) cid) = Done"
assume a4: "snd (cs (S t (j + 1)) cid) = Recording"
assume a5: "ps (S t (j + 1)) p \<noteq> None"
assume a6: "\<And>t. trace init t final \<Longrightarrow> final = S t (length t)"
assume a7: "\<And>i j t p cid q. \<lbrakk>i < j; j \<le> length t; trace init t final; ps (S t i) p \<noteq> None; snd (cs (S t i) cid) = Recording; snd (cs (S t j) cid) = Done; channel cid = Some (p, q)\<rbrakk> \<Longrightarrow> map Msg (fst (cs (S t j) cid)) = map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
have "Suc j < length t"
using a3 a2 a1 by (metis (no_types) False Suc_eq_plus1 Suc_lessI chan cs_done_implies_has_snapshotted done_only_from_recv_marker snap_q(1) step_q)
then show ?thesis
using a7 a6 a5 a4 a3 a2 by (metis (no_types) Suc_eq_plus1 chan nat_le_linear)
qed
qed
ultimately show ?thesis by simp
qed
ultimately show ?thesis
by (metis (no_types, lifting) Nil_is_map_conv assms(1) assms(3) chan cs_done cs_done_implies_has_snapshotted cs_not_nil_implies_postrecording_event nat_le_linear nsq self_append_conv2 snapshot_stable_ver_3)
qed
next
case Recording
then obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have snap_q: "has_snapshotted (S t i) q"
using Recording assms(1) chan cs_recording_implies_snapshot by blast
have fst_cs_empty: "cs (S t i) cid = ([], Recording)" (is ?P)
proof (rule ccontr)
assume "~ ?P"
have "snd (cs (S t i) cid) = Recording" using Recording by simp
- moreover have "fst (cs (S t i) cid) \<noteq> []" using `~ ?P` prod.collapse calculation by metis
+ moreover have "fst (cs (S t i) cid) \<noteq> []" using \<open>~ ?P\<close> prod.collapse calculation by metis
ultimately have "\<exists>j. j < i \<and> postrecording_event t j"
using assms(1) assms(4) chan cs_not_nil_implies_postrecording_event by blast
then show False using assms by auto
qed
then show ?thesis
proof -
have i_less_len_t: "i < length t"
proof (rule ccontr)
assume "~ i < length t"
then have "snd (cs (S t i) cid) = Done"
by (metis assms(1) cs_done le_eq_less_or_eq nat_le_linear no_change_if_ge_length_t)
then show False using Recording by simp
qed
then have "map Msg (fst (cs final cid))
= map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
proof (cases "j < i")
case True
then have "has_snapshotted (S t i) p"
by (metis Suc_eq_plus1 Suc_leI assms(1) snap_p(2) snapshot_stable_ver_3)
moreover have "length t \<le> length t" by simp
ultimately show ?thesis
using Recording chan assms(1) cs_done cs_when_recording final_is_s_t_len_t i_less_len_t by blast
next
case False
text \<open>need to show that next message that comes into the channel must be marker\<close>
have "\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ occurs_on (t ! k) = p" (is ?P)
proof (rule ccontr)
assume "~ ?P"
then obtain k where "i \<le> k" "k < j" "occurs_on (t ! k) = p" by blast
then show False
proof (cases "regular_event (t ! k)")
case True
then have "prerecording_event t k"
by (metis (no_types, hide_lams) \<open>k < j\<close> \<open>occurs_on (t ! k) = p\<close> all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t computation.prerecording_event computation_axioms less_trans nat_le_linear not_less snap_p(1) snapshot_stable_ver_2)
- then show ?thesis using assms `i \<le> k` by auto
+ then show ?thesis using assms \<open>i \<le> k\<close> by auto
next
case False
then have step_k: "(S t k) \<turnstile> (t ! k) \<mapsto> (S t (Suc k))"
by (metis (no_types, lifting) Suc_leI \<open>k < j\<close> all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t le_Suc_eq less_imp_Suc_add linorder_not_less no_change_if_ge_length_t snap_p(1) step_Suc)
then have "has_snapshotted (S t (Suc k)) p"
by (metis False \<open>occurs_on (t ! k) = p\<close> nonregular_event_induces_snapshot snapshot_state_unchanged)
then have "k \<ge> j"
by (metis Suc_leI \<open>k < j\<close> assms(1) snap_p(1) snapshot_stable_ver_3)
- then show False using `k < j` by simp
+ then show False using \<open>k < j\<close> by simp
qed
qed
moreover have "~ has_snapshotted (S t i) p"
using False assms(1) snap_p(1) snapshot_stable_ver_3 by auto
ultimately have to_snapshot: "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t i) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t i) cid)"
using False chan Recording assms(1) cs_when_recording_2 by auto
have step_j: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis Suc_eq_plus1 Suc_le_eq assms(1) distributed_system.step_Suc distributed_system_axioms computation.no_change_if_ge_length_t computation_axioms le_add1 not_less_eq_eq snap_p(1) snap_p(2))
then have "map Msg (fst (cs (S t j) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t j) cid)
= map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)"
proof -
have o: "~ regular_event (t ! j) \<and> occurs_on (t ! j) = p"
by (metis (no_types, hide_lams) distributed_system.no_state_change_if_no_event distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_j)
then show ?thesis
using chan snapshot_step_cs_preservation_p step_j by blast
qed
moreover have "map Msg (fst (cs final cid))
= map Msg (fst (cs (S t (j+1)) cid)) @ takeWhile ((\<noteq>) Marker) (msgs (S t (j+1)) cid)"
proof -
have "snd (cs (S t (j+1)) cid) = Recording"
proof -
have f1: "ps (S t j) p = None"
by (meson snap_p(1))
then have f2: "j < length t"
by (metis (no_types) all_processes_snapshotted_in_final_state assms(1) final_is_s_t_len_t linorder_not_le snapshot_stable_ver_3)
have "t ! j \<noteq> RecvMarker cid q p"
using f1 by (metis (no_types) Suc_eq_plus1 assms(1) recv_marker_means_snapshotted step_j)
then show ?thesis
using f2 f1 by (meson False assms(1) chan cs_done_implies_both_snapshotted(1) cs_in_initial_state_implies_not_snapshotted cs_not_not_started_stable done_only_from_recv_marker linorder_not_le recording_state.exhaust snap_q snapshot_stable_ver_3 step_j)
qed
moreover have "j+1 < length t"
proof (rule ccontr)
assume "~ j+1 < length t"
then have "snd (cs (S t (j+1)) cid) = Done"
by (metis assms(1) cs_done le_Suc_eq less_imp_Suc_add linorder_not_le no_change_if_ge_length_t)
then show False using calculation by auto
qed
ultimately show ?thesis
using chan snap_p(2) assms final_is_s_t_len_t cs_when_recording cs_done by blast
qed
ultimately show ?thesis using to_snapshot by simp
qed
then show ?thesis using fst_cs_empty by simp
qed
next
case Done
text \<open>msgs must be empty, and cs must also be empty\<close>
have fst_cs_empty: "fst (cs (S t i) cid) = []"
proof (rule ccontr)
assume "~ fst (cs (S t i) cid) = []"
then have "fst (cs (S t 0) cid) \<noteq> fst (cs (S t i) cid)"
by (metis chan assms(1) cs_not_nil_implies_postrecording_event gr_implies_not0 le0)
then have "\<exists>j. j < i \<and> postrecording_event t j"
using chan \<open>fst (cs (S t i) cid) \<noteq> []\<close> assms(1) assms(4) cs_not_nil_implies_postrecording_event by blast
then show False using assms by auto
qed
moreover have "msgs (S t i) cid = []"
proof -
have no_marker: "Marker \<notin> set (msgs (S t i) cid)" (is ?P)
proof (rule ccontr)
assume "~ ?P"
then have "Marker : set (msgs (S t i) cid)" by simp
then have "snd (cs (S t i) cid) \<noteq> Done"
by (metis Marker_in_channel_implies_not_done chan assms(1) nat_le_linear s_def take_all)
then show False using Done by simp
qed
have snap_both: "has_snapshotted (S t i) p \<and> has_snapshotted (S t i) q"
by (metis chan Done assms(1) cs_done_implies_both_snapshotted(1) cs_done_implies_has_snapshotted final_is_s_t_len_t computation.all_processes_snapshotted_in_final_state computation_axioms le_refl not_less s_def take_all)
obtain j where snap_p: "~ has_snapshotted (S t j) p" "has_snapshotted (S t (j+1)) p"
by (metis Suc_eq_plus1 assms(1) exists_snapshot_for_all_p)
have "j < i"
by (meson assms(1) not_le_imp_less snap_both snap_p(1) snapshot_stable_ver_2)
have step_j: "(S t j) \<turnstile> (t ! j) \<mapsto> (S t (j+1))"
by (metis Suc_eq_plus1 assms(1) distributed_system.step_Suc distributed_system_axioms computation.no_change_if_ge_length_t computation_axioms le_add1 linorder_not_less snap_p(1) snap_p(2))
have nonreg_j: "~ regular_event (t ! j)"
by (meson distributed_system.regular_event_cannot_induce_snapshot distributed_system_axioms snap_p(1) snap_p(2) step_j)
have oc_j: "occurs_on (t ! j) = p"
using no_state_change_if_no_event snap_p(1) snap_p(2) step_j by force
have "msgs (S t i) cid = [] \<or> (msgs (S t i) cid \<noteq> [] \<and> last (msgs (S t i) cid) = Marker)"
proof -
have "msgs (S t (j+1)) cid \<noteq> [] \<and> last (msgs (S t (j+1)) cid) = Marker"
proof -
have "msgs (S t (j+1)) cid = msgs (S t j) cid @ [Marker]"
proof -
have "isSnapshot (t ! j) \<or> isRecvMarker (t ! j)" using nonregular_event nonreg_j by blast
then show ?thesis
proof (elim disjE, goal_cases)
case 1
then have "t ! j = Snapshot p" using oc_j by auto
then show ?thesis using step_j chan by auto
next
case 2
then obtain cid' r where RecvMarker: "t ! j = RecvMarker cid' p r"
by (metis event.collapse(5) oc_j)
have "cid \<noteq> cid'"
proof (rule ccontr)
assume "~ cid \<noteq> cid'"
then have "channel cid = channel cid'" by auto
then have "Some (p, q) = Some (r, p)"
by (metis RecvMarker RecvMarker_implies_Marker_in_set assms(1) chan computation.no_marker_if_no_snapshot computation_axioms snap_p(1) step_j)
then show False using no_self_channel chan by simp
qed
then show ?thesis using oc_j snap_p step_j chan RecvMarker by auto
qed
qed
then show ?thesis by auto
qed
moreover have "i \<le> length t" using assms by simp
- moreover have "j+1 \<le> i" using `j < i` by simp
+ moreover have "j+1 \<le> i" using \<open>j < i\<close> by simp
moreover have "\<forall>k. j+1 \<le> k \<and> k < i \<and> regular_event (t ! k) \<longrightarrow> ~ occurs_on (t ! k) = p" (is ?R)
proof (rule ccontr)
assume "~ ?R"
then obtain k where range: "j+1 \<le> k" "k < i" and "regular_event (t ! k)" "occurs_on (t ! k) = p"
by blast
then have "postrecording_event t k" using snap_p
by (meson assms(1) calculation(2) le_trans linorder_not_less pre_if_regular_and_not_post prerecording_event snapshot_stable_ver_2)
then show False using assms range by auto
qed
ultimately show ?thesis
using assms(1) chan last_unchanged_or_empty_if_no_events snap_p(2) by auto
qed
then show ?thesis using no_marker last_in_set by fastforce
qed
ultimately show ?thesis
using chan Done assms(1) assms(4) final_is_s_t_len_t computation.cs_done_implies_same_snapshots computation_axioms by fastforce
qed
ultimately show "filter ((\<noteq>) Marker) (msgs (S t i) cid) = map Msg (fst (cs final cid))" by simp
qed
lemma snapshot_after_all_prerecording_events:
assumes
"trace init t final" and
"\<forall>i'. i' \<ge> i \<longrightarrow> ~ prerecording_event t i'" and
"\<forall>j'. j' < i \<longrightarrow> ~ postrecording_event t j'" and
"i \<le> length t"
shows
"state_equal_to_snapshot (S t i) final"
proof (unfold state_equal_to_snapshot_def, rule conjI)
show "ps_equal_to_snapshot (S t i) final"
using assms ps_after_all_prerecording_events by auto
show "cs_equal_to_snapshot (S t i) final"
using assms cs_after_all_prerecording_events by auto
qed
subsection \<open>Obtaining the desired traces\<close>
abbreviation all_prerecording_before_postrecording where
"all_prerecording_before_postrecording t \<equiv> \<exists>i. (\<forall>j. j < i \<longrightarrow> ~ postrecording_event t j)
\<and> (\<forall>j. j \<ge> i \<longrightarrow> ~ prerecording_event t j)
\<and> i \<le> length t
\<and> trace init t final"
definition count_violations :: "('a, 'b, 'c) trace \<Rightarrow> nat" where
"count_violations t = sum (\<lambda>i. if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0)
{0..<length t}"
lemma violations_ge_0:
shows
"(if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0) \<ge> 0"
by simp
lemma count_violations_ge_0:
shows
"count_violations t \<ge> 0"
by simp
lemma violations_0_implies_all_subterms_0:
assumes
"count_violations t = 0"
shows
"\<forall>i \<in> {0..<length t}. (if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0) = 0"
proof -
have "sum (\<lambda>i. if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0)
{0..<length t} = 0" using count_violations_def assms by simp
then show "\<forall>i \<in> {0..<length t}. (if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0) = 0"
by auto
qed
lemma exists_postrecording_violation_if_count_greater_0:
assumes
"count_violations t > 0"
shows
"\<exists>i. postrecording_event t i \<and> card {j \<in> {i+1..<length t}. prerecording_event t j} > 0" (is ?P)
proof (rule ccontr)
assume "~ ?P"
then have "\<forall>i. ~ postrecording_event t i \<or> card {j \<in> {i+1..<length t}. prerecording_event t j} = 0" by simp
have "count_violations t = 0"
proof (unfold count_violations_def)
have "\<forall>i. (if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0) = 0"
using \<open>\<forall>i. \<not> postrecording_event t i \<or> card {j \<in> {i + 1..<length t}. prerecording_event t j} = 0\<close> by auto
then show "sum (\<lambda>i. if postrecording_event t i
then card {j \<in> {i+1..<length t}. prerecording_event t j}
else 0) {0..<length t} = 0" by simp
qed
then show False using assms by simp
qed
lemma exists_prerecording_violation_when_card_greater_0:
assumes
"card {j \<in> {i+1..<length t}. prerecording_event t j} > 0"
shows
"\<exists>j \<in> {i+1..<length t}. prerecording_event t j"
by (metis (no_types, lifting) Collect_empty_eq assms card_0_eq empty_subsetI finite_atLeastLessThan not_gr_zero subset_card_intvl_is_intvl)
lemma card_greater_0_if_post_after_pre:
assumes
"i < j" and
"postrecording_event t i" and
"prerecording_event t j"
shows
"card {j \<in> {i+1..<length t}. prerecording_event t j} > 0"
proof -
have "j < length t" using prerecording_event assms by auto
have "{j \<in> {i+1..<length t}. prerecording_event t j} \<noteq> empty"
using Suc_eq_plus1 \<open>j < length t\<close> assms(1) assms(3) less_imp_triv by auto
then show ?thesis by fastforce
qed
lemma exists_neighboring_violation_pair:
assumes
"trace init t final" and
"count_violations t > 0"
shows
"\<exists>i j. i < j \<and> postrecording_event t i \<and> prerecording_event t j
\<and> (\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)) \<and> j < length t"
proof -
let ?I = "{i. postrecording_event t i \<and> card {j \<in> {i+1..<length t}. prerecording_event t j} > 0}"
have nonempty_I: "?I \<noteq> empty" using assms exists_postrecording_violation_if_count_greater_0 by blast
have fin_I: "finite ?I"
proof (rule ccontr)
assume "~ finite ?I"
then obtain i where "i > length t" "postrecording_event t i"
by (simp add: postrecording_event)
then show False using postrecording_event by simp
qed
let ?i = "Max ?I"
have no_greater_postrec_violation: "\<forall>i. i > ?i \<longrightarrow> ~ (postrecording_event t i \<and> card {j \<in> {i+1..<length t}. prerecording_event t j} > 0)"
using Max_gr_iff fin_I by blast
have post_i: "postrecording_event t ?i"
using Max_in fin_I nonempty_I by blast
have "card {j \<in> {?i+1..<length t}. prerecording_event t j} > 0"
proof -
have "?i \<in> ?I"
using Max_in fin_I nonempty_I by blast
then show ?thesis by simp
qed
let ?J = "{j \<in> {?i+1..<length t}. prerecording_event t j}"
have nonempty_J: "?J \<noteq> empty"
- using `card {j \<in> {?i+1..<length t}. prerecording_event t j} > 0` exists_prerecording_violation_when_card_greater_0
+ using \<open>card {j \<in> {?i+1..<length t}. prerecording_event t j} > 0\<close> exists_prerecording_violation_when_card_greater_0
by blast
have fin_J: "finite ?J" by auto
let ?j = "Min ?J"
have pre_j: "prerecording_event t ?j"
using Min_in fin_J nonempty_J by blast
have no_smaller_prerec_violation: "\<forall>j \<in> {?i+1..<length t}. j < ?j \<longrightarrow> ~ prerecording_event t j"
using Min_less_iff fin_J by blast
have j_less_len_t: "?j < length t"
using pre_j prerecording_event by blast
have "\<forall>k. (?i < k \<and> k < ?j) \<longrightarrow> ~ regular_event (t ! k)"
proof (rule allI, rule impI)
fix k
assume asm: "?i < k \<and> k < ?j"
then show "~ regular_event (t ! k)"
proof -
have "0_le_k": "0 \<le> k" by simp
have k_less_len_t: "k < length t" using j_less_len_t asm by auto
show ?thesis
proof (rule ccontr)
assume reg_event: "~ ~ regular_event (t ! k)"
then show False
proof (cases "has_snapshotted (S t k) (occurs_on (t ! k))")
case True
then have post_k: "postrecording_event t k" using reg_event k_less_len_t postrecording_event by simp
moreover have "card {j \<in> {k+1..<length t}. prerecording_event t j} > 0"
using post_k pre_j card_greater_0_if_post_after_pre asm pre_j by blast
ultimately show False using no_greater_postrec_violation asm by blast
next
case False
then have pre_k: "prerecording_event t k" using reg_event k_less_len_t prerecording_event by simp
moreover have "k \<in> {?i+1..<length t}" using asm k_less_len_t by simp
ultimately show False using no_smaller_prerec_violation asm by blast
qed
qed
qed
qed
moreover have "?i < ?j" using nonempty_J by auto
ultimately show ?thesis using pre_j post_i j_less_len_t by blast
qed
lemma same_cardinality_post_swap_1:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"trace init t final"
shows
"{k \<in> {0..<i}. prerecording_event t k}
= {k \<in> {0..<i}. prerecording_event (swap_events i j t) k}"
proof -
let ?t = "swap_events i j t"
have same_begin: "take i t = take i ?t" using swap_identical_heads assms by blast
have same_length: "length t = length (swap_events i j t)" using swap_identical_length assms by blast
have a: "\<forall>k. k < i \<longrightarrow> t ! k = ?t ! k"
by (metis nth_take same_begin)
then have "\<forall>k. k < i \<longrightarrow> prerecording_event t k = prerecording_event ?t k"
proof -
have "\<forall>k. k < i \<longrightarrow> S t k = S ?t k" using assms swap_events by simp
then show ?thesis unfolding prerecording_event using a same_length by presburger
qed
then show ?thesis by auto
qed
lemma same_cardinality_post_swap_2:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"trace init t final"
shows
"card {k \<in> {i..<j+1}. prerecording_event t k}
= card {k \<in> {i..<j+1}. prerecording_event (swap_events i j t) k}"
proof -
let ?t = "swap_events i j t"
have "card {k \<in> {i..<j+1}. prerecording_event t k} = 1"
proof -
have "\<forall>k. i \<le> k \<and> k < j \<longrightarrow> ~ prerecording_event t k"
proof (rule allI, rule impI)
fix k
assume asm: "i \<le> k \<and> k < j"
then show "~ prerecording_event t k"
proof (cases "k = i")
case True
then have "postrecording_event t k" using assms by simp
then show ?thesis
by (meson computation.postrecording_event computation.prerecording_event computation_axioms)
next
case False
then have "i < k \<and> k < j" using asm by force
then have "~ regular_event (t ! k)" using assms by simp
then show ?thesis unfolding prerecording_event by simp
qed
qed
then have "{k \<in> {i..<j}. prerecording_event t k} = empty" by simp
moreover have "{k \<in> {j..<j+1}. prerecording_event t k} = {j}"
proof -
have "{j..<j+1} = {j}" by auto
moreover have "prerecording_event t j" using assms by simp
ultimately show ?thesis by blast
qed
ultimately have "{k \<in> {i..<j+1}. prerecording_event t k} = {j}" using assms(3-4) by auto
then show ?thesis by simp
qed
moreover have "card {k \<in> {i..<j+1}. prerecording_event ?t k} = 1"
proof -
have swap_ind: "prerecording_event ?t i
\<and> postrecording_event ?t (i+1)
\<and> (\<forall>k. k > i+1 \<and> k < j+1 \<longrightarrow> ~ regular_event (?t ! k))"
using assms swap_events by blast
have "\<forall>k. i+1 \<le> k \<and> k < j+1 \<longrightarrow> ~ prerecording_event ?t k"
proof (rule allI, rule impI)
fix k
assume asm: "i+1 \<le> k \<and> k < j+1"
then show "~ prerecording_event ?t k"
proof (cases "k = i+1")
case True
then have "postrecording_event ?t k" using swap_ind by blast
then show ?thesis
by (meson computation.postrecording_event computation.prerecording_event computation_axioms)
next
case False
then have "i+1 < k \<and> k < j+1" using asm by linarith
then have "~ regular_event (?t ! k)" using asm assms swap_ind by blast
then show ?thesis unfolding prerecording_event by simp
qed
qed
then have "{k \<in> {i+1..<j+1}. prerecording_event ?t k} = empty" by simp
moreover have "{k \<in> {i..<i+1}. prerecording_event ?t k} = {i}"
proof -
have "{i..<i+1} = {i}" by simp
moreover have "prerecording_event ?t i" using swap_ind by blast
ultimately show ?thesis by blast
qed
ultimately have "{k \<in> {i..<j+1}. prerecording_event ?t k} = {i}" using assms(3-4) by auto
then show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
lemma same_cardinality_post_swap_3:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"trace init t final"
shows
"{k \<in> {j+1..<length t}. prerecording_event t k}
= {k \<in> {j+1..<length (swap_events i j t)}. prerecording_event (swap_events i j t) k}"
proof -
let ?t = "swap_events i j t"
have "drop (j+1) t = drop (j+1) ?t"
using assms(3) assms(4) swap_identical_tails by blast
have same_length: "length t = length ?t" using swap_identical_length assms by blast
have a: "\<forall>k. j+1 \<le> k \<and> k < length t \<longrightarrow> ?t ! k = t ! k"
proof (rule allI, rule impI)
fix k
assume "j+1 \<le> k \<and> k < length t"
then have "?t ! k = drop (j+1) (swap_events i j t) ! (k-(j+1))"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_leI assms(4) le_add_diff_inverse nth_drop same_length)
moreover have "t ! k = drop (j+1) t ! (k-(j+1))"
using \<open>j + 1 \<le> k \<and> k < length t\<close> by auto
ultimately have "drop (j+1) ?t ! (k-(j+1)) = drop (j+1) t ! (k-(j+1))"
using assms swap_identical_tails by metis
then show "?t ! k = t ! k"
using \<open>?t ! k = drop (j + 1) ?t ! (k - (j + 1))\<close> \<open>t ! k = drop (j + 1) t ! (k - (j + 1))\<close> by auto
qed
then have "\<forall>k. j+1 \<le> k \<and> k < length t \<longrightarrow> prerecording_event t k = prerecording_event ?t k"
proof -
have "\<forall>k. k \<ge> (j+1) \<longrightarrow> S t k = S ?t k" using assms swap_events by simp
then show ?thesis unfolding prerecording_event using a by auto
qed
then have "{k \<in> {j+1..<length t}. prerecording_event t k}
= {k \<in> {j+1..<length t}. prerecording_event ?t k}"
by auto
then show ?thesis using same_length by metis
qed
lemma card_ip1_to_j_is_1_in_normal_events:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"count_violations t = Suc n" and
"trace init t final"
shows
"card {k \<in> {i+1..<j+1}. prerecording_event t k} = 1"
proof -
have "\<forall>k. i < k \<and> k < j \<longrightarrow> ~ prerecording_event t k"
proof (rule allI, rule impI)
fix k
assume asm: "i < k \<and> k < j"
then show "~ prerecording_event t k"
proof -
have "~ regular_event (t ! k)" using asm assms by blast
then show ?thesis unfolding prerecording_event by simp
qed
qed
then have "{k \<in> {i+1..<j}. prerecording_event t k} = empty" by auto
moreover have "{k \<in> {j..<j+1}. prerecording_event t k} = {j}"
proof -
have "{j..<j+1} = {j}" by auto
moreover have "prerecording_event t j" using assms by simp
then show ?thesis by auto
qed
ultimately have "{k \<in> {i+1..<j+1}. prerecording_event t k} = {j}" using assms by auto
then show ?thesis by simp
qed
lemma card_ip1_to_j_is_0_in_swapped_events:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"count_violations t = Suc n" and
"trace init t final"
shows
"card {k \<in> {i+1..<j+1}. prerecording_event (swap_events i j t) k} = 0"
proof -
let ?t = "swap_events i j t"
have postrec_ip1: "postrecording_event ?t (i+1)" using assms swap_events by blast
have neigh_shift: "\<forall>k. i+1 < k \<and> k < j+1 \<longrightarrow> ~ regular_event (?t ! k)" using assms swap_events by blast
have "\<forall>k. i+1 \<le> k \<and> k < j+1 \<longrightarrow> ~ prerecording_event ?t k"
proof (rule allI, rule impI)
fix k
assume asm: "i+1 \<le> k \<and> k < j+1"
then show "~ prerecording_event ?t k"
proof (cases "k = i+1")
case True
then show ?thesis using postrec_ip1
by (meson computation.postrecording_event computation.prerecording_event computation_axioms)
next
case False
then have "i+1 < k \<and> k < j+1" using asm by simp
then have "~ regular_event (?t ! k)" using neigh_shift by blast
then show ?thesis unfolding prerecording_event by simp
qed
qed
then have "{k \<in> {i+1..<j+1}. prerecording_event ?t k} = empty" by auto
then show ?thesis by simp
qed
lemma count_violations_swap:
assumes
"prerecording_event t j" and
"postrecording_event t i" and
"i < j" and
"j < length t" and
"count_violations t = Suc n" and
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)" and
"count_violations t = Suc n" and
"trace init t final"
shows
"count_violations (swap_events i j t) = n"
proof -
let ?t = "swap_events i j t"
let ?f = "(\<lambda>i. if postrecording_event t i then card {j \<in> {i+1..<length t}. prerecording_event t j} else 0)"
let ?f' = "(\<lambda>i. if postrecording_event ?t i then card {j \<in> {i+1..<length ?t}. prerecording_event ?t j} else 0)"
have same_postrec_prefix: "\<forall>k. k < i \<longrightarrow> postrecording_event t k = postrecording_event ?t k"
proof -
have "\<forall>k. k < i \<longrightarrow> S t k = S ?t k" using assms swap_events by auto
then show ?thesis unfolding postrecording_event
proof -
assume a1: "\<forall>k<i. S t k = S (swap_events i j t) k"
{ fix nn :: nat
have "\<And>n na es nb. \<not> n < na \<or> \<not> na < length es \<or> \<not> nb < n \<or> swap_events n na es ! nb = (es ! nb::('a, 'b, 'c) event)"
by (metis (no_types) nth_take swap_identical_heads)
then have "\<not> nn < i \<or> \<not> nn < length t \<and> \<not> nn < length (swap_events i j t) \<or> \<not> regular_event (t ! nn) \<and> \<not> regular_event (swap_events i j t ! nn) \<or> ps (S t nn) (occurs_on (t ! nn)) = None \<and> ps (S (swap_events i j t) nn) (occurs_on (swap_events i j t ! nn)) = None \<or> regular_event (t ! nn) \<and> regular_event (swap_events i j t ! nn) \<and> nn < length t \<and> nn < length (swap_events i j t) \<and> ps (S t nn) (occurs_on (t ! nn)) \<noteq> None \<and> ps (S (swap_events i j t) nn) (occurs_on (swap_events i j t ! nn)) \<noteq> None"
using a1 by (metis (no_types) assms(3) assms(4) swap_identical_length) }
then show "\<forall>n<i. (n < length t \<and> regular_event (t ! n) \<and> ps (S t n) (occurs_on (t ! n)) \<noteq> None) = (n < length (swap_events i j t) \<and> regular_event (swap_events i j t ! n) \<and> ps (S (swap_events i j t) n) (occurs_on (swap_events i j t ! n)) \<noteq> None)"
by (metis (no_types))
qed
qed
have same_postrec_suffix: "\<forall>k. k \<ge> j+1 \<longrightarrow> postrecording_event t k = postrecording_event ?t k"
proof -
have post_equal_states: "\<forall>k. k \<ge> j+1 \<longrightarrow> S t k = S ?t k" using assms swap_events by presburger
show ?thesis
proof (rule allI, rule impI)
fix k
assume "j+1 \<le> k"
then show "postrecording_event t k = postrecording_event ?t k"
proof (cases "k < length t")
case False
then have "~ postrecording_event t k" using postrecording_event by simp
moreover have "~ postrecording_event ?t k"
using postrecording_event swap_identical_length False assms by metis
ultimately show ?thesis by simp
next
case True
then show "postrecording_event t k = postrecording_event ?t k"
using post_equal_states
proof -
assume a1: "\<forall>k\<ge>j + 1. S t k = S (swap_events i j t) k"
assume a2: "k < length t"
have f3: "length t = length (swap_events i j t)"
using assms(3) assms(4) swap_identical_length by blast
have f4: "k - (j + 1) + (j + 1) = k"
using \<open>j + 1 \<le> k\<close> le_add_diff_inverse2 by blast
have "drop (j + 1) t = drop (j + 1) (swap_events i j t)"
using assms(3) assms(4) swap_identical_tails by blast
then have "swap_events i j t ! k = t ! k"
using f4 f3 a2 by (metis (no_types) drop_drop hd_drop_conv_nth)
then show ?thesis
using f3 a1 \<open>j + 1 \<le> k\<close> postrecording_event by presburger
qed
qed
qed
qed
have sum_decomp_f: "sum ?f {0..<length t} = sum ?f {0..<i} + sum ?f {i..<j+1} + sum ?f {j+1..<length t}"
by (metis (no_types, lifting) One_nat_def add.right_neutral add_Suc_right assms(3) assms(4) discrete le_add2 less_SucI less_or_eq_imp_le sum.atLeastLessThan_concat)
have sum_decomp_f': "sum ?f' {0..<length t} = sum ?f' {0..<i} + sum ?f' {i..<j+1} + sum ?f' {j+1..<length t}"
by (metis (no_types, lifting) One_nat_def add.right_neutral add_Suc_right assms(3) assms(4) discrete le_add2 less_SucI less_or_eq_imp_le sum.atLeastLessThan_concat)
have prefix_sum: "sum ?f {0..<i} = sum ?f' {0..<i}"
proof -
have "\<forall>l. 0 \<le> l \<and> l < i \<longrightarrow> ?f l = ?f' l"
proof (rule allI, rule impI)
fix l
assume "0 \<le> l \<and> l < i"
then have "l < i" by simp
show "?f l = ?f' l"
proof (cases "postrecording_event t l")
case True
let ?G = "{k \<in> {l+1..<length t}. prerecording_event t k}"
let ?G' = "{k \<in> {l+1..<length t}. prerecording_event ?t k}"
let ?A = "{k \<in> {l+1..<i}. prerecording_event t k}"
let ?B = "{k \<in> {i..<j+1}. prerecording_event t k}"
let ?C = "{k \<in> {j+1..<length t}. prerecording_event t k}"
let ?A' = "{k \<in> {l+1..<i}. prerecording_event ?t k}"
let ?B' = "{k \<in> {i..<j+1}. prerecording_event ?t k}"
let ?C' = "{k \<in> {j+1..<length t}. prerecording_event ?t k}"
have card_G: "card ?G = card ?A + card ?B + card ?C"
proof -
- have "?G = ?A \<union> (?B \<union> ?C)" using assms `l < i` by auto
+ have "?G = ?A \<union> (?B \<union> ?C)" using assms \<open>l < i\<close> by auto
then have "card ?G = card (?A \<union> (?B \<union> ?C))" by simp
also have "card (?A \<union> (?B \<union> ?C)) = card ?A + card (?B \<union> ?C)"
proof -
- have "?A \<inter> (?B \<union> ?C) = {}" using `l < i` assms by auto
+ have "?A \<inter> (?B \<union> ?C) = {}" using \<open>l < i\<close> assms by auto
then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
also have "card ?A + card (?B \<union> ?C) = card ?A + card ?B + card ?C"
proof -
have "?B \<inter> ?C = {}" by auto
then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
finally show ?thesis by simp
qed
have card_G': "card ?G' = card ?A' + card ?B' + card ?C'"
proof -
- have "?G' = ?A' \<union> (?B' \<union> ?C')" using assms `l < i` by auto
+ have "?G' = ?A' \<union> (?B' \<union> ?C')" using assms \<open>l < i\<close> by auto
then have "card ?G' = card (?A' \<union> (?B' \<union> ?C'))" by simp
also have "card (?A' \<union> (?B' \<union> ?C')) = card ?A' + card (?B' \<union> ?C')"
proof -
- have "?A' \<inter> (?B' \<union> ?C') = {}" using `l < i` assms by auto
+ have "?A' \<inter> (?B' \<union> ?C') = {}" using \<open>l < i\<close> assms by auto
then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
also have "card ?A' + card (?B' \<union> ?C') = card ?A' + card ?B' + card ?C'"
proof -
have "?B' \<inter> ?C' = {}" by auto
then show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
finally show ?thesis by simp
qed
have "card ?G = card ?G'"
proof -
have "card ?A = card ?A'"
proof -
have "{k \<in> {0..<i}. prerecording_event t k} = {k \<in> {0..<i}. prerecording_event ?t k}"
using assms same_cardinality_post_swap_1 by blast
then have "?A = ?A'" by auto
then show ?thesis by simp
qed
moreover have "card ?B = card ?B'" using assms same_cardinality_post_swap_2 by blast
moreover have "card ?C = card ?C'"
proof -
have "?C = ?C'" using assms same_cardinality_post_swap_3 by auto
then show ?thesis by simp
qed
ultimately show ?thesis using card_G card_G' by linarith
qed
- moreover have "postrecording_event ?t l" using True same_postrec_prefix `l < i` by blast
+ moreover have "postrecording_event ?t l" using True same_postrec_prefix \<open>l < i\<close> by blast
moreover have "length ?t = length t" using assms(3) assms(4) by fastforce
ultimately show ?thesis using True by presburger
next
case False
- then have "~ postrecording_event ?t l" using same_postrec_prefix `l < i` by blast
+ then have "~ postrecording_event ?t l" using same_postrec_prefix \<open>l < i\<close> by blast
then show ?thesis using False by simp
qed
qed
then show ?thesis using sum_eq_if_same_subterms by auto
qed
have infix_sum: "sum ?f {i..<j+1} = sum ?f' {i..<j+1} + 1"
proof -
have sum_decomp_f: "sum ?f {i..<j+1} = sum ?f {i..<i+2} + sum ?f {i+2..<j+1}"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_leI add_2_eq_Suc' assms(3) less_Suc_eq linorder_not_le sum.atLeastLessThan_concat)
have sum_decomp_f': "sum ?f' {i..<j+1} = sum ?f' {i..<i+2} + sum ?f' {i+2..<j+1}"
by (metis (no_types, lifting) Suc_eq_plus1 Suc_mono add.assoc assms(3) discrete le_add1 one_add_one sum.atLeastLessThan_concat)
have "sum ?f {i+2..<j+1} = sum ?f' {i+2..<j+1}"
proof -
have "\<forall>l. i+2 \<le> l \<and> l < j+1 \<longrightarrow> ?f l = ?f' l"
proof (rule allI, rule impI)
fix l
assume asm: "i+2 \<le> l \<and> l < j+1"
have "?f l = 0"
proof (cases "l = j")
case True
then have "~ postrecording_event t l"
using assms(1) postrecording_event prerecording_event by auto
then show ?thesis by simp
next
case False
then have "i < l \<and> l < j" using assms asm by simp
then have "~ regular_event (t ! l)" using assms by blast
then have "~ postrecording_event t l" unfolding postrecording_event by simp
then show ?thesis by simp
qed
moreover have "?f' l = 0"
proof -
have "\<forall>k. i+1 < k \<and> k < j+1 \<longrightarrow> ~ regular_event (?t ! k)" using assms swap_events by blast
then have "~ regular_event (?t ! l)" using asm by simp
then have "~ postrecording_event ?t l" unfolding postrecording_event by simp
then show ?thesis by simp
qed
ultimately show "?f l = ?f' l" by simp
qed
then show ?thesis using sum_eq_if_same_subterms by simp
qed
moreover have "sum ?f {i..<i+2} = 1 + sum ?f' {i..<i+2}"
proof -
have int_def: "{i..<i+2} = {i,i+1}" by auto
then have "sum ?f {i,i+1} = ?f i + ?f (i+1)" by simp
moreover have "sum ?f' {i,i+1} = ?f' i + ?f' (i+1)" using int_def by simp
moreover have "?f (i+1) = 0"
proof (cases "j = i+1")
case True
then have "prerecording_event t (i+1)" using assms by simp
then have "~ postrecording_event t (i+1)"
unfolding postrecording_event using prerecording_event by simp
then show ?thesis by simp
next
case False
then have "~ regular_event (t ! (i+1))" using assms by simp
then have "~ postrecording_event t (i+1)" unfolding postrecording_event by simp
then show ?thesis by simp
qed
moreover have "?f' i = 0"
proof -
have "prerecording_event ?t i" using assms swap_events by blast
then have "~ postrecording_event ?t i"
unfolding postrecording_event using prerecording_event by simp
then show ?thesis by simp
qed
moreover have "?f i = ?f' (i+1) + 1"
proof -
have pi: "postrecording_event t i" using assms by simp
moreover have pip1: "postrecording_event ?t (i+1)" using assms swap_events by blast
let ?G = "{k \<in> {i+1..<length t}. prerecording_event t k}"
let ?G' = "{k \<in> {i+2..<length ?t}. prerecording_event ?t k}"
let ?A = "{k \<in> {i+1..<j+1}. prerecording_event t k}"
let ?B = "{k \<in> {j+1..<length t}. prerecording_event t k}"
let ?A' = "{k \<in> {i+2..<j+1}. prerecording_event ?t k}"
let ?B' = "{k \<in> {j+1..<length ?t}. prerecording_event ?t k}"
have card_G: "card ?G = card ?A + card ?B"
proof -
have "?G = ?A \<union> ?B" using assms by auto
moreover have "?A \<inter> ?B = {}" by auto
ultimately show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
have card_G': "card ?G' = card ?A' + card ?B'"
proof -
have "?G' = ?A' \<union> ?B'" using assms by auto
moreover have "?A' \<inter> ?B' = {}" by auto
ultimately show ?thesis by (simp add: card_Un_disjoint disjoint_iff_not_equal)
qed
have "card ?G = card ?G' + 1"
proof -
have "card ?A = card ?A' + 1"
proof -
have "card ?A = 1" using assms card_ip1_to_j_is_1_in_normal_events by blast
moreover have "card ?A' = 0" using assms card_ip1_to_j_is_0_in_swapped_events by force
ultimately show ?thesis by algebra
qed
moreover have "card ?B = card ?B'" using assms same_cardinality_post_swap_3 by force
ultimately show ?thesis using card_G card_G' by presburger
qed
moreover have "card ?G = ?f i" using pi by simp
moreover have "card ?G' = ?f' (i+1)" using pip1 by simp
ultimately show ?thesis by argo
qed
ultimately show ?thesis by auto
qed
ultimately show ?thesis using sum_decomp_f sum_decomp_f' by linarith
qed
have suffix_sum: "sum ?f {j+1..<length t} = sum ?f' {j+1..<length t}"
proof -
have "\<forall>l. l > j \<longrightarrow> ?f l = ?f' l"
proof (rule allI, rule impI)
fix l
assume "l > j"
then show "?f l = ?f' l"
proof (cases "postrecording_event t l")
case True
let ?G = "{k \<in> {l+1..<length t}. prerecording_event t k}"
let ?G' = "{k \<in> {l+1..<length t}. prerecording_event ?t k}"
let ?C = "{k \<in> {j+1..<length t}. prerecording_event t k}"
let ?C' = "{k \<in> {j+1..<length t}. prerecording_event ?t k}"
have "card ?G = card ?G'"
proof -
have "?C = ?C'" using assms same_cardinality_post_swap_3 by auto
- then have "?G = ?G'" using `l > j` by fastforce
+ then have "?G = ?G'" using \<open>l > j\<close> by fastforce
then show ?thesis by simp
qed
- moreover have "postrecording_event ?t l" using True same_postrec_suffix `l > j` by simp
+ moreover have "postrecording_event ?t l" using True same_postrec_suffix \<open>l > j\<close> by simp
moreover have "length ?t = length t" using assms(3) assms(4) by fastforce
ultimately show ?thesis using True by presburger
next
case False
- then have "~ postrecording_event ?t l" using same_postrec_suffix `l > j` by simp
+ then have "~ postrecording_event ?t l" using same_postrec_suffix \<open>l > j\<close> by simp
then show ?thesis using False by simp
qed
qed
then have "\<forall>k. j+1 \<le> k \<and> k < length t \<longrightarrow> ?f k = ?f' k"
using discrete by blast
moreover have "length t = length ?t"
using assms(3) assms(4) swap_identical_length by blast
ultimately show ?thesis by (blast intro:sum_eq_if_same_subterms)
qed
have "sum ?f {0..<length t} = sum ?f' {0..<length t} + 1"
proof -
have "sum ?f {0..<i} = sum ?f' {0..<i}" using prefix_sum by simp
moreover have "sum ?f {i..<j+1} = sum ?f' {i..<j+1} + 1" using infix_sum by simp
moreover have "sum ?f {j+1..<length t} = sum ?f' {j+1..<length t}" using suffix_sum by simp
ultimately show ?thesis using sum_decomp_f sum_decomp_f' by presburger
qed
moreover have "length ?t = length t"
using assms(3) assms(4) by auto
moreover have "sum ?f {0..<length t} = n + 1" using assms count_violations_def by simp
ultimately have "sum ?f' {0..<length ?t} = n" by presburger
then show ?thesis unfolding count_violations_def by presburger
qed
lemma desired_trace_always_exists:
assumes
"trace init t final"
shows
"\<exists>t'. perm t' t
\<and> all_prerecording_before_postrecording t'"
using assms proof (induct "count_violations t" arbitrary: t)
case 0
then show ?case
proof (cases "\<exists>i. prerecording_event t i")
case False
then have "\<forall>j. ~ prerecording_event t j" by auto
then have "\<forall>j. j \<le> 0 \<longrightarrow> ~ postrecording_event t j"
using "0.prems" init_is_s_t_0 no_initial_process_snapshot postrecording_event by auto
moreover have "\<forall>j. j > 0 \<longrightarrow> ~ prerecording_event t j" using False by auto
moreover have "length t > 0"
by (metis "0.prems" all_processes_snapshotted_in_final_state length_greater_0_conv no_initial_process_snapshot tr_init trace_and_start_determines_end)
ultimately show ?thesis using "0.prems" False by auto
next
case True
let ?Is = "{i. prerecording_event t i}"
have "?Is \<noteq> empty"
by (simp add: True)
moreover have fin_Is: "finite ?Is"
proof (rule ccontr)
assume "~ finite ?Is"
then obtain i where "i > length t" "prerecording_event t i"
by (simp add: prerecording_event)
then show False using prerecording_event by auto
qed
let ?i = "Max ?Is"
have pi: "prerecording_event t ?i"
using Max_in calculation fin_Is by blast
have "?i < length t"
proof (rule ccontr)
assume "~ ?i < length t"
then show False
using calculation fin_Is computation.prerecording_event computation_axioms by force
qed
moreover have "\<forall>j. j \<ge> ?i+1 \<longrightarrow> ~ prerecording_event t j"
proof -
have "\<forall>j. j > ?i \<longrightarrow> ~ prerecording_event t j"
using Max_less_iff fin_Is by auto
then show ?thesis by auto
qed
moreover have "\<forall>j. j < ?i+1 \<longrightarrow> ~ postrecording_event t j"
proof -
have "\<forall>j. j \<le> ?i \<longrightarrow> ~ postrecording_event t j"
proof (rule allI, rule impI, rule ccontr)
fix j
assume "j \<le> ?i" "~ ~ postrecording_event t j"
then have "j < ?i"
by (metis add_diff_inverse_nat dual_order.antisym le_add1 pi postrecording_event prerecording_event)
then have "count_violations t > 0"
proof -
have "(if postrecording_event t j
then card {l \<in> {j+1..<length t}. prerecording_event t l}
else 0) = card {l \<in> {j+1..<length t}. prerecording_event t l}"
- using `~ ~ postrecording_event t j` by simp
+ using \<open>~ ~ postrecording_event t j\<close> by simp
moreover have "card {l \<in> {j+1..<length t}. prerecording_event t l} > 0"
proof -
have "j + 1 \<le> ?i \<and> ?i < length t"
using \<open>Max {i. prerecording_event t i} < length t\<close> \<open>j < Max {i. prerecording_event t i}\<close> discrete by blast
moreover have "prerecording_event t ?i" using pi by simp
ultimately have "{l \<in> {j+1..<length t}. prerecording_event t l} \<noteq> empty" by fastforce
then show ?thesis by fastforce
qed
ultimately show ?thesis
by (metis (no_types, lifting) violations_0_implies_all_subterms_0 \<open>Max {i. prerecording_event t i} < length t\<close> \<open>j < Max {i. prerecording_event t i}\<close> atLeastLessThan_iff less_trans linorder_not_le neq0_conv)
qed
then show False using "0" by simp
qed
then show ?thesis by auto
qed
moreover have "?i+1 \<le> length t"
using calculation(2) discrete by blast
ultimately show ?thesis using "0.prems" by blast
qed
next
case (Suc n)
then obtain i j where ind: "postrecording_event t i" "prerecording_event t j"
"\<forall>k. (i < k \<and> k < j) \<longrightarrow> ~ regular_event (t ! k)"
"i < j" "j < length t" using exists_neighboring_violation_pair Suc by force
then have "trace init (swap_events i j t) final
\<and> (\<forall>k. k \<ge> j + 1 \<longrightarrow> S (swap_events i j t) k = S t k)
\<and> (\<forall>k. k \<le> i \<longrightarrow> S (swap_events i j t) k = S t k)"
using Suc swap_events by presburger
moreover have "perm (swap_events i j t) t" using swap_events_perm ind by blast
moreover have "count_violations (swap_events i j t) = n"
using count_violations_swap Suc ind by simp
ultimately show ?case using Suc.hyps by blast
qed
theorem snapshot_algorithm_is_correct:
assumes
"trace init t final"
shows
"\<exists>t' i. trace init t' final \<and> perm t' t
\<and> state_equal_to_snapshot (S t' i) final \<and> i \<le> length t'"
proof -
obtain t' where "perm t' t" and
"all_prerecording_before_postrecording t'"
using assms desired_trace_always_exists by blast
then show ?thesis using snapshot_after_all_prerecording_events
by blast
qed
subsection \<open>Stable property detection\<close>
text \<open>Finally, we show that the computed snapshot is indeed
suitable for stable property detection, as claimed in~\cite{chandy}.\<close>
definition stable where
"stable p \<equiv> (\<forall>c. p c \<longrightarrow> (\<forall>t c'. trace c t c' \<longrightarrow> p c'))"
lemma has_snapshot_stable:
assumes
"trace init t final"
shows
"stable (\<lambda>c. has_snapshotted c p)"
using snapshot_stable stable_def by auto
definition some_snapshot_state where
"some_snapshot_state t \<equiv>
SOME (t', i). trace init t final
\<and> trace init t' final \<and> perm t' t
\<and> state_equal_to_snapshot (S t' i) final"
lemma split_S:
assumes
"trace init t final"
shows
"trace (S t i) (drop i t) final"
proof -
have "t = take i t @ drop i t" by simp
then show ?thesis
by (metis split_trace assms exists_trace_for_any_i
trace_and_start_determines_end)
qed
theorem Stable_Property_Detection:
assumes
"stable p" and
"trace init t final" and
"(t', i) = some_snapshot_state t" and
"p (S t' i)"
shows
"p final"
proof -
have "\<exists>t' i. trace init t final
\<and> trace init t' final \<and> perm t' t
\<and> state_equal_to_snapshot (S t' i) final"
using snapshot_algorithm_is_correct assms(2) by blast
then have "trace init t' final"
using assms
unfolding some_snapshot_state_def
by (metis (lifting) case_prodD case_prodI someI_ex)
then show ?thesis
using assms stable_def split_S by metis
qed
end (* locale computation *)
end (* theory Snapshot *)
diff --git a/thys/Chandy_Lamport/Swap.thy b/thys/Chandy_Lamport/Swap.thy
--- a/thys/Chandy_Lamport/Swap.thy
+++ b/thys/Chandy_Lamport/Swap.thy
@@ -1,1553 +1,1553 @@
section \<open>Swap lemmas\<close>
text \<open>\<close>
theory Swap
imports
Distributed_System
begin
context distributed_system
begin
lemma swap_msgs_Trans_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain u u' where "ev = Trans ?p u u'"
by (metis assms(3) event.collapse(1))
obtain u'' u''' where "ev' = Trans ?q u'' u'''"
by (metis assms(4) event.collapse(1))
then have "msgs d' i = msgs d i"
by (metis Trans_msg assms(1) assms(3) assms(4) assms(5))
then have "msgs e i = msgs e' i"
using Trans_msg assms(2) assms(3) assms(4) assms(6) by auto
then show ?thesis by blast
qed
lemma swap_msgs_Send_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send_ev: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
obtain i'' s u'' u''' m' where Send_ev': "ev' = Send i'' ?q s u'' u''' m'"
by (metis assms(4) event.collapse(2))
have "i' \<noteq> i''"
by (metis (mono_tags, lifting) \<open>ev = Send i' (occurs_on ev) r u u' m\<close> \<open>ev' = Send i'' (occurs_on ev') s u'' u''' m'\<close> assms(1) assms(2) assms(7) can_occur_def event.simps(27) happen_implies_can_occur option.simps(1) prod.simps(1))
then show ?thesis
proof (cases "i = i' \<or> i = i''")
case True
then show ?thesis
proof (elim disjE)
assume "i = i'"
then have "msgs d i = msgs c i @ [Msg m]"
by (metis \<open>ev = Send i' (occurs_on ev) r u u' m\<close> assms(1) next_send)
moreover have "msgs e i = msgs d i"
by (metis \<open>ev' = Send i'' (occurs_on ev') s u'' u''' m'\<close> \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(2) assms(4) event.sel(8) msgs_unchanged_for_other_is regular_event)
moreover have "msgs d' i = msgs c i"
by (metis \<open>ev' = Send i'' (occurs_on ev') s u'' u''' m'\<close> \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(4) assms(5) event.sel(8) msgs_unchanged_for_other_is regular_event)
moreover have "msgs e' i = msgs d' i @ [Msg m]"
by (metis \<open>ev = Send i' (occurs_on ev) r u u' m\<close> \<open>i = i'\<close> assms(6) next_send)
ultimately show ?thesis by simp
next
assume "i = i''"
then have "msgs d i = msgs c i"
by (metis Send_ev \<open>i' \<noteq> i''\<close> assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
moreover have "msgs e i = msgs d i @ [Msg m']"
by (metis Send_ev' \<open>i = i''\<close> assms(2) next_send)
moreover have "msgs d' i = msgs c i @ [Msg m']"
by (metis Send_ev' \<open>i = i''\<close> assms(5) next_send)
moreover have "msgs e' i = msgs d' i"
by (metis Send_ev \<open>i = i''\<close> \<open>i' \<noteq> i''\<close> assms(3) assms(6) event.sel(8) msgs_unchanged_for_other_is regular_event)
ultimately show ?thesis by simp
qed
next
case False
then have "msgs e i = msgs d i" using Send_ev' assms
by (metis event.sel(8) msgs_unchanged_for_other_is regular_event)
also have "... = msgs c i"
by (metis False Send_ev assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
also have "... = msgs d' i"
by (metis (no_types, hide_lams) \<open>msgs d i = msgs c i\<close> assms(2) assms(4) assms(5) calculation regular_event same_messages_imply_same_resulting_messages)
also have "... = msgs e' i"
by (metis (no_types, hide_lams) \<open>msgs c i = msgs d' i\<close> \<open>msgs d i = msgs c i\<close> assms(1) assms(3) assms(6) regular_event same_messages_imply_same_resulting_messages)
finally show ?thesis by simp
qed
qed
lemma swap_msgs_Recv_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Recv_ev: "ev = Recv i' ?p r u u' m"
by (metis assms(3) event.collapse(3))
obtain i'' s u'' u''' m' where Recv_ev': "ev' = Recv i'' ?q s u'' u''' m'"
by (metis assms(4) event.collapse(3))
have "i' \<noteq> i''"
by (metis Recv_ev Recv_ev' assms(1) assms(2) assms(7) can_occur_Recv happen_implies_can_occur option.simps(1) prod.simps(1))
show ?thesis
proof (cases "i = i' \<or> i = i''")
case True
then show ?thesis
proof (elim disjE)
assume "i = i'"
then have "Msg m # msgs d i = msgs c i" using Recv_ev assms by (metis next_recv)
moreover have "msgs e i = msgs d i"
by (metis Recv_ev' \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(2) assms(4) event.sel(9) msgs_unchanged_for_other_is regular_event)
moreover have "msgs d' i = msgs c i"
by (metis Recv_ev' \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(4) assms(5) event.sel(9) msgs_unchanged_for_other_is regular_event)
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv_ev \<open>i = i'\<close> assms(6) next_recv)
ultimately show ?thesis by (metis list.inject)
next
assume "i = i''"
then have "msgs d i = msgs c i"
by (metis Recv_ev \<open>i' \<noteq> i''\<close> assms(1) assms(3) event.sel(9) msgs_unchanged_for_other_is regular_event)
moreover have "Msg m' # msgs e i = msgs d i"
by (metis Recv_ev' \<open>i = i''\<close> assms(2) next_recv)
moreover have "Msg m' # msgs d' i = msgs c i"
by (metis Recv_ev' \<open>i = i''\<close> assms(5) next_recv)
moreover have "msgs e' i = msgs d' i"
by (metis Recv_ev \<open>i = i''\<close> \<open>i' \<noteq> i''\<close> assms(3) assms(6) event.sel(9) msgs_unchanged_for_other_is regular_event)
ultimately show ?thesis by (metis list.inject)
qed
next
case False
then have "msgs e i = msgs d i"
by (metis Recv_ev' assms(2) assms(4) event.sel(9) msgs_unchanged_for_other_is regular_event)
also have "... = msgs c i"
by (metis False Recv_ev assms(1) assms(3) event.sel(9) msgs_unchanged_for_other_is regular_event)
also have "... = msgs d' i"
by (metis (no_types, hide_lams) \<open>msgs d i = msgs c i\<close> assms(2) assms(4) assms(5) calculation regular_event same_messages_imply_same_resulting_messages)
also have "... = msgs e' i"
by (metis (no_types, lifting) \<open>msgs c i = msgs d' i\<close> \<open>msgs d i = msgs c i\<close> assms(1) assms(3) assms(6) regular_event same_messages_imply_same_resulting_messages)
finally show ?thesis by simp
qed
qed
lemma swap_msgs_Send_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
obtain u'' u''' where Trans: "ev' = Trans ?q u'' u'''"
by (metis assms(4) event.collapse(1))
show ?thesis
proof (cases "i = i'")
case True
then have "msgs d i = msgs c i @ [Msg m]"
by (metis Send assms(1) next_send)
moreover have "msgs e i = msgs d i"
using Trans_msg assms(2) assms(4) by auto
moreover have "msgs d' i = msgs c i"
using Trans_msg assms(4) assms(5) by auto
moreover have "msgs e' i = msgs d' i @ [Msg m]"
by (metis Send True assms(6) next_send)
ultimately show ?thesis by simp
next
case False
then have "msgs e i = msgs d i"
using Trans_msg assms(2) assms(4) by auto
also have "... = msgs c i"
by (metis False Send assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
also have "... = msgs d' i"
using Trans_msg assms(4) assms(5) by blast
also have "... = msgs e' i"
by (metis (no_types, lifting) \<open>msgs c i = msgs d' i\<close> \<open>msgs d i = msgs c i\<close> assms(1) assms(3) assms(6) regular_event same_messages_imply_same_resulting_messages)
finally show ?thesis by simp
qed
qed
lemma swap_msgs_Trans_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_msgs_Send_Trans by auto
lemma swap_msgs_Recv_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Recv: "ev = Recv i' ?p r u u' m"
by (metis assms(3) event.collapse(3))
obtain u'' u''' where Trans: "ev' = Trans ?q u'' u'''"
by (metis assms(4) event.collapse(1))
show ?thesis
proof (cases "i = i'")
case True
then have "Msg m # msgs d i = msgs c i"
by (metis Recv assms(1) next_recv)
moreover have "msgs e i = msgs d i"
using Trans_msg assms(2) assms(4) by auto
moreover have "msgs d' i = msgs c i"
using Trans_msg assms(4) assms(5) by auto
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv True assms(6) next_recv)
ultimately show ?thesis by (metis list.inject)
next
case False
then have "msgs e i = msgs d i"
using Trans_msg assms(2) assms(4) by auto
also have "... = msgs c i"
by (metis False Recv assms(1) assms(3) event.sel(9) msgs_unchanged_for_other_is regular_event)
also have "... = msgs d' i"
using Trans_msg assms(4) assms(5) by blast
also have "... = msgs e' i"
by (metis False Recv assms(6) next_recv)
finally show ?thesis by simp
qed
qed
lemma swap_msgs_Trans_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_msgs_Recv_Trans by auto
lemma swap_msgs_Send_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
obtain i'' s u'' u''' m' where Recv: "ev' = Recv i'' ?q s u'' u''' m'"
by (metis assms(4) event.collapse(3))
show ?thesis
proof (cases "i = i'"; cases "i = i''", goal_cases)
case 1
then have "msgs e' i = msgs d' i @ [Msg m]"
by (metis Send assms(6) next_send)
moreover have "Msg m' # msgs d' i = msgs c i"
by (metis "1"(2) Recv assms(5) next_recv)
moreover have "msgs d i = msgs c i @ [Msg m]"
by (metis "1"(1) Send assms(1) next_send)
moreover have "Msg m' # msgs e i = msgs d i"
by (metis "1"(2) Recv assms(2) next_recv)
ultimately show ?thesis
by (metis list.sel(2) list.sel(3) not_Cons_self2 tl_append2)
next
case 2
then have "msgs d i = msgs c i @ [Msg m]"
by (metis Send assms(1) next_send)
moreover have "msgs e i = msgs d i"
by (metis "2"(2) Recv assms(2) assms(4) event.sel(9) msgs_unchanged_for_other_is regular_event)
moreover have "msgs d' i = msgs c i"
by (metis "2"(2) Recv assms(4) assms(5) event.sel(9) msgs_unchanged_for_other_is regular_event)
moreover have "msgs e' i = msgs d' i @ [Msg m]"
by (metis Send 2(1) assms(6) next_send)
ultimately show ?thesis by simp
next
assume 3: "i \<noteq> i'" "i = i''"
then have "msgs d i = msgs c i"
by (metis Send assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
moreover have "Msg m' # msgs e i = msgs d i" using 3 Recv assms by (metis next_recv)
moreover have "Msg m' # msgs d' i = msgs c i"
by (metis "3"(2) Recv assms(5) next_recv)
moreover have "msgs e' i = msgs d' i"
by (metis "3"(1) Send assms(6) next_send)
ultimately show ?thesis by (metis list.inject)
next
assume 4: "i \<noteq> i'" "i \<noteq> i''"
then have "msgs e i = msgs d i"
by (metis Recv assms(2) assms(4) event.sel(9) msgs_unchanged_for_other_is regular_event)
also have "... = msgs c i"
by (metis "4"(1) Send assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
also have "... = msgs d' i"
by (metis "4"(2) Recv assms(5) next_recv)
also have "... = msgs e' i"
by (metis "4"(1) Send assms(6) next_send)
finally show ?thesis by simp
qed
qed
lemma swap_msgs_Recv_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_msgs_Send_Recv by auto
lemma same_cs_implies_same_resulting_cs:
assumes
"cs c i = cs d i"
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'" and
"regular_event ev"
shows
"cs c' i = cs d' i"
proof -
have "isTrans ev \<or> isSend ev \<or> isRecv ev" using assms by simp
then show ?thesis
proof (elim disjE)
assume "isTrans ev"
then show ?thesis
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) event.distinct_disc(4) no_cs_change_if_no_event)
next
assume "isSend ev"
then show ?thesis
by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) event.distinct_disc(10) no_cs_change_if_no_event)
next
assume "isRecv ev"
then obtain i' r s u u' m where Recv: "ev = Recv i' r s u u' m"
by (meson isRecv_def)
then show ?thesis
proof (cases "i' = i")
case True
with assms Recv show ?thesis by (cases "snd (cs c i) = Recording", auto)
next
case False
then show ?thesis using assms Recv by simp
qed
qed
qed
lemma regular_event_implies_same_channel_snapshot_Recv_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"cs e i = cs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Recv_ev: "ev = Recv i' ?p r u u' m"
by (metis assms(3) event.collapse(3))
obtain i'' s u'' u''' m' where Recv_ev': "ev' = Recv i'' ?q s u'' u''' m'"
by (metis assms(4) event.collapse(3))
have "i' \<noteq> i''"
by (metis Recv_ev Recv_ev' assms(1) assms(5) assms(7) can_occur_Recv happen_implies_can_occur option.simps(1) prod.simps(1))
show ?thesis
proof (cases "i = i' \<or> i = i''")
case True
then show ?thesis
proof (elim disjE)
assume "i = i'"
then have "cs d' i = cs c i"
using assms(4) assms(5) assms(7) no_cs_change_if_no_event
by (metis Recv_ev' \<open>i' \<noteq> i''\<close> event.sel(9) regular_event)
then have "cs e' i = cs d i"
using assms(1) assms(3) assms(6) distributed_system.same_cs_implies_same_resulting_cs distributed_system_axioms regular_event by blast
then have "cs d i = cs e i"
by (metis Recv_ev' \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(2) assms(4) event.sel(9) no_cs_change_if_no_event regular_event)
then show ?thesis
by (simp add: \<open>cs e' i = cs d i\<close>)
next
assume "i = i''"
then have "cs d i = cs c i"
by (metis Recv_ev \<open>i' \<noteq> i''\<close> assms(1) assms(3) event.sel(9) no_cs_change_if_no_event regular_event)
then have "cs e i = cs d' i"
using assms(2) assms(4) assms(5) regular_event same_cs_implies_same_resulting_cs by blast
then have "cs d' i = cs e' i"
by (metis Recv_ev \<open>i = i''\<close> \<open>i' \<noteq> i''\<close> assms(3) assms(6) event.sel(9) no_cs_change_if_no_event regular_event)
then show ?thesis
by (simp add: \<open>cs e i = cs d' i\<close>)
qed
next
case False
then show ?thesis
by (metis Recv_ev Recv_ev' assms(1) assms(2) assms(5) assms(6) next_recv)
qed
qed
lemma regular_event_implies_same_channel_snapshot_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"~ isRecv ev" and
"regular_event ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"cs e i = cs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' s u'' u''' m' where Recv: "ev' = Recv i' ?q s u'' u''' m'"
by (metis assms(5) event.collapse(3))
show ?thesis
proof (cases "i = i'")
case True
then have "cs d i = cs c i"
- using assms(1) assms(3) assms(7) no_cs_change_if_no_event `regular_event ev` `~ isRecv ev` by auto
+ using assms(1) assms(3) assms(7) no_cs_change_if_no_event \<open>regular_event ev\<close> \<open>~ isRecv ev\<close> by auto
then have "cs e i = cs d' i"
using assms(2) assms(5) assms(6) regular_event same_cs_implies_same_resulting_cs by blast
then have "cs d' i = cs e' i"
- using True assms(3) assms(6) assms(7) no_cs_change_if_no_event `regular_event ev` `~ isRecv ev` by auto
+ using True assms(3) assms(6) assms(7) no_cs_change_if_no_event \<open>regular_event ev\<close> \<open>~ isRecv ev\<close> by auto
then show ?thesis
by (simp add: \<open>cs e i = cs d' i\<close>)
next
case False
then have "cs d i = cs c i"
using assms(1) assms(3) assms(4) no_cs_change_if_no_event by auto
then have "cs d' i = cs e i"
by (metis (no_types, lifting) assms(2) assms(5) assms(6) regular_event same_cs_implies_same_resulting_cs)
then show "cs e i = cs e' i"
using assms(3) assms(4) assms(7) no_cs_change_if_no_event by auto
qed
qed
lemma same_messages_2:
assumes
"\<forall>p. has_snapshotted c p = has_snapshotted d p" and
"msgs c i = msgs d i" and
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'" and
"~ regular_event ev"
shows
"msgs c' i = msgs d' i"
proof (cases "channel i = None")
case True
then show ?thesis
using assms(2) assms(3) assms(4) no_msgs_change_if_no_channel by auto
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
have "isSnapshot ev \<or> isRecvMarker ev"
using assms(5) event.exhaust_disc by auto
then show ?thesis
proof (elim disjE)
assume "isSnapshot ev"
then obtain r where Snapshot: "ev = Snapshot r" by (meson isSnapshot_def)
then show ?thesis
proof (cases "r = p")
case True
then have "msgs c' i = msgs c i @ [Marker]" using chan Snapshot assms by simp
moreover have "msgs d' i = msgs d i @ [Marker]" using chan Snapshot assms True by simp
ultimately show ?thesis using assms by simp
next
case False
then have "msgs c' i = msgs c i" using chan Snapshot assms by simp
moreover have "msgs d' i = msgs d i" using chan Snapshot assms False by simp
ultimately show ?thesis using assms by simp
qed
next
assume "isRecvMarker ev"
then obtain i' r s where RecvMarker: "ev = RecvMarker i' r s"
by (meson isRecvMarker_def)
then show ?thesis
proof (cases "has_snapshotted c r")
case snap: True
then show ?thesis
proof (cases "i = i'")
case True
then have "Marker # msgs c' i = msgs c i" using chan RecvMarker assms snap by simp
moreover have "Marker # msgs d' i = msgs d i" using chan RecvMarker assms snap True by simp
ultimately show ?thesis using assms by (metis list.inject)
next
case False
then have "msgs d' i = msgs d i"
using RecvMarker assms(1) assms(4) snap by auto
also have "... = msgs c i" using assms by simp
also have "... = msgs c' i"
using False RecvMarker snap assms by auto
finally show ?thesis using snap by simp
qed
next
case no_snap: False
then show ?thesis
proof (cases "i = i'")
case True
then have "Marker # msgs c' i = msgs c i" using chan RecvMarker assms by simp
moreover have "Marker # msgs d' i = msgs d i" using chan RecvMarker assms True by simp
ultimately show ?thesis using assms by (metis list.inject)
next
case not_i: False
then show ?thesis
proof (cases "r = p")
case True
then have "msgs c' i = msgs c i @ [Marker]"
using no_snap RecvMarker assms True chan not_i by auto
moreover have "msgs d' i = msgs d i @ [Marker]"
proof -
have "~ has_snapshotted d r" using assms no_snap True by simp
then show ?thesis
using no_snap RecvMarker assms True chan not_i by auto
qed
ultimately show ?thesis using assms by simp
next
case False
then have "msgs c i = msgs c' i" using False RecvMarker no_snap chan assms not_i by simp
moreover have "msgs d' i = msgs d i"
proof -
have "~ has_snapshotted d r" using assms no_snap False by simp
then show ?thesis
using False RecvMarker no_snap chan assms not_i by simp
qed
ultimately show ?thesis using assms by simp
qed
qed
qed
qed
qed
lemma same_cs_2:
assumes
"\<forall>p. has_snapshotted c p = has_snapshotted d p" and
"cs c i = cs d i" and
"c \<turnstile> ev \<mapsto> c'" and
"d \<turnstile> ev \<mapsto> d'"
shows
"cs c' i = cs d' i"
proof (cases "channel i = None")
case True
then show ?thesis
using assms(2) assms(3) assms(4) no_cs_change_if_no_channel by auto
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
then show ?thesis
proof (cases ev)
case (Snapshot r)
with assms chan show ?thesis by (cases "r = q", auto)
next
case (RecvMarker i' r s)
then show ?thesis
proof (cases "has_snapshotted c r")
case snap: True
then have sdr: "has_snapshotted d r" using assms by auto
then show ?thesis
proof (cases "i = i'")
case True
then have "cs c' i = (fst (cs c i), Done)"
using RecvMarker assms(3) next_recv_marker by blast
also have "... = cs d' i"
using RecvMarker True assms(2) assms(4) by auto
finally show ?thesis using True by simp
next
case False
then have "cs c' i = cs c i" using RecvMarker assms snap by auto
also have "... = cs d' i" using RecvMarker assms snap sdr False by auto
finally show ?thesis by simp
qed
next
case no_snap: False
then have nsdr: "~ has_snapshotted d r" using assms by blast
show ?thesis
proof (cases "i = i'")
case True
then have "cs c' i = (fst (cs c i), Done)"
using RecvMarker assms(3) next_recv_marker by blast
also have "... = cs d' i"
using RecvMarker True assms(2) assms(4) by auto
finally show ?thesis using True by simp
next
case not_i: False
with assms RecvMarker chan no_snap show ?thesis by (cases "r = q", auto)
qed
qed
next
case (Trans r u u')
then show ?thesis
by (metis assms(2) assms(3) assms(4) event.disc(1) regular_event same_cs_implies_same_resulting_cs)
next
case (Send i' r s u u' m)
then show ?thesis
by (metis assms(2) assms(3) assms(4) event.disc(7) regular_event same_cs_implies_same_resulting_cs)
next
case (Recv i' r s u u' m)
then show ?thesis
by (metis assms(2) assms(3) assms(4) event.disc(13) regular_event same_cs_implies_same_resulting_cs)
qed
qed
lemma swap_Snapshot_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof -
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
have "ev = Snapshot ?p"
by (metis assms(3) event.collapse(4))
obtain u'' u''' where "ev' = Trans ?q u'' u'''"
by (metis assms(4) event.collapse(1))
have "msgs c i = msgs d' i"
using Trans_msg assms(4) assms(5) by blast
then have "msgs e' i = msgs d i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d' p"
using assms(4) assms(5) regular_event_preserves_process_snapshots by auto
- moreover have "msgs c i = msgs d' i" using `msgs c i = msgs d' i` by auto
+ moreover have "msgs c i = msgs d' i" using \<open>msgs c i = msgs d' i\<close> by auto
moreover have "c \<turnstile> ev \<mapsto> d" using assms by auto
moreover have "d' \<turnstile> ev \<mapsto> e'" using assms by auto
moreover have "~ regular_event ev" using assms by auto
ultimately show ?thesis by (blast intro: same_messages_2[symmetric])
qed
then have "msgs d i = msgs e i"
using Trans_msg assms(2) assms(4) by blast
then show ?thesis
by (simp add: \<open>msgs e' i = msgs d i\<close>)
qed
lemma swap_msgs_Trans_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecvMarker ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e' i = msgs e i"
proof -
have nr: "~ regular_event ev"
using assms(3) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r where RecvMarker: "ev = RecvMarker i' ?p r"
by (metis assms(3) event.collapse(5))
obtain u'' u''' where Trans: "ev' = Trans ?q u'' u'''"
by (metis assms(4) event.collapse(1))
have "msgs c i = msgs d' i"
using Trans_msg assms(4) assms(5) by blast
then have "msgs e' i = msgs d i"
proof -
have "\<forall>p. has_snapshotted d' p = has_snapshotted c p"
using assms(4) assms(5) regular_event_preserves_process_snapshots by auto
moreover have "~ regular_event ev" using assms by auto
moreover have "\<forall>n. msgs d' n = msgs c n" (* why does he need this assumption? *)
by (metis Trans assms(5) local.next.simps(3))
ultimately show ?thesis
using assms(1) assms(6) same_messages_2 by blast
qed
thm same_messages_2
then have "msgs d i = msgs e i"
using Trans_msg assms(2) assms(4) by blast
then show ?thesis
by (simp add: \<open>msgs e' i = msgs d i\<close>)
qed
lemma swap_Trans_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_Snapshot_Trans by auto
lemma swap_Send_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_msgs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
have Snapshot: "ev' = Snapshot ?q"
by (metis assms(4) event.collapse(4))
show ?thesis
proof (cases "i = i'"; cases "p = ?q")
assume asm: "i = i'" "p = ?q"
then have "?p = p"
proof -
have "channel i' = Some (p, q)" using chan asm by simp
then show ?thesis using assms can_occur_def Send chan
by (metis (mono_tags, lifting) event.simps(27) happen_implies_can_occur option.inject prod.inject)
qed
then show ?thesis using assms asm by simp
next
assume "i = i'" "p \<noteq> ?q"
then have "msgs d i = msgs c i @ [Msg m]"
by (metis Send assms(1) next_send)
moreover have "msgs e i = msgs d i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(2) chan next_snapshot option.inject)
moreover have "msgs d' i = msgs c i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(5) chan next_snapshot option.inject)
moreover have "msgs e' i = msgs d' i @ [Msg m]"
by (metis Send \<open>i = i'\<close> assms(6) next_send)
ultimately show ?thesis by simp
next
assume asm: "i \<noteq> i'" "p = ?q"
then have "msgs d i = msgs c i"
by (metis Send assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
moreover have "msgs e i = msgs c i @ [Marker]"
by (metis (full_types) Snapshot asm(2) assms(2) calculation chan next_snapshot)
moreover have "msgs d' i = msgs c i @ [Marker]"
by (metis (full_types) Snapshot asm(2) assms(5) chan next_snapshot)
moreover have "msgs e' i = msgs d' i"
by (metis Send asm(1) assms(6) next_send)
ultimately show ?thesis by simp
next
assume "i \<noteq> i'" "p \<noteq> ?q"
then have "msgs c i = msgs d i"
by (metis Send assms(1) assms(3) event.sel(8) msgs_unchanged_for_other_is regular_event)
then have "msgs e i = msgs d' i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(2,5) chan next_snapshot option.inject)
then show ?thesis
by (metis Send \<open>i \<noteq> i'\<close> assms(6) next_send)
qed
qed
lemma swap_Snapshot_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_Send_Snapshot by auto
lemma swap_Recv_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_msgs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Recv: "ev = Recv i' ?p r u u' m"
by (metis assms(3) event.collapse(3))
have Snapshot: "ev' = Snapshot ?q"
by (metis assms(4) event.collapse(4))
show ?thesis
proof (cases "i = i'"; cases "p = ?q")
assume "i = i'" "p = ?q"
then have "Msg m # msgs d i = msgs c i"
by (metis Recv assms(1) next_recv)
moreover have "msgs e i = msgs d i @ [Marker]"
by (metis (full_types) Snapshot \<open>p = occurs_on ev'\<close> assms(2) chan next_snapshot)
moreover have "msgs d' i = msgs c i @ [Marker]"
by (metis (full_types) Snapshot \<open>p = occurs_on ev'\<close> assms(5) chan next_snapshot)
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv \<open>i = i'\<close> assms(6) next_recv)
ultimately show ?thesis
by (metis list.sel(3) neq_Nil_conv tl_append2)
next
assume "i = i'" "p \<noteq> ?q"
then have "Msg m # msgs d i = msgs c i"
by (metis Recv assms(1) next_recv)
moreover have "msgs e i = msgs d i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(2) chan next_snapshot option.inject)
moreover have "msgs d' i = msgs c i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(5) chan next_snapshot option.inject)
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv \<open>i = i'\<close> assms(6) next_recv)
ultimately show ?thesis by (metis list.inject)
next
assume "i \<noteq> i'" "p = ?q"
then have "msgs d i = msgs c i"
by (metis Recv assms(1) next_recv)
moreover have "msgs e i = msgs d i @ [Marker]"
by (metis (full_types) Snapshot \<open>p = occurs_on ev'\<close> assms(2) chan next_snapshot)
moreover have "msgs d' i = msgs c i @ [Marker]"
by (metis (full_types) Snapshot \<open>p = occurs_on ev'\<close> assms(5) chan next_snapshot)
moreover have "msgs e' i = msgs d' i"
by (metis Recv \<open>i ~= i'\<close> assms(6) next_recv)
ultimately show ?thesis by simp
next
assume "i \<noteq> i'" "p \<noteq> ?q"
then have "msgs d i = msgs c i"
by (metis Recv assms(1) next_recv)
moreover have "msgs e i = msgs d i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(2) chan next_snapshot option.inject)
moreover have "msgs d' i = msgs c i"
by (metis Pair_inject Snapshot \<open>p \<noteq> occurs_on ev'\<close> assms(5) chan next_snapshot option.inject)
moreover have "msgs e' i = msgs d' i"
by (metis Recv \<open>i ~= i'\<close> assms(6) next_recv)
ultimately show ?thesis by auto
qed
qed
lemma swap_Snapshot_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_Recv_Snapshot by auto
lemma swap_msgs_Recv_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isRecvMarker ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_msgs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
obtain i' p' r u u' m where Recv: "ev = Recv i' p' r u u' m"
by (metis assms(3) event.collapse(3))
obtain i'' q' s where RecvMarker: "ev' = RecvMarker i'' q' s"
by (metis assms(4) event.collapse(5))
have "i' \<noteq> i''"
proof (rule ccontr)
assume "~ i' \<noteq> i''"
then have "channel i' = channel i''" by auto
then have "Some (r, p') = Some (s, q')" using assms can_occur_def Recv RecvMarker by simp
then show False using assms
by (metis Recv RecvMarker event.sel(3,5) option.inject prod.inject)
qed
then show ?thesis
proof (cases "i = i' \<or> i = i''")
case True
then show ?thesis
proof (elim disjE)
assume "i = i'"
then have pqrp: "(p, q) = (r, p')"
by (metis Recv assms(1) chan distributed_system.can_occur_Recv distributed_system_axioms next_recv option.inject)
then show ?thesis
proof (cases "has_snapshotted c q'")
case snap: True
then have "Msg m # msgs d i = msgs c i"
by (metis Recv \<open>i = i'\<close> assms(1) next_recv)
moreover have "msgs c i = msgs d' i"
using RecvMarker \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(5) msgs_unchanged_if_snapshotted_RecvMarker_for_other_is snap by blast
moreover have "msgs d i = msgs e i"
using RecvMarker \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms(1) assms(2) snap snapshot_state_unchanged by auto
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv \<open>i = i'\<close> assms(6) next_recv)
ultimately show ?thesis by (metis list.inject)
next
case no_snap: False
then have msgs_d: "Msg m # msgs d i = msgs c i"
by (metis Recv \<open>i = i'\<close> assms(1) next_recv)
then show ?thesis
proof (cases "q' = r")
case True
then have "msgs d' i = msgs c i @ [Marker]"
proof -
have "channel i = Some (q', q)"
using True chan pqrp by blast
then show ?thesis using RecvMarker assms no_snap
by (simp add: no_snap \<open>i = i'\<close> \<open>i' \<noteq> i''\<close>)
qed
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv \<open>i = i'\<close> assms(6) next_recv)
moreover have "msgs e i = msgs d i @ [Marker]"
proof -
have "ps d q' = ps c q'"
using assms(1) assms(7) no_state_change_if_no_event RecvMarker by auto
then show ?thesis
using RecvMarker \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> assms True chan no_snap pqrp by simp
qed
ultimately show ?thesis using msgs_d
by (metis append_self_conv2 list.inject list.sel(3) message.distinct(1) tl_append2)
next
case False
then have "msgs e i = msgs d i"
proof -
have "~ has_snapshotted d q'"
using assms(1) assms(7) no_snap no_state_change_if_no_event RecvMarker by auto
moreover have "\<nexists>r. channel i = Some (q', r)" using chan False pqrp by auto
- moreover have "i \<noteq> i''" using `i = i'` `i' \<noteq> i''` by simp
+ moreover have "i \<noteq> i''" using \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> by simp
ultimately show ?thesis using RecvMarker assms by simp
qed
moreover have "msgs d' i = msgs c i"
proof -
have "\<nexists>r. channel i = Some (q', r)"
using False chan pqrp by auto
- moreover have "i \<noteq> i''" using `i = i'` `i' \<noteq> i''` by simp
+ moreover have "i \<noteq> i''" using \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> by simp
ultimately show ?thesis using RecvMarker assms(5) no_snap by auto
qed
moreover have "Msg m # msgs e' i = msgs d' i"
by (metis Recv \<open>i = i'\<close> assms(6) next_recv)
ultimately show ?thesis using msgs_d
by (metis list.inject)
qed
qed
next
assume "i = i''"
then have "msgs d i = msgs c i" using assms
by (metis Recv \<open>i' \<noteq> i''\<close> next_recv)
moreover have "msgs e i = msgs d' i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
by (metis Recv assms(1) next_recv)
then show ?thesis
by (meson assms(2) assms(5) calculation same_messages_2 same_messages_imply_same_resulting_messages)
qed
moreover have "msgs e' i = msgs d' i"
- using assms by (metis Recv \<open>i' \<noteq> i''\<close> `i = i''` next_recv)
+ using assms by (metis Recv \<open>i' \<noteq> i''\<close> \<open>i = i''\<close> next_recv)
ultimately show ?thesis by simp
qed
next
assume asm: "~ (i = i' \<or> i = i'')"
then have "msgs c i = msgs d i"
by (metis Recv assms(1) assms(3) event.distinct_disc(16,18) event.sel(9) msgs_unchanged_for_other_is nonregular_event)
then have "msgs d' i = msgs e i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
by (meson \<open>msgs c i = msgs d i\<close> assms(2) assms(5) same_messages_2 same_messages_imply_same_resulting_messages)
qed
then show ?thesis
by (metis Recv asm assms(3) assms(6) event.distinct_disc(16,18) event.sel(9) msgs_unchanged_for_other_is nonregular_event)
qed
qed
lemma swap_RecvMarker_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecvMarker ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_msgs_Recv_RecvMarker by auto
lemma swap_msgs_Send_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isRecvMarker ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_msgs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' p' r u u' m where Send: "ev = Send i' p' r u u' m"
by (metis assms(3) event.collapse(2))
obtain i'' q' s where RecvMarker: "ev' = RecvMarker i'' q' s"
by (metis assms(4) event.collapse(5))
have "p' \<noteq> q'" using Send RecvMarker assms by simp
show ?thesis
proof (cases "i = i'"; cases "i = i''", goal_cases)
case 1
then have "msgs e' i = msgs d' i @ [Msg m]"
by (metis Send assms(6) next_send)
- moreover have "Marker # msgs d' i = msgs c i" using `i = i''` RecvMarker assms by simp
+ moreover have "Marker # msgs d' i = msgs c i" using \<open>i = i''\<close> RecvMarker assms by simp
moreover have "msgs d i = msgs c i @ [Msg m]"
by (metis "1"(1) Send assms(1) next_send)
- moreover have "Marker # msgs e i = msgs d i" using `i = i''` RecvMarker assms by simp
+ moreover have "Marker # msgs e i = msgs d i" using \<open>i = i''\<close> RecvMarker assms by simp
ultimately show ?thesis
by (metis append_self_conv2 list.inject list.sel(3) message.distinct(1) tl_append2)
next
case 2
then have pqpr: "(p, q) = (p', r)" using chan Send can_occur_def assms by simp
then have "msgs d i = msgs c i @ [Msg m]"
by (metis 2(1) Send assms(1) next_send)
moreover have "msgs e' i = msgs d' i @ [Msg m]"
by (metis "2"(1) Send assms(6) next_send)
moreover have "msgs d' i = msgs c i"
proof -
- have "\<nexists>r. channel i = Some (q', r)" using `p' \<noteq> q'` chan pqpr by simp
- with RecvMarker `i \<noteq> i''` `i = i'` assms show ?thesis by (cases "has_snapshotted c q'", auto)
+ have "\<nexists>r. channel i = Some (q', r)" using \<open>p' \<noteq> q'\<close> chan pqpr by simp
+ with RecvMarker \<open>i \<noteq> i''\<close> \<open>i = i'\<close> assms show ?thesis by (cases "has_snapshotted c q'", auto)
qed
moreover have "msgs e i = msgs d i"
proof -
- have "\<nexists>r. channel i = Some (q', r)" using `p' \<noteq> q'` chan pqpr by simp
- with RecvMarker `i \<noteq> i''` `i = i'` assms show ?thesis by (cases "has_snapshotted d q'", auto)
+ have "\<nexists>r. channel i = Some (q', r)" using \<open>p' \<noteq> q'\<close> chan pqpr by simp
+ with RecvMarker \<open>i \<noteq> i''\<close> \<open>i = i'\<close> assms show ?thesis by (cases "has_snapshotted d q'", auto)
qed
ultimately show ?thesis by simp
next
assume 3: "i \<noteq> i'" "i = i''"
then have mcd: "msgs c i = msgs d i"
by (metis Send assms(1) next_send)
moreover have "msgs e i = msgs d' i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
moreover have "~ regular_event ev'" using assms by auto
ultimately show ?thesis using mcd assms(2,5) by (blast intro: same_messages_2[symmetric])
qed
moreover have "msgs e' i = msgs d' i"
by (metis "3"(1) Send assms(6) next_send)
ultimately show ?thesis by simp
next
assume 4: "i \<noteq> i'" "i \<noteq> i''"
have mcd: "msgs c i = msgs d i"
by (metis "4"(1) Send assms(1) assms(3) event.distinct_disc(12,14) event.sel(8) msgs_unchanged_for_other_is nonregular_event)
have "msgs d' i = msgs e i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
moreover have "~ regular_event ev'" using assms by auto
ultimately show ?thesis using mcd assms(2,5) same_messages_2 by blast
qed
moreover have "msgs e' i = msgs d' i"
by (metis "4"(1) Send assms(6) next_send)
ultimately show ?thesis by simp
qed
qed
lemma swap_RecvMarker_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecvMarker ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"msgs e i = msgs e' i"
using assms swap_msgs_Send_RecvMarker by auto
lemma swap_cs_Trans_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain u'' u''' where "ev = Trans ?p u'' u'''"
by (metis assms(3) event.collapse(1))
have "ev' = Snapshot ?q"
by (metis assms(4) event.collapse(4))
have "cs d i = cs c i"
by (metis assms(1) assms(3) event.distinct_disc(4) no_cs_change_if_no_event regular_event)
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
using assms(3) assms(6) no_cs_change_if_no_event regular_event by blast
finally show ?thesis by simp
qed
lemma swap_cs_Snapshot_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
using swap_cs_Trans_Snapshot assms by auto
lemma swap_cs_Send_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
have Snapshot: "ev' = Snapshot ?q"
by (metis assms(4) event.collapse(4))
have "cs d i = cs c i"
by (metis Send assms(1) next_send)
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
using assms(3) assms(6) no_cs_change_if_no_event regular_event by blast
finally show ?thesis by simp
qed
lemma swap_cs_Snapshot_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
using swap_cs_Send_Snapshot assms by auto
lemma swap_cs_Recv_Snapshot:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isSnapshot ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Recv: "ev = Recv i' ?p r u u' m"
by (metis assms(3) event.collapse(3))
have Snapshot: "ev' = Snapshot ?q"
by (metis assms(4) event.collapse(4))
show ?thesis
proof (cases "i = i'")
case True
then show ?thesis
proof (cases "snd (cs c i) = Recording")
case True
- then have "cs d i = (fst (cs c i) @ [m], Recording)" using Recv assms True `i = i'` chan
+ then have "cs d i = (fst (cs c i) @ [m], Recording)" using Recv assms True \<open>i = i'\<close> chan
by (metis next_recv)
moreover have "cs e i = cs d i"
by (metis Snapshot assms(2) calculation fst_conv next_snapshot)
moreover have "cs c i = cs d' i"
by (metis Snapshot True assms(5) next_snapshot prod.collapse)
moreover have "cs e' i = (fst (cs d' i) @ [m], Recording)"
by (metis (mono_tags, lifting) Recv assms(1) assms(6) calculation(1) calculation(3) next_recv)
ultimately show ?thesis by simp
next
case False
have "cs d i = cs c i"
by (metis False Recv assms(1) next_recv)
have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
moreover have "cs d' i = cs e' i"
proof -
have "cs d' i = cs c i"
by (metis Pair_inject Recv Snapshot True assms(1) assms(5) assms(7) can_occur_Recv distributed_system.happen_implies_can_occur distributed_system.next_snapshot distributed_system_axioms option.inject)
- then show ?thesis using chan `i = i'` False Recv assms
+ then show ?thesis using chan \<open>i = i'\<close> False Recv assms
by (metis next_recv)
qed
ultimately show ?thesis by simp
qed
next
case False
have "cs d i = cs c i"
by (metis False Recv assms(1) next_recv)
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
by (metis False Recv assms(6) next_recv)
finally show ?thesis by simp
qed
qed
lemma swap_cs_Snapshot_Recv:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSnapshot ev" and
"isRecv ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"cs e i = cs e' i"
using swap_cs_Recv_Snapshot assms by auto
lemma swap_cs_Trans_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isTrans ev" and
"isRecvMarker ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain u'' u''' where "ev = Trans ?p u'' u'''"
by (metis assms(3) event.collapse(1))
obtain i' s where "ev' = RecvMarker i' ?q s"
by (metis assms(4) event.collapse(5))
have "cs d i = cs c i"
by (metis assms(1) assms(3) event.distinct_disc(4) no_cs_change_if_no_event regular_event)
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
using assms(3) assms(6) no_cs_change_if_no_event regular_event by blast
finally show ?thesis by simp
qed
lemma swap_cs_RecvMarker_Trans:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecvMarker ev" and
"isTrans ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
using swap_cs_Trans_RecvMarker assms by auto
lemma swap_cs_Send_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isSend ev" and
"isRecvMarker ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
let ?p = "occurs_on ev"
let ?q = "occurs_on ev'"
obtain i' r u u' m where Send: "ev = Send i' ?p r u u' m"
by (metis assms(3) event.collapse(2))
obtain i'' s where RecvMarker: "ev' = RecvMarker i'' ?q s"
by (metis assms(4) event.collapse(5))
have "cs d i = cs c i"
by (metis assms(1) assms(3) event.distinct_disc(10,12,14) no_cs_change_if_no_event nonregular_event)
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
using assms(3) assms(6) no_cs_change_if_no_event regular_event by blast
finally show ?thesis by simp
qed
lemma swap_cs_RecvMarker_Send:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecvMarker ev" and
"isSend ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'"
shows
"cs e i = cs e' i"
using swap_cs_Send_RecvMarker assms by auto
lemma swap_cs_Recv_RecvMarker:
assumes
"c \<turnstile> ev \<mapsto> d" and
"d \<turnstile> ev' \<mapsto> e" and
"isRecv ev" and
"isRecvMarker ev'" and
"c \<turnstile> ev' \<mapsto> d'" and
"d' \<turnstile> ev \<mapsto> e'" and
"occurs_on ev \<noteq> occurs_on ev'"
shows
"cs e i = cs e' i"
proof (cases "channel i = None")
case True
then show ?thesis
by (metis assms(1) assms(2) assms(5) assms(6) no_cs_change_if_no_channel)
next
case False
then obtain p q where chan: "channel i = Some (p, q)" by auto
have nr: "~ regular_event ev'"
using assms(4) nonregular_event by blast
obtain i' p' r u u' m where Recv: "ev = Recv i' p' r u u' m"
by (metis assms(3) event.collapse(3))
obtain i'' q' s where RecvMarker: "ev' = RecvMarker i'' q' s"
by (metis assms(4) event.collapse(5))
have "i' \<noteq> i''"
proof (rule ccontr)
assume "~ i' \<noteq> i''"
then have "channel i' = channel i''" by simp
then have "(r, p') = (s, q')" using Recv RecvMarker assms can_occur_def by simp
then show False using Recv RecvMarker assms can_occur_def by simp
qed
show ?thesis
proof (cases "i = i'")
case True
then have pqrp: "(p, q) = (r, p')" using Recv assms can_occur_def chan by simp
then show ?thesis
proof (cases "snd (cs c i)")
case NotStarted
- then have "cs d i = cs c i" using assms Recv `i = i'` by simp
+ then have "cs d i = cs c i" using assms Recv \<open>i = i'\<close> by simp
moreover have "cs d' i = cs e i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
with assms(2,5) calculation show ?thesis by (blast intro: same_cs_2[symmetric])
qed
thm same_cs_2
moreover have "cs d' i = cs e' i"
proof -
have "cs d' i = cs c i"
proof -
have "\<nexists>r. channel i = Some (r, q')"
using Recv RecvMarker assms(7) chan pqrp by auto
- with RecvMarker assms chan `i = i'` `i' \<noteq> i''` show ?thesis
+ with RecvMarker assms chan \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> show ?thesis
by (cases "has_snapshotted c q'", auto)
qed
- then show ?thesis using assms Recv `i = i'` NotStarted by simp
+ then show ?thesis using assms Recv \<open>i = i'\<close> NotStarted by simp
qed
ultimately show ?thesis by simp
next
case Done
- then have "cs d i = cs c i" using assms Recv `i = i'` by simp
+ then have "cs d i = cs c i" using assms Recv \<open>i = i'\<close> by simp
moreover have "cs d' i = cs e i"
proof -
have "\<forall>p. has_snapshotted c p = has_snapshotted d p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis using assms(2,5) calculation by (blast intro: same_cs_2[symmetric])
qed
moreover have "cs d' i = cs e' i"
proof -
have "cs d' i = cs c i"
proof -
have "\<nexists>r. channel i = Some (r, q')"
using Recv RecvMarker assms(7) chan pqrp by auto
- with RecvMarker assms chan `i = i'` `i' \<noteq> i''` show ?thesis
+ with RecvMarker assms chan \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> show ?thesis
by (cases "has_snapshotted c q'", auto)
qed
- then show ?thesis using assms Recv `i = i'` Done by simp
+ then show ?thesis using assms Recv \<open>i = i'\<close> Done by simp
qed
ultimately show ?thesis by simp
next
case Recording
have "cs d i = (fst (cs c i) @ [m], Recording)"
using Recording Recv True assms(1) by auto
moreover have "cs e i = cs d i"
proof -
have "\<nexists>r. channel i = Some (r, q')"
using Recv RecvMarker assms(7) chan pqrp by auto
- with RecvMarker assms chan `i = i'` `i' \<noteq> i''` show ?thesis
+ with RecvMarker assms chan \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> show ?thesis
by (cases "has_snapshotted d q'", auto)
qed
moreover have "cs c i = cs d' i "
proof -
have "\<nexists>r. channel i = Some (r, q')"
using Recv RecvMarker assms(7) chan pqrp by auto
- with RecvMarker assms chan `i = i'` `i' \<noteq> i''` show ?thesis
+ with RecvMarker assms chan \<open>i = i'\<close> \<open>i' \<noteq> i''\<close> show ?thesis
by (cases "has_snapshotted c q'", auto)
qed
moreover have "cs e' i = (fst (cs d' i) @ [m], Recording)"
using Recording Recv True assms(6) calculation(3) by auto
ultimately show ?thesis by simp
qed
next
case False
have "cs d i = cs c i"
using False Recv assms(1) by auto
then have "cs e i = cs d' i"
proof -
have "\<forall>p. has_snapshotted d p = has_snapshotted c p"
using assms(1) assms(3) regular_event_preserves_process_snapshots by auto
then show ?thesis
using \<open>cs d i = cs c i\<close> assms(2) assms(5) same_cs_2 by blast
qed
also have "... = cs e' i"
using False Recv assms(6) by auto
finally show ?thesis by simp
qed
qed
end (* context distributed_system *)
end (* theory Swap *)
diff --git a/thys/Chandy_Lamport/Trace.thy b/thys/Chandy_Lamport/Trace.thy
--- a/thys/Chandy_Lamport/Trace.thy
+++ b/thys/Chandy_Lamport/Trace.thy
@@ -1,581 +1,581 @@
section \<open>Traces\<close>
text \<open>Traces extend transitions to finitely many intermediate events.\<close>
theory Trace
imports
"HOL-Library.Sublist"
Distributed_System
begin
context distributed_system
begin
text \<open>We can think of a trace as the transitive closure of the next
relation. A trace consists of initial and final configurations $c$ and
$c'$, with an ordered list of events $t$ occurring sequentially on $c$,
yielding $c'$.\<close>
inductive (in distributed_system) trace where
tr_init: "trace c [] c"
| tr_step: "\<lbrakk> c \<turnstile> ev \<mapsto> c'; trace c' t c'' \<rbrakk>
\<Longrightarrow> trace c (ev # t) c''"
subsection \<open>Properties of traces\<close>
lemma trace_trans:
shows
"\<lbrakk> trace c t c';
trace c' t' c''
\<rbrakk> \<Longrightarrow> trace c (t @ t') c''"
proof (induct c t c' rule:trace.induct)
case tr_init
then show ?case by simp
next
case tr_step
then show ?case using trace.tr_step by auto
qed
lemma trace_decomp_head:
assumes
"trace c (ev # t) c'"
shows
"\<exists>c''. c \<turnstile> ev \<mapsto> c'' \<and> trace c'' t c'"
using assms trace.simps by blast
lemma trace_decomp_tail:
shows
"trace c (t @ [ev]) c' \<Longrightarrow> \<exists>c''. trace c t c'' \<and> c'' \<turnstile> ev \<mapsto> c'"
proof (induct t arbitrary: c)
case Nil
then show ?case
by (metis (mono_tags, lifting) append_Nil distributed_system.trace.simps distributed_system_axioms list.discI list.sel(1) list.sel(3))
next
case (Cons ev' t)
then obtain d where step: "c \<turnstile> ev' \<mapsto> d" and "trace d (t @ [ev]) c'" using trace_decomp_head by force
then obtain d' where tr: "trace d t d'" and "d' \<turnstile> ev \<mapsto> c'" using Cons.hyps by blast
moreover have "trace c (ev' # t) d'" using step tr trace.tr_step by simp
ultimately show ?case by auto
qed
lemma trace_snoc:
assumes
"trace c t c'" and
"c' \<turnstile> ev \<mapsto> c''"
shows
"trace c (t @ [ev]) c''"
using assms(1) assms(2) tr_init tr_step trace_trans by auto
lemma trace_rev_induct [consumes 1, case_names tr_rev_init tr_rev_step]:
"\<lbrakk> trace c t c';
(\<And>c. P c [] c);
(\<And>c t c' ev c''. trace c t c' \<Longrightarrow> P c t c' \<Longrightarrow> c' \<turnstile> ev \<mapsto> c'' \<Longrightarrow> P c (t @ [ev]) c'')
\<rbrakk> \<Longrightarrow> P c t c'"
proof (induct t arbitrary: c' rule:rev_induct)
case Nil
then show ?case
using distributed_system.trace.cases distributed_system_axioms by blast
next
case (snoc ev t)
then obtain c'' where "trace c t c''" "c'' \<turnstile> ev \<mapsto> c'" using trace_decomp_tail by blast
then show ?case using snoc by simp
qed
lemma trace_and_start_determines_end:
shows
"trace c t c' \<Longrightarrow> trace c t d' \<Longrightarrow> c' = d'"
proof (induct c t c' arbitrary: d' rule:trace_rev_induct)
case tr_rev_init
then show ?case using trace.cases by fastforce
next
case (tr_rev_step c t c' ev c'')
then obtain d'' where "trace c t d''" "d'' \<turnstile> ev \<mapsto> d'" using trace_decomp_tail by blast
then show ?case using tr_rev_step state_and_event_determine_next by simp
qed
lemma suffix_split_trace:
shows
"\<lbrakk> trace c t c';
suffix t' t
\<rbrakk> \<Longrightarrow> \<exists>c''. trace c'' t' c'"
proof (induct t arbitrary: c)
case Nil
then have "t' = []" by simp
then have "trace c' t' c'" using tr_init by simp
then show ?case by blast
next
case (Cons ev t'')
from Cons.prems have q: "suffix t' t'' \<or> t' = ev # t''" by (meson suffix_Cons)
thus ?case
proof (cases "suffix t' t''")
case True
then show ?thesis using Cons.hyps Cons.prems(1) trace.simps by blast
next
case False
hence "t' = ev # t''" using q by simp
thus ?thesis using Cons.hyps Cons.prems by blast
qed
qed
lemma prefix_split_trace:
fixes
c :: "('p, 's, 'm) configuration" and
t :: "('p, 's, 'm) trace"
shows
"\<lbrakk> \<exists>c'. trace c t c';
prefix t' t
\<rbrakk> \<Longrightarrow> \<exists>c''. trace c t' c''"
proof (induct t rule:rev_induct)
case Nil
then show ?case by simp
next
case (snoc ev t'')
from snoc.prems have q: "prefix t' t'' \<or> t' = t'' @ [ev]" by auto
thus ?case
proof (cases "prefix t' t''")
case True
thus ?thesis using trace_decomp_tail using snoc.hyps snoc.prems(1) trace.simps by blast
next
case False
thus ?thesis using q snoc.prems by fast
qed
qed
lemma split_trace:
shows
"\<lbrakk> trace c t c';
t = t' @ t''
\<rbrakk> \<Longrightarrow> \<exists>c''. trace c t' c'' \<and> trace c'' t'' c'"
proof (induct t'' arbitrary: t')
case Nil
then show ?case using tr_init by auto
next
case (Cons ev t'')
obtain c'' where p: "trace c (t' @ [ev]) c''"
using Cons.prems prefix_split_trace rotate1.simps(2) by force
then have "trace c'' t'' c'"
using Cons.hyps Cons.prems trace_and_start_determines_end by force
then show ?case
by (meson distributed_system.tr_step distributed_system.trace_decomp_tail distributed_system_axioms p)
qed
subsection \<open>Describing intermediate configurations\<close>
definition construct_fun_from_rel :: "('a * 'b) set \<Rightarrow> 'a \<Rightarrow> 'b" where
"construct_fun_from_rel R x = (THE y. (x,y) \<in> R)"
definition trace_rel where
"trace_rel \<equiv> {((x, t'), y). trace x t' y}"
lemma fun_must_admit_trace:
shows
"single_valued R \<Longrightarrow> x \<in> Domain R
\<Longrightarrow> (x, construct_fun_from_rel R x) \<in> R"
unfolding construct_fun_from_rel_def
by (rule theI') (auto simp add: single_valued_def)
lemma single_valued_trace_rel:
shows
"single_valued trace_rel"
proof (rule single_valuedI)
fix x y y'
assume asm: "(x, y) \<in> trace_rel" "(x, y') \<in> trace_rel"
then obtain x' t where "x = (x', t)"
by (meson surj_pair)
then have "trace x' t y" "trace x' t y'"
using asm trace_rel_def by auto
then show "y = y'"
using trace_and_start_determines_end by blast
qed
definition run_trace where
"run_trace \<equiv> construct_fun_from_rel trace_rel"
text \<open>In order to describe intermediate configurations
of a trace we introduce the $s$ function definition, which,
given an initial configuration $c$, a trace $t$ and an index $i \in \mathbb{N}$,
determines the unique state after the first $i$ events of $t$.\<close>
definition s where
"s c t i = (THE c'. trace c (take i t) c')"
lemma s_is_partial_execution:
shows
"s c t i = run_trace (c, take i t)"
unfolding s_def run_trace_def
construct_fun_from_rel_def trace_rel_def
by auto
lemma exists_trace_for_any_i:
assumes
"\<exists>c'. trace c t c'"
shows
"trace c (take i t) (s c t i)"
proof -
have "prefix (take i t) t" using take_is_prefix by auto
then obtain c'' where tr: "trace c (take i t) c''" using assms prefix_split_trace by blast
then show ?thesis
proof -
have "((c, take i t), s c t i) \<in> trace_rel"
unfolding s_def trace_rel_def construct_fun_from_rel_def
by (metis case_prod_conv distributed_system.trace_and_start_determines_end distributed_system_axioms mem_Collect_eq the_equality tr)
then show ?thesis by (simp add: trace_rel_def)
qed
qed
lemma exists_trace_for_any_i_j:
assumes
"\<exists>c'. trace c t c'" and
"i \<le> j"
shows
"trace (s c t i) (take (j - i) (drop i t)) (s c t j)"
proof -
have "trace c (take j t) (s c t j)" using exists_trace_for_any_i assms by simp
- from `j \<ge> i` have "take j t = take i t @ (take (j - i) (drop i t))"
+ from \<open>j \<ge> i\<close> have "take j t = take i t @ (take (j - i) (drop i t))"
by (metis le_add_diff_inverse take_add)
then have "trace c (take i t) (s c t i) \<and> trace (s c t i) (take (j - i) (drop i t)) (s c t j)"
by (metis (no_types, lifting) assms(1) exists_trace_for_any_i split_trace trace_and_start_determines_end)
then show ?thesis by simp
qed
lemma step_Suc:
assumes
"i < length t" and
valid: "trace c t c'"
shows "(s c t i) \<turnstile> (t ! i) \<mapsto> (s c t (Suc i))"
proof -
have ex_trace: "trace (s c t i) (take (Suc i - i) (drop i t)) (s c t (Suc i))"
using exists_trace_for_any_i_j le_less valid by blast
moreover have "Suc i - i = 1" by auto
moreover have "take 1 (drop i t) = [t ! i]"
by (metis \<open>Suc i - i = 1\<close> assms(1) hd_drop_conv_nth le_add_diff_inverse lessI nat_less_le same_append_eq take_add take_hd_drop)
ultimately show ?thesis
by (metis list.discI trace.simps trace_decomp_head)
qed
subsection \<open>Trace-related lemmas\<close>
lemma snapshot_state_unchanged_trace:
assumes
"trace c t c'" and
"ps c p = Some u"
shows
"ps c' p = Some u"
using assms snapshot_state_unchanged by (induct c t c', auto)
lemma no_state_change_if_only_nonregular_events:
shows
"\<lbrakk> trace c t c';
\<nexists>ev. ev \<in> set t \<and> regular_event ev \<and> occurs_on ev = p;
states c p = st
\<rbrakk> \<Longrightarrow> states c' p = st"
proof (induct c t c' rule:trace_rev_induct)
case (tr_rev_init c)
then show ?case by simp
next
case (tr_rev_step c t c' ev c'')
then have "states c' p = st"
proof -
have "\<nexists>ev. ev : set t \<and> regular_event ev \<and> occurs_on ev = p"
using tr_rev_step by auto
then show ?thesis using tr_rev_step by blast
qed
then show ?case
using tr_rev_step no_state_change_if_no_event no_state_change_if_nonregular_event
by auto
qed
lemma message_must_be_delivered_2_trace:
assumes
"trace c t c'" and
"m : set (msgs c i)" and
"m \<notin> set (msgs c' i)" and
"channel i = Some (q, p)"
shows
"\<exists>ev \<in> set t. (\<exists>p q. ev = RecvMarker i p q \<and> m = Marker) \<or> (\<exists>p q s s' m'. ev = Recv i q p s s' m' \<and> m = Msg m')"
proof (rule ccontr)
assume "~ (\<exists>ev \<in> set t. (\<exists>p q. ev = RecvMarker i p q \<and> m = Marker) \<or> (\<exists>p q s s' m'. ev = Recv i q p s s' m' \<and> m = Msg m'))" (is ?P)
have "\<lbrakk> trace c t c'; m : set (msgs c i); ?P \<rbrakk> \<Longrightarrow> m : set (msgs c' i)"
proof (induct c t c' rule:trace_rev_induct)
case (tr_rev_init c)
then show ?case by simp
next
case (tr_rev_step c t d ev c')
then have m_in_set: "m : set (msgs d i)"
using tr_rev_step by auto
then show ?case
proof (cases ev)
case (Snapshot r)
then show ?thesis
using distributed_system.message_must_be_delivered_2 distributed_system_axioms m_in_set tr_rev_step.hyps(3) by blast
next
case (RecvMarker i' r s)
then show ?thesis
proof (cases "m = Marker")
case True
then have "i' \<noteq> i" using tr_rev_step RecvMarker by simp
then show ?thesis
using RecvMarker m_in_set message_must_be_delivered_2 tr_rev_step.hyps(3) by blast
next
case False
then show ?thesis
using RecvMarker tr_rev_step.hyps(3) m_in_set message_must_be_delivered_2 by blast
qed
next
case (Trans r u u')
then show ?thesis
using tr_rev_step.hyps(3) m_in_set by auto
next
case (Send i' r s u u' m')
then show ?thesis
using distributed_system.message_must_be_delivered_2 distributed_system_axioms m_in_set tr_rev_step.hyps(3) by blast
next
case (Recv i' r s u u' m')
then show ?thesis
proof (cases "Msg m' = m")
case True
then have "i' \<noteq> i" using Recv tr_rev_step by auto
then show ?thesis
using Recv m_in_set tr_rev_step(3) by auto
next
case False
then show ?thesis
by (metis Recv event.distinct(17) event.inject(3) m_in_set message_must_be_delivered_2 tr_rev_step.hyps(3))
qed
qed
qed
- then have "m : set (msgs c' i)" using assms `?P` by auto
+ then have "m : set (msgs c' i)" using assms \<open>?P\<close> by auto
then show False using assms by simp
qed
lemma marker_must_be_delivered_2_trace:
assumes
"trace c t c'" and
"Marker : set (msgs c i)" and
"Marker \<notin> set (msgs c' i)" and
"channel i = Some (p, q)"
shows
"\<exists>ev \<in> set t. (\<exists>p q. ev = RecvMarker i p q)"
proof -
show "\<exists>ev \<in> set t. (\<exists>p q. ev = RecvMarker i p q)"
using assms message_must_be_delivered_2_trace by fast
qed
lemma snapshot_stable:
shows
"\<lbrakk> trace c t c';
has_snapshotted c p
\<rbrakk> \<Longrightarrow> has_snapshotted c' p"
proof (induct c t c' rule:trace_rev_induct)
case (tr_rev_init c)
then show ?case by blast
next
case (tr_rev_step c t c' ev c'')
then show ?case
proof (cases ev)
case (Snapshot q)
then have "p \<noteq> q" using tr_rev_step next_snapshot can_occur_def by auto
then show ?thesis using Snapshot tr_rev_step by auto
next
case (RecvMarker i q r)
with tr_rev_step show ?thesis
by (cases "p = q"; auto)
qed simp_all
qed
lemma snapshot_stable_2:
shows
"trace c t c' \<Longrightarrow> ~ has_snapshotted c' p \<Longrightarrow> ~ has_snapshotted c p"
using snapshot_stable by blast
lemma no_markers_if_all_snapshotted:
shows
"\<lbrakk> trace c t c';
\<forall>p. has_snapshotted c p;
Marker \<notin> set (msgs c i)
\<rbrakk> \<Longrightarrow> Marker \<notin> set (msgs c' i)"
proof (induct c t c' rule:trace_rev_induct)
case (tr_rev_init c)
then show ?case by simp
next
case (tr_rev_step c t c' ev c'')
have all_snapshotted: "\<forall>p. has_snapshotted c' p" using snapshot_stable tr_rev_step by auto
have no_marker: "Marker \<notin> set (msgs c' i)" using tr_rev_step by blast
then show ?case
proof (cases ev)
case (Snapshot r)
then show ?thesis using can_occur_def tr_rev_step all_snapshotted by auto
next
case (RecvMarker i' r s)
have "i' \<noteq> i"
proof (rule ccontr)
assume "~ i' \<noteq> i"
then have "Marker : set (msgs c i)"
using can_occur_def RecvMarker tr_rev_step RecvMarker_implies_Marker_in_set by blast
then show False using tr_rev_step by simp
qed
then show ?thesis using tr_rev_step all_snapshotted no_marker RecvMarker by auto
next
case (Trans p s s')
then show ?thesis using tr_rev_step no_marker by auto
next
case (Send i' r s u u' m)
then show ?thesis
proof (cases "i' = i")
case True
then have "msgs c'' i = msgs c' i @ [Msg m]" using tr_rev_step Send by auto
then show ?thesis using no_marker by auto
next
case False
then show ?thesis using Send tr_rev_step no_marker by auto
qed
next
case (Recv i' r s u u' m)
then show ?thesis
proof (cases "i = i'")
case True
then have "msgs c'' i = tl (msgs c' i)" using tr_rev_step Recv by auto
then show ?thesis using no_marker by (metis list.sel(2) list.set_sel(2))
next
case False
then show ?thesis using Recv tr_rev_step no_marker by auto
qed
qed
qed
lemma event_stays_valid_if_no_occurrence_trace:
shows
"\<lbrakk> trace c t c';
list_all (\<lambda>ev. occurs_on ev \<noteq> occurs_on ev') t;
can_occur ev' c
\<rbrakk> \<Longrightarrow> can_occur ev' c'"
proof (induct c t c' rule:trace_rev_induct)
case tr_rev_init
then show ?case by blast
next
case tr_rev_step
then show ?case using event_stays_valid_if_no_occurrence by auto
qed
lemma event_can_go_back_if_no_sender_trace:
shows
"\<lbrakk> trace c t c';
list_all (\<lambda>ev. occurs_on ev \<noteq> occurs_on ev') t;
can_occur ev' c';
~ isRecvMarker ev';
list_all (\<lambda>ev. ~ isSend ev) t
\<rbrakk> \<Longrightarrow> can_occur ev' c"
proof (induct c t c' rule:trace_rev_induct)
case tr_rev_init
then show ?case by blast
next
case tr_rev_step
then show ?case using event_can_go_back_if_no_sender by auto
qed
lemma done_only_from_recv_marker_trace:
assumes
"trace c t c'" and
"t \<noteq> []" and
"snd (cs c cid) \<noteq> Done" and
"snd (cs c' cid) = Done" and
"channel cid = Some (p, q)"
shows
"RecvMarker cid q p \<in> set t"
proof (rule ccontr)
assume "~ RecvMarker cid q p \<in> set t"
moreover have "\<lbrakk> trace c t c'; ~ RecvMarker cid q p \<in> set t; snd (cs c cid) \<noteq> Done; channel cid = Some (p, q) \<rbrakk>
\<Longrightarrow> snd (cs c' cid) \<noteq> Done"
proof (induct t arbitrary: c' rule:rev_induct)
case Nil
then show ?case
by (metis list.discI trace.simps)
next
case (snoc ev t)
then obtain d where ind: "trace c t d" and step: "d \<turnstile> ev \<mapsto> c'"
using trace_decomp_tail by blast
then have "snd (cs d cid) \<noteq> Done"
proof -
have "~ RecvMarker cid q p \<in> set t"
using snoc.prems(2) by auto
then show ?thesis using snoc ind by blast
qed
then show ?case
using done_only_from_recv_marker local.step snoc.prems(2) snoc.prems(4) by auto
qed
ultimately have "snd (cs c' cid) \<noteq> Done" using assms by blast
then show False using assms by simp
qed
lemma cs_not_not_started_stable_trace:
shows
"\<lbrakk> trace c t c'; snd (cs c cid) \<noteq> NotStarted; channel cid = Some (p, q) \<rbrakk> \<Longrightarrow> snd (cs c' cid) \<noteq> NotStarted"
proof (induct t arbitrary:c' rule:rev_induct)
case Nil
then show ?case
by (metis list.discI trace.simps)
next
case (snoc ev t)
then obtain d where tr: "trace c t d" and step: "d \<turnstile> ev \<mapsto> c'"
using trace_decomp_tail by blast
then have "snd (cs d cid) \<noteq> NotStarted" using snoc by blast
then show ?case using cs_not_not_started_stable snoc step by blast
qed
lemma no_messages_introduced_if_no_channel:
assumes
trace: "trace init t final" and
no_msgs_if_no_channel: "\<forall>i. channel i = None \<longrightarrow> msgs init i = []"
shows
"channel cid = None \<Longrightarrow> msgs (s init t i) cid = []"
proof (induct i)
case 0
then show ?case
by (metis assms exists_trace_for_any_i no_msgs_if_no_channel take0 tr_init trace_and_start_determines_end)
next
case (Suc n)
have f: "trace (s init t n) (take ((Suc n) - n) (drop n t)) (s init t (Suc n))"
using exists_trace_for_any_i_j order_le_less trace assms by blast
then show ?case
proof (cases "drop n t = Nil")
case True
then show ?thesis using Suc.hyps Suc.prems
by (metis f tr_init trace_and_start_determines_end take_Nil)
next
case False
have suc_n_minus_n: "Suc n - n = 1" by auto
then have "length (take ((Suc n) - n) (drop n t)) = 1" using False by auto
then obtain ev where "ev # Nil = take ((Suc n) - n) (drop n t)"
by (metis False One_nat_def suc_n_minus_n length_greater_0_conv self_append_conv2 take_eq_Nil take_hd_drop)
then have g: "(s init t n) \<turnstile> ev \<mapsto> (s init t (Suc n))"
by (metis f tr_init trace_and_start_determines_end trace_decomp_head)
then show ?thesis
proof (cases ev)
case (Snapshot r)
then show ?thesis
using Suc.hyps Suc.prems g by auto
next
case (RecvMarker cid' sr r)
have "cid' \<noteq> cid" using RecvMarker can_occur_def g Suc by auto
with RecvMarker Suc g show ?thesis by (cases "has_snapshotted (s init t n) sr", auto)
next
case (Trans r u u')
then show ?thesis
by (metis Suc.hyps Suc.prems g next_trans)
next
case (Send cid' r s u u' m)
have "cid' \<noteq> cid" using Send can_occur_def g Suc by auto
then show ?thesis using Suc g Send by simp
next
case (Recv cid' s r u u' m)
have "cid' \<noteq> cid" using Recv can_occur_def g Suc by auto
then show ?thesis using Suc g Recv by simp
qed
qed
qed
end (* context distributed_system *)
end (* theory Trace *)
diff --git a/thys/Complex_Geometry/Angles.thy b/thys/Complex_Geometry/Angles.thy
--- a/thys/Complex_Geometry/Angles.thy
+++ b/thys/Complex_Geometry/Angles.thy
@@ -1,510 +1,510 @@
(* ---------------------------------------------------------------------------- *)
subsection \<open>Angle between two vectors\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>In this section we introduce different measures of angle between two vectors (represented by complex numbers).\<close>
theory Angles
imports More_Transcendental Canonical_Angle More_Complex
begin
(* ---------------------------------------------------------------------------- *)
subsubsection \<open>Oriented angle\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>Oriented angle between two vectors (it is always in the interval $(-\pi, \pi]$).\<close>
definition ang_vec ("\<angle>") where
[simp]: "\<angle> z1 z2 \<equiv> \<downharpoonright>arg z2 - arg z1\<downharpoonleft>"
lemma ang_vec_bounded:
shows "-pi < \<angle> z1 z2 \<and> \<angle> z1 z2 \<le> pi"
by (simp add: canon_ang(1) canon_ang(2))
lemma ang_vec_sym:
assumes "\<angle> z1 z2 \<noteq> pi"
shows "\<angle> z1 z2 = - \<angle> z2 z1"
using assms
unfolding ang_vec_def
using canon_ang_uminus[of "arg z2 - arg z1"]
by simp
lemma ang_vec_sym_pi:
assumes "\<angle> z1 z2 = pi"
shows "\<angle> z1 z2 = \<angle> z2 z1"
using assms
unfolding ang_vec_def
using canon_ang_uminus_pi[of "arg z2 - arg z1"]
by simp
lemma ang_vec_plus_pi1:
assumes "\<angle> z1 z2 > 0"
shows "\<downharpoonright>\<angle> z1 z2 + pi\<downharpoonleft> = \<angle> z1 z2 - pi"
proof (rule canon_ang_eqI)
show "\<exists> x::int. \<angle> z1 z2 - pi - (\<angle> z1 z2 + pi) = 2 * real_of_int x * pi"
by (rule_tac x="-1" in exI) auto
next
show "- pi < \<angle> z1 z2 - pi \<and> \<angle> z1 z2 - pi \<le> pi"
using assms
unfolding ang_vec_def
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
by auto
qed
lemma ang_vec_plus_pi2:
assumes "\<angle> z1 z2 \<le> 0"
shows "\<downharpoonright>\<angle> z1 z2 + pi\<downharpoonleft> = \<angle> z1 z2 + pi"
proof (rule canon_ang_id)
show "- pi < \<angle> z1 z2 + pi \<and> \<angle> z1 z2 + pi \<le> pi"
using assms
unfolding ang_vec_def
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
by auto
qed
lemma ang_vec_opposite1:
assumes "z1 \<noteq> 0"
shows "\<angle> (-z1) z2 = \<downharpoonright>\<angle> z1 z2 - pi\<downharpoonleft>"
proof-
have "\<angle> (-z1) z2 = \<downharpoonright>arg z2 - (arg z1 + pi)\<downharpoonleft>"
unfolding ang_vec_def
using arg_uminus[OF assms]
using canon_ang_arg[of z2, symmetric]
using canon_ang_diff[of "arg z2" "arg z1 + pi", symmetric]
by simp
moreover
have "\<downharpoonright>\<angle> z1 z2 - pi\<downharpoonleft> = \<downharpoonright>arg z2 - arg z1 - pi\<downharpoonleft>"
using canon_ang_id[of pi, symmetric]
using canon_ang_diff[of "arg z2 - arg z1" "pi", symmetric]
by simp_all
ultimately
show ?thesis
by (simp add: field_simps)
qed
lemma ang_vec_opposite2:
assumes "z2 \<noteq> 0"
shows "\<angle> z1 (-z2) = \<downharpoonright>\<angle> z1 z2 + pi\<downharpoonleft>"
unfolding ang_vec_def
using arg_mult[of "-1" "z2"] assms
using arg_complex_of_real_negative[of "-1"]
using canon_ang_diff[of "arg (-1) + arg z2" "arg z1", symmetric]
using canon_ang_sum[of "arg z2 - arg z1" "pi", symmetric]
using canon_ang_id[of pi] canon_ang_arg[of z1]
by (auto simp: algebra_simps)
lemma ang_vec_opposite_opposite:
assumes "z1 \<noteq> 0" and "z2 \<noteq> 0"
shows "\<angle> (-z1) (-z2) = \<angle> z1 z2"
proof-
have "\<angle> (-z1) (-z2) = \<downharpoonright>\<downharpoonright>\<angle> z1 z2 + pi\<downharpoonleft> - \<downharpoonright>pi\<downharpoonleft>\<downharpoonleft>"
using ang_vec_opposite1[OF assms(1)]
using ang_vec_opposite2[OF assms(2)]
using canon_ang_id[of pi, symmetric]
by simp_all
also have "... = \<downharpoonright>\<angle> z1 z2\<downharpoonleft>"
by (subst canon_ang_diff[symmetric], simp)
finally
show ?thesis
by (metis ang_vec_def canon_ang(1) canon_ang(2) canon_ang_id)
qed
lemma ang_vec_opposite_opposite':
assumes "z1 \<noteq> z" and "z2 \<noteq> z"
shows "\<angle> (z - z1) (z - z2) = \<angle> (z1 - z) (z2 - z)"
using ang_vec_opposite_opposite[of "z - z1" "z - z2"] assms
by (simp add: field_simps del: ang_vec_def)
text \<open>Cosine, scalar product and the law of cosines\<close>
lemma cos_cmod_scalprod:
shows "cmod z1 * cmod z2 * (cos (\<angle> z1 z2)) = Re (scalprod z1 z2)"
proof (cases "z1 = 0 \<or> z2 = 0")
case True
thus ?thesis
by auto
next
case False
thus ?thesis
by (simp add: cos_diff cos_arg sin_arg field_simps)
qed
lemma cos0_scalprod0:
assumes "z1 \<noteq> 0" and "z2 \<noteq> 0"
shows "cos (\<angle> z1 z2) = 0 \<longleftrightarrow> scalprod z1 z2 = 0"
using assms
using cnj_mix_real[of z1 z2]
using cos_cmod_scalprod[of z1 z2]
by (auto simp add: complex_eq_if_Re_eq)
lemma ortho_scalprod0:
assumes "z1 \<noteq> 0" and "z2 \<noteq> 0"
shows "\<angle> z1 z2 = pi/2 \<or> \<angle> z1 z2 = -pi/2 \<longleftrightarrow> scalprod z1 z2 = 0"
using cos0_scalprod0[OF assms]
using ang_vec_bounded[of z1 z2]
using cos_0_iff_canon[of "\<angle> z1 z2"]
by (metis cos_minus cos_pi_half divide_minus_left)
lemma law_of_cosines:
shows "(cdist B C)\<^sup>2 = (cdist A C)\<^sup>2 + (cdist A B)\<^sup>2 - 2*(cdist A C)*(cdist A B)*(cos (\<angle> (C-A) (B-A)))"
proof-
let ?a = "C-B" and ?b = "C-A" and ?c = "B-A"
have "?a = ?b - ?c"
by simp
hence "(cmod ?a)\<^sup>2 = (cmod (?b - ?c))\<^sup>2"
by metis
also have "... = Re (scalprod (?b-?c) (?b-?c))"
by (simp add: cmod_square)
also have "... = (cmod ?b)\<^sup>2 + (cmod ?c)\<^sup>2 - 2*Re (scalprod ?b ?c)"
by (simp add: cmod_square field_simps)
finally
show ?thesis
using cos_cmod_scalprod[of ?b ?c]
by simp
qed
(* ---------------------------------------------------------------------------- *)
subsubsection \<open>Unoriented angle\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>Convex unoriented angle between two vectors (it is always in the interval $[0, pi]$).\<close>
definition ang_vec_c ("\<angle>c") where
[simp]:"\<angle>c z1 z2 \<equiv> abs (\<angle> z1 z2)"
lemma ang_vec_c_sym:
shows "\<angle>c z1 z2 = \<angle>c z2 z1"
unfolding ang_vec_c_def
using ang_vec_sym_pi[of z1 z2] ang_vec_sym[of z1 z2]
by (cases "\<angle> z1 z2 = pi") auto
lemma ang_vec_c_bounded: "0 \<le> \<angle>c z1 z2 \<and> \<angle>c z1 z2 \<le> pi"
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
by auto
text \<open>Cosine and scalar product\<close>
lemma cos_c_: "cos (\<angle>c z1 z2) = cos (\<angle> z1 z2)"
unfolding ang_vec_c_def
by (smt cos_minus)
lemma ortho_c_scalprod0:
assumes "z1 \<noteq> 0" and "z2 \<noteq> 0"
shows "\<angle>c z1 z2 = pi/2 \<longleftrightarrow> scalprod z1 z2 = 0"
proof-
have "\<angle> z1 z2 = pi / 2 \<or> \<angle> z1 z2 = - pi / 2 \<longleftrightarrow> \<angle>c z1 z2 = pi/2"
unfolding ang_vec_c_def
using arctan
by force
thus ?thesis
using ortho_scalprod0[OF assms]
by simp
qed
(* ---------------------------------------------------------------------------- *)
subsubsection \<open>Acute angle\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>Acute or right angle (non-obtuse) between two vectors (it is always in the interval $[0, \frac{\pi}{2}$]).
We will use this to measure angle between two circles, since it can always be acute (or right).\<close>
definition acute_ang where
[simp]: "acute_ang \<alpha> = (if \<alpha> > pi / 2 then pi - \<alpha> else \<alpha>)"
definition ang_vec_a ("\<angle>a") where
[simp]: "\<angle>a z1 z2 \<equiv> acute_ang (\<angle>c z1 z2)"
lemma ang_vec_a_sym:
"\<angle>a z1 z2 = \<angle>a z2 z1"
unfolding ang_vec_a_def
using ang_vec_c_sym
by auto
lemma ang_vec_a_opposite2:
"\<angle>a z1 z2 = \<angle>a z1 (-z2)"
proof(cases "z2 = 0")
case True
thus ?thesis
by (metis minus_zero)
next
case False
thus ?thesis
proof(cases "\<angle> z1 z2 < -pi / 2")
case True
hence "\<angle> z1 z2 < 0"
using pi_not_less_zero
by linarith
have "\<angle>a z1 z2 = pi + \<angle> z1 z2"
using True \<open>\<angle> z1 z2 < 0\<close>
unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def
by auto
moreover
have "\<angle>a z1 (-z2) = pi + \<angle> z1 z2"
unfolding ang_vec_a_def ang_vec_c_def abs_real_def
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
using ang_vec_plus_pi2[of z1 z2] True \<open>\<angle> z1 z2 < 0\<close> \<open>z2 \<noteq> 0\<close>
using ang_vec_opposite2[of z2 z1]
by auto
ultimately
show ?thesis
by auto
next
case False
show ?thesis
proof (cases "\<angle> z1 z2 \<le> 0")
case True
have "\<angle>a z1 z2 = - \<angle> z1 z2"
using \<open>\<not> \<angle> z1 z2 < - pi / 2\<close> True
unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def
by auto
moreover
have "\<angle>a z1 (-z2) = - \<angle> z1 z2"
using \<open>\<not> \<angle> z1 z2 < - pi / 2\<close> True
unfolding ang_vec_a_def ang_vec_c_def abs_real_def
using ang_vec_plus_pi2[of z1 z2]
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
using \<open>z2 \<noteq> 0\<close> ang_vec_opposite2[of z2 z1]
by auto
ultimately
show ?thesis
by simp
next
case False
show ?thesis
proof (cases "\<angle> z1 z2 < pi / 2")
case True
have "\<angle>a z1 z2 = \<angle> z1 z2"
using \<open>\<not> \<angle> z1 z2 \<le> 0\<close> True
unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def
by auto
moreover
have "\<angle>a z1 (-z2) = \<angle> z1 z2"
using \<open>\<not> \<angle> z1 z2 \<le> 0\<close> True
unfolding ang_vec_a_def ang_vec_c_def abs_real_def
using ang_vec_plus_pi1[of z1 z2]
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
using \<open>z2 \<noteq> 0\<close> ang_vec_opposite2[of z2 z1]
by auto
ultimately
show ?thesis
by simp
next
case False
have "\<angle> z1 z2 > 0"
using False
by (metis less_linear less_trans pi_half_gt_zero)
have "\<angle>a z1 z2 = pi - \<angle> z1 z2"
using False \<open>\<angle> z1 z2 > 0\<close>
unfolding ang_vec_a_def ang_vec_c_def ang_vec_a_def abs_real_def
by auto
moreover
have "\<angle>a z1 (-z2) = pi - \<angle> z1 z2"
unfolding ang_vec_a_def ang_vec_c_def abs_real_def
using False \<open>\<angle> z1 z2 > 0\<close>
using ang_vec_plus_pi1[of z1 z2]
using canon_ang(1)[of "arg z2 - arg z1"] canon_ang(2)[of "arg z2 - arg z1"]
using \<open>z2 \<noteq> 0\<close> ang_vec_opposite2[of z2 z1]
by auto
ultimately
show ?thesis
by auto
qed
qed
qed
qed
lemma ang_vec_a_opposite1:
shows "\<angle>a z1 z2 = \<angle>a (-z1) z2"
using ang_vec_a_sym[of "-z1" z2] ang_vec_a_opposite2[of z2 z1] ang_vec_a_sym[of z2 z1]
by auto
lemma ang_vec_a_scale1:
assumes "k \<noteq> 0"
shows "\<angle>a (cor k * z1) z2 = \<angle>a z1 z2"
proof (cases "k > 0")
case True
thus ?thesis
unfolding ang_vec_a_def ang_vec_c_def ang_vec_def
using arg_mult_real_positive[of k z1]
by auto
next
case False
hence "k < 0"
using assms
by auto
thus ?thesis
using arg_mult_real_negative[of k z1]
using ang_vec_a_opposite1[of z1 z2]
unfolding ang_vec_a_def ang_vec_c_def ang_vec_def
by simp
qed
lemma ang_vec_a_scale2:
assumes "k \<noteq> 0"
shows "\<angle>a z1 (cor k * z2) = \<angle>a z1 z2"
using ang_vec_a_sym[of z1 "complex_of_real k * z2"]
using ang_vec_a_scale1[OF assms, of z2 z1]
using ang_vec_a_sym[of z1 z2]
by auto
lemma ang_vec_a_scale:
assumes "k1 \<noteq> 0" and "k2 \<noteq> 0"
shows "\<angle>a (cor k1 * z1) (cor k2 * z2) = \<angle>a z1 z2"
using ang_vec_a_scale1[OF assms(1)] ang_vec_a_scale2[OF assms(2)]
by auto
lemma ang_a_cnj_cnj:
shows "\<angle>a z1 z2 = \<angle>a (cnj z1) (cnj z2)"
unfolding ang_vec_a_def ang_vec_c_def ang_vec_def
proof(cases "arg z1 \<noteq> pi \<and> arg z2 \<noteq> pi")
case True
thus "acute_ang \<bar>\<downharpoonright>arg z2 - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg (cnj z2) - arg (cnj z1)\<downharpoonleft>\<bar>"
using arg_cnj_not_pi[of z1] arg_cnj_not_pi[of z2]
apply (auto simp del:acute_ang_def)
proof(cases "\<downharpoonright>arg z2 - arg z1\<downharpoonleft> = pi")
case True
thus "acute_ang \<bar>\<downharpoonright>arg z2 - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg z1 - arg z2\<downharpoonleft>\<bar>"
using canon_ang_uminus_pi[of "arg z2 - arg z1"]
by (auto simp add:field_simps)
next
case False
thus "acute_ang \<bar>\<downharpoonright>arg z2 - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg z1 - arg z2\<downharpoonleft>\<bar>"
using canon_ang_uminus[of "arg z2 - arg z1"]
by (auto simp add:field_simps)
qed
next
case False
thus "acute_ang \<bar>\<downharpoonright>arg z2 - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg (cnj z2) - arg (cnj z1)\<downharpoonleft>\<bar>"
proof(cases "arg z1 = pi")
case False
hence "arg z2 = pi"
using \<open> \<not> (arg z1 \<noteq> pi \<and> arg z2 \<noteq> pi)\<close>
by auto
thus ?thesis
using False
using arg_cnj_not_pi[of z1] arg_cnj_pi[of z2]
apply (auto simp del:acute_ang_def)
proof(cases "arg z1 > 0")
case True
hence "-arg z1 \<le> 0"
by auto
thus "acute_ang \<bar>\<downharpoonright>pi - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>pi + arg z1\<downharpoonleft>\<bar>"
using True canon_ang_plus_pi1[of "arg z1"]
using arg_bounded[of z1] canon_ang_plus_pi2[of "-arg z1"]
by (auto simp add:field_simps)
next
case False
hence "-arg z1 \<ge> 0"
by simp
thus "acute_ang \<bar>\<downharpoonright>pi - arg z1\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>pi + arg z1\<downharpoonleft>\<bar>"
proof(cases "arg z1 = 0")
case True
thus ?thesis
by (auto simp del:acute_ang_def)
next
case False
hence "-arg z1 > 0"
using \<open>-arg z1 \<ge> 0\<close>
by auto
thus ?thesis
using False canon_ang_plus_pi1[of "-arg z1"]
using arg_bounded[of z1] canon_ang_plus_pi2[of "arg z1"]
by (auto simp add:field_simps)
qed
qed
next
case True
thus ?thesis
using arg_cnj_pi[of z1]
apply (auto simp del:acute_ang_def)
proof(cases "arg z2 = pi")
case True
thus "acute_ang \<bar>\<downharpoonright>arg z2 - pi\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg (cnj z2) - pi\<downharpoonleft>\<bar>"
using arg_cnj_pi[of z2]
by auto
next
case False
thus "acute_ang \<bar>\<downharpoonright>arg z2 - pi\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>arg (cnj z2) - pi\<downharpoonleft>\<bar>"
using arg_cnj_not_pi[of z2]
apply (auto simp del:acute_ang_def)
proof(cases "arg z2 > 0")
case True
hence "-arg z2 \<le> 0"
by auto
thus "acute_ang \<bar>\<downharpoonright>arg z2 - pi\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>- arg z2 - pi\<downharpoonleft>\<bar>"
using True canon_ang_minus_pi1[of "arg z2"]
using arg_bounded[of z2] canon_ang_minus_pi2[of "-arg z2"]
by (auto simp add: field_simps)
next
case False
hence "-arg z2 \<ge> 0"
by simp
thus "acute_ang \<bar>\<downharpoonright>arg z2 - pi\<downharpoonleft>\<bar> = acute_ang \<bar>\<downharpoonright>- arg z2 - pi\<downharpoonleft>\<bar>"
proof(cases "arg z2 = 0")
case True
thus ?thesis
by (auto simp del:acute_ang_def)
next
case False
hence "-arg z2 > 0"
using \<open>-arg z2 \<ge> 0\<close>
by auto
thus ?thesis
using False canon_ang_minus_pi1[of "-arg z2"]
using arg_bounded[of z2] canon_ang_minus_pi2[of "arg z2"]
by (auto simp add:field_simps)
qed
qed
qed
qed
qed
text \<open>Cosine and scalar product\<close>
lemma ortho_a_scalprod0:
assumes "z1 \<noteq> 0" and "z2 \<noteq> 0"
shows "\<angle>a z1 z2 = pi/2 \<longleftrightarrow> scalprod z1 z2 = 0"
unfolding ang_vec_a_def
using assms ortho_c_scalprod0[of z1 z2]
by auto
declare ang_vec_c_def[simp del]
lemma cos_a_c: "cos (\<angle>a z1 z2) = abs (cos (\<angle>c z1 z2))"
proof-
have "0 \<le> \<angle>c z1 z2" "\<angle>c z1 z2 \<le> pi"
using ang_vec_c_bounded[of z1 z2]
by auto
show ?thesis
proof (cases "\<angle>c z1 z2 = pi/2")
case True
thus ?thesis
unfolding ang_vec_a_def acute_ang_def
by (smt cos_pi_half pi_def pi_half)
next
case False
show ?thesis
proof (cases "\<angle>c z1 z2 < pi / 2")
case True
thus ?thesis
- using `0 \<le> \<angle>c z1 z2`
+ using \<open>0 \<le> \<angle>c z1 z2\<close>
using cos_gt_zero_pi[of "\<angle>c z1 z2"]
unfolding ang_vec_a_def
by simp
next
case False
hence "\<angle>c z1 z2 > pi/2"
- using `\<angle>c z1 z2 \<noteq> pi/2`
+ using \<open>\<angle>c z1 z2 \<noteq> pi/2\<close>
by simp
hence "cos (\<angle>c z1 z2) < 0"
- using `\<angle>c z1 z2 \<le> pi`
+ using \<open>\<angle>c z1 z2 \<le> pi\<close>
using cos_lt_zero_on_pi2_pi[of "\<angle>c z1 z2"]
by simp
thus ?thesis
- using `\<angle>c z1 z2 > pi/2`
+ using \<open>\<angle>c z1 z2 > pi/2\<close>
unfolding ang_vec_a_def
by simp
qed
qed
qed
end
diff --git a/thys/Complex_Geometry/Linear_Systems.thy b/thys/Complex_Geometry/Linear_Systems.thy
--- a/thys/Complex_Geometry/Linear_Systems.thy
+++ b/thys/Complex_Geometry/Linear_Systems.thy
@@ -1,217 +1,217 @@
(* ---------------------------------------------------------------------------- *)
subsection \<open>Systems of linear equations\<close>
(* ---------------------------------------------------------------------------- *)
(* TODO: merge with matrices *)
text \<open>In this section some simple properties of systems of linear equations with two or three unknowns are derived.
Existence and uniqueness of solutions of regular and singular homogenous and non-homogenous systems is characterized.\<close>
theory Linear_Systems
imports Main
begin
text \<open>Determinant of 2x2 matrix\<close>
definition det2 :: "('a::field) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
[simp]: "det2 a11 a12 a21 a22 \<equiv> a11*a22 - a12*a21"
text \<open>Regular homogenous system has only trivial solution\<close>
lemma regular_homogenous_system:
fixes a11 a12 a21 a22 x1 x2 :: "'a::field"
assumes "det2 a11 a12 a21 a22 \<noteq> 0"
assumes "a11*x1 + a12*x2 = 0" and
"a21*x1 + a22*x2 = 0"
shows "x1 = 0 \<and> x2 = 0"
proof (cases "a11 = 0")
case True
with assms(1) have "a12 \<noteq> 0" "a21 \<noteq> 0"
by auto
thus ?thesis
using \<open>a11 = 0\<close> assms(2) assms(3)
by auto
next
case False
hence "x1 = - a12*x2 / a11"
using assms(2)
by (metis eq_neg_iff_add_eq_0 mult_minus_left nonzero_mult_div_cancel_left)
hence "a21 * (- a12 * x2 / a11) + a22 * x2 = 0"
using assms(3)
by simp
hence "a21 * (- a12 * x2) + a22 * x2 * a11 = 0"
using \<open>a11 \<noteq> 0\<close>
by auto
hence "(a11*a22 - a12*a21)*x2 = 0"
by (simp add: field_simps)
thus ?thesis
using assms(1) assms(2) \<open>a11 \<noteq> 0\<close>
by auto
qed
text \<open>Regular system has a unique solution\<close>
lemma regular_system:
fixes a11 a12 a21 a22 b1 b2 :: "'a::field"
assumes "det2 a11 a12 a21 a22 \<noteq> 0"
shows "\<exists>! x. a11*(fst x) + a12*(snd x) = b1 \<and>
a21*(fst x) + a22*(snd x) = b2"
proof
let ?d = "a11*a22 - a12*a21" and ?d1 = "b1*a22 - b2*a12" and ?d2 = "b2*a11 - b1*a21"
let ?x = "(?d1 / ?d, ?d2 / ?d)"
have "a11 * ?d1 + a12 * ?d2 = b1*?d" "a21 * ?d1 + a22 * ?d2 = b2*?d"
by (auto simp add: field_simps)
thus "a11 * fst ?x + a12 * snd ?x = b1 \<and> a21 * fst ?x + a22 * snd ?x = b2"
using assms
by (metis (hide_lams, no_types) det2_def add_divide_distrib eq_divide_imp fst_eqD snd_eqD times_divide_eq_right)
fix x'
assume "a11 * fst x' + a12 * snd x' = b1 \<and> a21 * fst x' + a22 * snd x' = b2"
with \<open>a11 * fst ?x + a12 * snd ?x = b1 \<and> a21 * fst ?x + a22 * snd ?x = b2\<close>
have "a11 * (fst x' - fst ?x) + a12 * (snd x' - snd ?x) = 0 \<and> a21 * (fst x' - fst ?x) + a22 * (snd x' - snd ?x) = 0"
by (auto simp add: field_simps)
thus "x' = ?x"
using regular_homogenous_system[OF assms, of "fst x' - fst ?x" "snd x' - snd ?x"]
by (cases x') auto
qed
text \<open>Singular system does not have a unique solution\<close>
lemma singular_system:
fixes a11 a12 a21 a22 ::"'a::field"
assumes "det2 a11 a12 a21 a22 = 0" and "a11 \<noteq> 0 \<or> a12 \<noteq> 0"
assumes x0: "a11*fst x0 + a12*snd x0 = b1"
"a21*fst x0 + a22*snd x0 = b2"
assumes x: "a11*fst x + a12*snd x = b1"
shows "a21*fst x + a22*snd x = b2"
proof (cases "a11 = 0")
case True
with assms have "a21 = 0" "a12 \<noteq> 0"
by auto
let ?k = "a22 / a12"
have "b2 = ?k * b1"
using x0 \<open>a11 = 0\<close> \<open>a21 = 0\<close> \<open>a12 \<noteq> 0\<close>
by auto
thus ?thesis
using \<open>a11 = 0\<close> \<open>a21 = 0\<close> \<open>a12 \<noteq> 0\<close> x
by auto
next
case False
let ?k = "a21 / a11"
from x
have "?k * a11 * fst x + ?k * a12 * snd x = ?k * b1"
using \<open>a11 \<noteq> 0\<close>
by (auto simp add: field_simps)
moreover
have "a21 = ?k * a11" "a22 = ?k * a12" "b2 = ?k * b1"
using assms(1) x0 \<open>a11 \<noteq> 0\<close>
by (auto simp add: field_simps)
ultimately
show ?thesis
by auto
qed
text \<open>All solutions of a homogenous system of 2 equations with 3 unknows are proportional\<close>
lemma linear_system_homogenous_3_2:
fixes a11 a12 a13 a21 a22 a23 x1 y1 z1 x2 y2 z2 :: "'a::field"
assumes "f1 = (\<lambda> x y z. a11 * x + a12 * y + a13 * z)"
assumes "f2 = (\<lambda> x y z. a21 * x + a22 * y + a23 * z)"
assumes "f1 x1 y1 z1 = 0" and "f2 x1 y1 z1 = 0"
assumes "f1 x2 y2 z2 = 0" and "f2 x2 y2 z2 = 0"
assumes "x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0"
assumes "det2 a11 a12 a21 a22 \<noteq> 0 \<or> det2 a11 a13 a21 a23 \<noteq> 0 \<or> det2 a12 a13 a22 a23 \<noteq> 0"
shows "\<exists> k. x1 = k * x2 \<and> y1 = k * y2 \<and> z1 = k * z2"
proof-
let ?Dz = "det2 a11 a12 a21 a22"
let ?Dy = "det2 a11 a13 a21 a23"
let ?Dx = "det2 a12 a13 a22 a23"
have "a21 * (f1 x1 y1 z1) - a11 * (f2 x1 y1 z1) = 0"
using assms
by simp
hence yz1: "?Dz*y1 + ?Dy*z1 = 0"
using assms
by (simp add: field_simps)
have "a21 * (f1 x2 y2 z2) - a11 * (f2 x2 y2 z2) = 0"
using assms
by simp
hence yz2: "?Dz*y2 + ?Dy*z2 = 0"
using assms
by (simp add: field_simps)
have "a22 * (f1 x1 y1 z1) - a12 * (f2 x1 y1 z1) = 0"
using assms
by simp
hence xz1: "-?Dz*x1 + ?Dx*z1 = 0"
using assms
by (simp add: field_simps)
have "a22 * (f1 x2 y2 z2) - a12 * (f2 x2 y2 z2) = 0"
using assms
by simp
hence xz2: "-?Dz*x2 + ?Dx*z2 = 0"
using assms
by (simp add: field_simps)
have "a23 * (f1 x1 y1 z1) - a13 * (f2 x1 y1 z1) = 0"
using assms
by simp
hence xy1: "?Dy*x1 + ?Dx*y1 = 0"
using assms
by (simp add: field_simps)
have "a23 * (f1 x2 y2 z2) - a13 * (f2 x2 y2 z2) = 0"
using assms
by simp
hence xy2: "?Dy*x2 + ?Dx*y2 = 0"
using assms
by (simp add: field_simps)
show ?thesis
- using `?Dz \<noteq> 0 \<or> ?Dy \<noteq> 0 \<or> ?Dx \<noteq> 0`
+ using \<open>?Dz \<noteq> 0 \<or> ?Dy \<noteq> 0 \<or> ?Dx \<noteq> 0\<close>
proof safe
assume "?Dz \<noteq> 0"
hence *:
"x2 = (?Dx / ?Dz) * z2" "y2 = - (?Dy / ?Dz) * z2"
"x1 = (?Dx / ?Dz) * z1" "y1 = - (?Dy / ?Dz) * z1"
using xy2 xz2 xy1 xz1 yz1 yz2
by (simp_all add: field_simps)
hence "z2 \<noteq> 0"
- using `x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0`
+ using \<open>x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0\<close>
by auto
thus ?thesis
- using * `?Dz \<noteq> 0`
+ using * \<open>?Dz \<noteq> 0\<close>
by (rule_tac x="z1/z2" in exI) auto
next
assume "?Dy \<noteq> 0"
hence *:
"x2 = - (?Dx / ?Dy) * y2" "z2 = - (?Dz / ?Dy) * y2"
"x1 = - (?Dx / ?Dy) * y1" "z1 = - (?Dz / ?Dy) * y1"
using xy2 xz2 xy1 xz1 yz1 yz2
by (simp_all add: field_simps)
hence "y2 \<noteq> 0"
- using `x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0`
+ using \<open>x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0\<close>
by auto
thus ?thesis
- using * `?Dy \<noteq> 0`
+ using * \<open>?Dy \<noteq> 0\<close>
by (rule_tac x="y1/y2" in exI) auto
next
assume "?Dx \<noteq> 0"
hence *:
"y2 = - (?Dy / ?Dx) * x2" "z2 = (?Dz / ?Dx) * x2"
"y1 = - (?Dy / ?Dx) * x1" "z1 = (?Dz / ?Dx) * x1"
using xy2 xz2 xy1 xz1 yz1 yz2
by (simp_all add: field_simps)
hence "x2 \<noteq> 0"
- using `x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0`
+ using \<open>x2 \<noteq> 0 \<or> y2 \<noteq> 0 \<or> z2 \<noteq> 0\<close>
by auto
thus ?thesis
- using * `?Dx \<noteq> 0`
+ using * \<open>?Dx \<noteq> 0\<close>
by (rule_tac x="x1/x2" in exI) auto
qed
qed
end
diff --git a/thys/Complex_Geometry/Oriented_Circlines.thy b/thys/Complex_Geometry/Oriented_Circlines.thy
--- a/thys/Complex_Geometry/Oriented_Circlines.thy
+++ b/thys/Complex_Geometry/Oriented_Circlines.thy
@@ -1,1373 +1,1373 @@
(* -------------------------------------------------------------------------- *)
section \<open>Oriented circlines\<close>
(* -------------------------------------------------------------------------- *)
theory Oriented_Circlines
imports Circlines
begin
(* ----------------------------------------------------------------- *)
subsection \<open>Oriented circlines definition\<close>
(* ----------------------------------------------------------------- *)
text \<open>In this section we describe how the orientation is introduced for the circlines. Similarly as
the set of circline points, the set of disc points is introduced using the quadratic form induced by
the circline matrix --- the set of points of the circline disc is the set of points such that
satisfy that $A\cdot z\cdot \overline{z} + B\cdot \overline{z} + C\cdot z + D < 0$, where
$(A, B, C, D)$ is a circline matrix representative Hermitean matrix. As the
set of disc points must be invariant to the choice of representative, it is clear that oriented
circlines matrices are equivalent only if they are proportional by a positive real factor (recall
that unoriented circline allowed arbitrary non-zero real factors).\<close>
definition ocircline_eq_cmat :: "complex_mat \<Rightarrow> complex_mat \<Rightarrow> bool" where
[simp]: "ocircline_eq_cmat A B \<longleftrightarrow>(\<exists> k::real. k > 0 \<and> B = cor k *\<^sub>s\<^sub>m A)"
lift_definition ocircline_eq_clmat :: "circline_mat \<Rightarrow> circline_mat \<Rightarrow> bool" is ocircline_eq_cmat
done
lemma ocircline_eq_cmat_id [simp]:
shows "ocircline_eq_cmat H H"
by (simp, rule_tac x=1 in exI, simp)
quotient_type ocircline = circline_mat / ocircline_eq_clmat
proof (rule equivpI)
show "reflp ocircline_eq_clmat"
unfolding reflp_def
by transfer (auto, rule_tac x="1" in exI, simp)
next
show "symp ocircline_eq_clmat"
unfolding symp_def
by transfer (simp only: ocircline_eq_cmat_def, safe, rule_tac x="1/k" in exI, simp)
next
show "transp ocircline_eq_clmat"
unfolding transp_def
by transfer (simp only: ocircline_eq_cmat_def, safe, rule_tac x="k*ka" in exI, simp)
qed
(* ----------------------------------------------------------------- *)
subsection \<open>Points on oriented circlines\<close>
(* ----------------------------------------------------------------- *)
text \<open>Boundary of the circline.\<close>
lift_definition on_ocircline :: "ocircline \<Rightarrow> complex_homo \<Rightarrow> bool" is on_circline_clmat_hcoords
by transfer (simp del: quad_form_def, (erule exE)+, simp add: quad_form_scale_m quad_form_scale_v del: quad_form_def)
definition ocircline_set :: "ocircline \<Rightarrow> complex_homo set" where
"ocircline_set H = {z. on_ocircline H z}"
lemma ocircline_set_I [simp]:
assumes "on_ocircline H z"
shows "z \<in> ocircline_set H"
using assms
unfolding ocircline_set_def
by simp
(* ----------------------------------------------------------------- *)
subsection \<open>Disc and disc complement - in and out points\<close>
(* ----------------------------------------------------------------- *)
text \<open>Interior and the exterior of an oriented circline.\<close>
definition in_ocircline_cmat_cvec :: "complex_mat \<Rightarrow> complex_vec \<Rightarrow> bool" where
[simp]: "in_ocircline_cmat_cvec H z \<longleftrightarrow> Re (quad_form z H) < 0"
lift_definition in_ocircline_clmat_hcoords :: "circline_mat \<Rightarrow> complex_homo_coords \<Rightarrow> bool" is in_ocircline_cmat_cvec
done
lift_definition in_ocircline :: "ocircline \<Rightarrow> complex_homo \<Rightarrow> bool" is in_ocircline_clmat_hcoords
proof transfer
fix H H' z z'
assume hh: "hermitean H \<and> H \<noteq> mat_zero" and "hermitean H' \<and> H' \<noteq> mat_zero" and
"z \<noteq> vec_zero" and "z' \<noteq> vec_zero"
assume "ocircline_eq_cmat H H'" and "z \<approx>\<^sub>v z'"
then obtain k k' where
*: "0 < k" "H' = cor k *\<^sub>s\<^sub>m H" "k' \<noteq> 0" "z' = k' *\<^sub>s\<^sub>v z"
by auto
hence "quad_form z' H' = cor k * cor ((cmod k')\<^sup>2) * quad_form z H"
by (simp add: quad_form_scale_v quad_form_scale_m del: vec_cnj_sv quad_form_def)
hence "Re (quad_form z' H') = k * (cmod k')\<^sup>2 * Re (quad_form z H)"
using hh quad_form_hermitean_real[of H]
by (simp add: power2_eq_square)
thus "in_ocircline_cmat_cvec H z = in_ocircline_cmat_cvec H' z'"
using \<open>k > 0\<close> \<open>k' \<noteq> 0\<close>
using mult_less_0_iff
by fastforce
qed
definition disc :: "ocircline \<Rightarrow> complex_homo set" where
"disc H = {z. in_ocircline H z}"
lemma disc_I [simp]:
assumes "in_ocircline H z"
shows "z \<in> disc H"
using assms
unfolding disc_def
by simp
definition out_ocircline_cmat_cvec :: "complex_mat \<Rightarrow> complex_vec \<Rightarrow> bool" where
[simp]: "out_ocircline_cmat_cvec H z \<longleftrightarrow> Re (quad_form z H) > 0"
lift_definition out_ocircline_clmat_hcoords :: "circline_mat \<Rightarrow> complex_homo_coords \<Rightarrow> bool" is out_ocircline_cmat_cvec
done
lift_definition out_ocircline :: "ocircline \<Rightarrow> complex_homo \<Rightarrow> bool" is out_ocircline_clmat_hcoords
proof transfer
fix H H' z z'
assume hh: "hermitean H \<and> H \<noteq> mat_zero" "hermitean H' \<and> H' \<noteq> mat_zero"
"z \<noteq> vec_zero" "z' \<noteq> vec_zero"
assume "ocircline_eq_cmat H H'" "z \<approx>\<^sub>v z'"
then obtain k k' where
*: "0 < k" "H' = cor k *\<^sub>s\<^sub>m H" "k' \<noteq> 0" "z' = k' *\<^sub>s\<^sub>v z"
by auto
hence "quad_form z' H' = cor k * cor ((cmod k')\<^sup>2) * quad_form z H"
by (simp add: quad_form_scale_v quad_form_scale_m del: vec_cnj_sv quad_form_def)
hence "Re (quad_form z' H') = k * (cmod k')\<^sup>2 * Re (quad_form z H)"
using hh quad_form_hermitean_real[of H]
by (simp add: power2_eq_square)
thus "out_ocircline_cmat_cvec H z = out_ocircline_cmat_cvec H' z'"
using \<open>k > 0\<close> \<open>k' \<noteq> 0\<close>
using zero_less_mult_pos
by fastforce
qed
definition disc_compl :: "ocircline \<Rightarrow> complex_homo set" where
"disc_compl H = {z. out_ocircline H z}"
text \<open>These three sets are mutually disjoint and they fill up the entire plane.\<close>
lemma disc_compl_I [simp]:
assumes "out_ocircline H z"
shows "z \<in> disc_compl H"
using assms
unfolding disc_compl_def
by simp
lemma in_on_out:
shows "in_ocircline H z \<or> on_ocircline H z \<or> out_ocircline H z"
apply (transfer, transfer)
using quad_form_hermitean_real
using complex_eq_if_Re_eq
by auto
lemma in_on_out_univ:
shows "disc H \<union> disc_compl H \<union> ocircline_set H = UNIV"
unfolding disc_def disc_compl_def ocircline_set_def
using in_on_out[of H]
by auto
lemma disc_inter_disc_compl [simp]:
shows "disc H \<inter> disc_compl H = {}"
unfolding disc_def disc_compl_def
by auto (transfer, transfer, simp)
lemma disc_inter_ocircline_set [simp]:
shows "disc H \<inter> ocircline_set H = {}"
unfolding disc_def ocircline_set_def
by auto (transfer, transfer, simp)
lemma disc_compl_inter_ocircline_set [simp]:
shows "disc_compl H \<inter> ocircline_set H = {}"
unfolding disc_compl_def ocircline_set_def
by auto (transfer, transfer, simp)
(* ----------------------------------------------------------------- *)
subsection \<open>Opposite orientation\<close>
(* ----------------------------------------------------------------- *)
text \<open>Finding opposite circline is idempotent, and opposite circlines share the same set of points,
but exchange disc and its complement.\<close>
definition opposite_ocircline_cmat :: "complex_mat \<Rightarrow> complex_mat" where
[simp]: "opposite_ocircline_cmat H = (-1) *\<^sub>s\<^sub>m H"
lift_definition opposite_ocircline_clmat :: "circline_mat \<Rightarrow> circline_mat" is opposite_ocircline_cmat
by (auto simp add: hermitean_def mat_adj_def mat_cnj_def)
lift_definition opposite_ocircline :: "ocircline \<Rightarrow> ocircline" is opposite_ocircline_clmat
by transfer auto
lemma opposite_ocircline_involution [simp]:
shows "opposite_ocircline (opposite_ocircline H) = H"
by (transfer, transfer) (auto, rule_tac x="1" in exI, simp)
lemma on_circline_opposite_ocircline_cmat [simp]:
assumes "hermitean H \<and> H \<noteq> mat_zero" and "z \<noteq> vec_zero"
shows "on_circline_cmat_cvec (opposite_ocircline_cmat H) z = on_circline_cmat_cvec H z"
using assms
by (simp add: quad_form_scale_m del: quad_form_def)
lemma on_circline_opposite_ocircline [simp]:
shows "on_ocircline (opposite_ocircline H) z \<longleftrightarrow> on_ocircline H z"
using on_circline_opposite_ocircline_cmat
by (transfer, transfer, simp)
lemma ocircline_set_opposite_ocircline [simp]:
shows "ocircline_set (opposite_ocircline H) = ocircline_set H"
unfolding ocircline_set_def
by auto
lemma disc_compl_opposite_ocircline [simp]:
shows "disc_compl (opposite_ocircline H) = disc H"
unfolding disc_def disc_compl_def
apply auto
apply (transfer, transfer)
apply (auto simp add: quad_form_scale_m simp del: quad_form_def)
apply (transfer ,transfer)
apply (auto simp add: quad_form_scale_m simp del: quad_form_def)
done
lemma disc_opposite_ocircline [simp]:
shows "disc (opposite_ocircline H) = disc_compl H"
using disc_compl_opposite_ocircline[of "opposite_ocircline H"]
by simp
(* ----------------------------------------------------------------- *)
subsection \<open>Positive orientation. Conversion between unoriented and oriented circlines\<close>
(* ----------------------------------------------------------------- *)
text \<open>Given an oriented circline, one can trivially obtain its unoriented counterpart, and these two
share the same set of points.\<close>
lift_definition of_ocircline :: "ocircline \<Rightarrow> circline" is "id::circline_mat \<Rightarrow> circline_mat"
by transfer (simp, erule exE, force)
lemma of_ocircline_opposite_ocircline [simp]:
shows "of_ocircline (opposite_ocircline H) = of_ocircline H"
by (transfer, transfer) (simp, erule exE, rule_tac x="-1" in exI, simp)
lemma on_ocircline_of_circline [simp]:
shows "on_circline (of_ocircline H) z \<longleftrightarrow> on_ocircline H z"
by (transfer, transfer, simp)
lemma circline_set_of_ocircline [simp]:
shows "circline_set (of_ocircline H) = ocircline_set H"
unfolding ocircline_set_def circline_set_def
by (safe) (transfer, simp)+
lemma inj_of_ocircline:
assumes "of_ocircline H = of_ocircline H'"
shows "H = H' \<or> H = opposite_ocircline H'"
using assms
by (transfer, transfer) (simp, metis linorder_neqE_linordered_idom minus_of_real_eq_of_real_iff mult_minus1 mult_sm_distribution neg_0_equal_iff_equal neg_less_0_iff_less)
lemma inj_ocircline_set:
assumes "ocircline_set H = ocircline_set H'" and "ocircline_set H \<noteq> {}"
shows "H = H' \<or> H = opposite_ocircline H'"
proof-
from assms
have "circline_set (of_ocircline H) = circline_set (of_ocircline H')"
"circline_set (of_ocircline H') \<noteq> {}"
by auto
hence "of_ocircline H = of_ocircline H'"
by (simp add: inj_circline_set)
thus ?thesis
by (rule inj_of_ocircline)
qed
text \<open>Positive orientation.\<close>
text \<open>Given a representative Hermitean matrix of a circline, it represents exactly one of the two
possible oriented circlines. The choice of what should be called a positive orientation is
arbitrary. We follow Schwerdtfeger \cite{schwerdtfeger}, use the leading coefficient $A$ as the
first criterion, and say that circline matrices with $A > 0$ are called positively oriented, and
with $A < 0$ negatively oriented. However, Schwerdtfeger did not discuss the possible case of $A =
0$ (the case of lines), so we had to extend his definition to achieve a total characterization.\<close>
definition pos_oriented_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "pos_oriented_cmat H \<longleftrightarrow>
(let (A, B, C, D) = H
in (Re A > 0 \<or> (Re A = 0 \<and> ((B \<noteq> 0 \<and> arg B > 0) \<or> (B = 0 \<and> Re D > 0)))))"
lift_definition pos_oriented_clmat :: "circline_mat \<Rightarrow> bool" is pos_oriented_cmat
done
lift_definition pos_oriented :: "ocircline \<Rightarrow> bool" is pos_oriented_clmat
by transfer
(case_tac circline_mat1, case_tac circline_mat2, simp, erule exE, simp,
metis mult_pos_pos zero_less_mult_pos)
lemma pos_oriented:
shows "pos_oriented H \<or> pos_oriented (opposite_ocircline H)"
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where HH: "H = (A, B, C, D)"
by (cases H) auto
moreover
hence "Re A = 0 \<and> Re D = 0 \<longrightarrow> B \<noteq> 0"
using hh hermitean_elems[of A B C D]
by (cases A, cases D) (auto simp add: Complex_eq)
moreover
have "B \<noteq> 0 \<and> \<not> 0 < arg B \<longrightarrow> 0 < arg (- B)"
using canon_ang_plus_pi2[of "arg B"] arg_bounded[of B]
by (auto simp add: arg_uminus)
ultimately
show "pos_oriented_cmat H \<or> pos_oriented_cmat (opposite_ocircline_cmat H)"
by auto
qed
lemma pos_oriented_opposite_ocircline_cmat [simp]:
assumes "hermitean H \<and> H \<noteq> mat_zero"
shows "pos_oriented_cmat (opposite_ocircline_cmat H) \<longleftrightarrow> \<not> pos_oriented_cmat H"
proof-
obtain A B C D where HH: "H = (A, B, C, D)"
by (cases H) auto
moreover
hence "Re A = 0 \<and> Re D = 0 \<longrightarrow> B \<noteq> 0"
using assms hermitean_elems[of A B C D]
by (cases A, cases D) (auto simp add: Complex_eq)
moreover
have "B \<noteq> 0 \<and> \<not> 0 < arg B \<longrightarrow> 0 < arg (- B)"
using canon_ang_plus_pi2[of "arg B"] arg_bounded[of B]
by (auto simp add: arg_uminus)
moreover
have "B \<noteq> 0 \<and> 0 < arg B \<longrightarrow> \<not> 0 < arg (- B)"
using canon_ang_plus_pi1[of "arg B"] arg_bounded[of B]
by (auto simp add: arg_uminus)
ultimately
show "pos_oriented_cmat (opposite_ocircline_cmat H) = (\<not> pos_oriented_cmat H)"
by simp (metis not_less_iff_gr_or_eq)
qed
lemma pos_oriented_opposite_ocircline [simp]:
shows "pos_oriented (opposite_ocircline H) \<longleftrightarrow> \<not> pos_oriented H"
using pos_oriented_opposite_ocircline_cmat
by (transfer, transfer, simp)
lemma pos_oriented_circle_inf:
assumes "\<infinity>\<^sub>h \<notin> ocircline_set H"
shows "pos_oriented H \<longleftrightarrow> \<infinity>\<^sub>h \<notin> disc H"
using assms
unfolding ocircline_set_def disc_def
apply simp
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where HH: "H = (A, B, C, D)"
by (cases H) auto
hence "is_real A"
using hh hermitean_elems
by auto
assume "\<not> on_circline_cmat_cvec H \<infinity>\<^sub>v"
thus "pos_oriented_cmat H = (\<not> in_ocircline_cmat_cvec H \<infinity>\<^sub>v)"
using HH \<open>is_real A\<close>
by (cases A) (auto simp add: vec_cnj_def Complex_eq)
qed
lemma pos_oriented_euclidean_circle:
assumes "is_circle (of_ocircline H)"
"(a, r) = euclidean_circle (of_ocircline H)"
"circline_type (of_ocircline H) < 0"
shows "pos_oriented H \<longleftrightarrow> of_complex a \<in> disc H"
using assms
unfolding disc_def
apply simp
proof (transfer, transfer)
fix H a r
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where HH: "H = (A, B, C, D)"
by (cases H) auto
hence "is_real A" "is_real D" "C = cnj B"
using hh hermitean_elems
by auto
assume *: "\<not> circline_A0_cmat (id H)" "(a, r) = euclidean_circle_cmat (id H)" "circline_type_cmat (id H) < 0"
hence "A \<noteq> 0" "Re A \<noteq> 0"
using HH \<open>is_real A\<close>
by (case_tac[!] A) (auto simp add: Complex_eq)
have "Re (A*D - B*C) < 0"
using \<open>circline_type_cmat (id H) < 0\<close> HH
by simp
have **: "(A * (D * cnj A) - B * (C * cnj A)) / (A * cnj A) = (A*D - B*C) / A"
using \<open>A \<noteq> 0\<close>
by (simp add: field_simps)
hence ***: "0 < Re A \<longleftrightarrow> Re ((A * (D * cnj A) - B * (C * cnj A)) / (A * cnj A)) < 0"
using \<open>is_real A\<close> \<open>A \<noteq> 0\<close> \<open>Re (A*D - B*C) < 0\<close>
by (simp add: Re_divide_real divide_less_0_iff)
have "Re D - Re (cnj B * B / cnj A) < Re ((C - cnj B * A / cnj A) * B / A)" if "Re A > 0"
using HH * \<open>is_real A\<close> that
by simp (smt "**" "***" cnj.simps(1) cnj.simps(2) complex_eq diff_divide_distrib left_diff_distrib'
minus_complex.simps(1) mult.commute nonzero_mult_div_cancel_right)?
moreover have "Re A > 0" if "Re D - Re (cnj B * B / cnj A) < Re ((C - cnj B * A / cnj A) * B / A)"
using HH * \<open>is_real A\<close> that
by simp (smt "**" "***" cnj.simps(1) cnj.simps(2) complex_eq diff_divide_distrib left_diff_distrib'
minus_complex.simps(1) mult.commute nonzero_mult_div_cancel_right)?
ultimately show "pos_oriented_cmat H = in_ocircline_cmat_cvec H (of_complex_cvec a)"
using HH \<open>Re A \<noteq> 0\<close> * \<open>is_real A\<close> by (auto simp add: vec_cnj_def)
qed
text \<open>Introduce positive orientation\<close>
definition of_circline_cmat :: "complex_mat \<Rightarrow> complex_mat" where
[simp]: "of_circline_cmat H = (if pos_oriented_cmat H then H else opposite_ocircline_cmat H)"
lift_definition of_circline_clmat :: "circline_mat \<Rightarrow> circline_mat" is of_circline_cmat
by (auto simp add: hermitean_def mat_adj_def mat_cnj_def)
lemma of_circline_clmat_def':
shows "of_circline_clmat H = (if pos_oriented_clmat H then H else opposite_ocircline_clmat H)"
by transfer simp
lemma pos_oriented_cmat_mult_positive':
assumes
"hermitean H1 \<and> H1 \<noteq> mat_zero" and
"hermitean H2 \<and> H2 \<noteq> mat_zero" and
"\<exists>k. k > 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1" and
"pos_oriented_cmat H1"
shows "pos_oriented_cmat H2"
proof-
obtain A1 B1 C1 D1 A2 B2 C2 D2
where HH: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2)
thus ?thesis
using assms
by fastforce
qed
lemma pos_oriented_cmat_mult_positive:
assumes
"hermitean H1 \<and> H1 \<noteq> mat_zero" and
"hermitean H2 \<and> H2 \<noteq> mat_zero" and
"\<exists>k. k > 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
shows
"pos_oriented_cmat H1 \<longleftrightarrow> pos_oriented_cmat H2"
proof-
from assms(3) obtain k where "k > 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
hence "\<exists>k. k > 0 \<and> H1 = cor k *\<^sub>s\<^sub>m H2"
by (rule_tac x="1/k" in exI, auto)
thus ?thesis
using assms pos_oriented_cmat_mult_positive'
by blast
qed
lemma pos_oriented_cmat_mult_negative:
assumes
"hermitean H1 \<and> H1 \<noteq> mat_zero" and
"hermitean H2 \<and> H2 \<noteq> mat_zero" and
"\<exists>k. k < 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
shows
"pos_oriented_cmat H1 \<longleftrightarrow> \<not> pos_oriented_cmat H2"
using assms
proof-
obtain A B C D A1 B1 C1 D1
where *: "H1 = (A, B, C, D)" "H2 = (A1, B1, C1, D1)"
by (cases H1, cases H2) auto
hence **: "is_real A" "is_real D" "is_real A1" "is_real D1" "B = 0 \<longleftrightarrow> C = 0" "B1 = 0 \<longleftrightarrow> C1 = 0"
using assms hermitean_elems[of A B C D] hermitean_elems[of A1 B1 C1 D1]
by auto
show ?thesis
proof (rule iffI)
assume H1: "pos_oriented_cmat H1"
show "\<not> pos_oriented_cmat H2"
proof (cases "Re A > 0")
case True
thus ?thesis
using assms * ** mult_neg_pos
by fastforce
next
case False
show ?thesis
proof (cases "B = 0")
case True
thus ?thesis
- using assms * ** H1 `\<not> Re A > 0` mult_neg_pos
+ using assms * ** H1 \<open>\<not> Re A > 0\<close> mult_neg_pos
by fastforce
next
case False
thus ?thesis
using arg_uminus_opposite_sign[of B] arg_mult_real_negative
- using assms * ** H1 `\<not> Re A > 0` mult_neg_pos
+ using assms * ** H1 \<open>\<not> Re A > 0\<close> mult_neg_pos
by fastforce
qed
qed
next
assume H2: "\<not> pos_oriented_cmat H2"
show "pos_oriented_cmat H1"
proof (cases "Re A > 0")
case True
thus ?thesis
using * ** mult_neg_pos
by fastforce
next
case False
show ?thesis
proof (cases "B = 0")
case True
thus ?thesis
- using assms * ** H2 `\<not> Re A > 0`
+ using assms * ** H2 \<open>\<not> Re A > 0\<close>
by simp (smt arg_0_iff arg_complex_of_real_negative arg_complex_of_real_positive arg_mult_eq complex_of_real_Re mult.right_neutral mult_eq_0_iff of_real_0 of_real_1 zero_complex.simps(1))
next
case False
thus ?thesis
- using assms `\<not> Re A > 0` H2 * **
+ using assms \<open>\<not> Re A > 0\<close> H2 * **
using arg_uminus_opposite_sign[of B]
by (cases "Re A = 0", auto simp add: mult_neg_neg)
qed
qed
qed
qed
lift_definition of_circline :: "circline \<Rightarrow> ocircline" is of_circline_clmat
proof transfer
fix H1 H2
assume hh:
"hermitean H1 \<and> H1 \<noteq> mat_zero"
"hermitean H2 \<and> H2 \<noteq> mat_zero"
assume "circline_eq_cmat H1 H2"
then obtain k where *: "k \<noteq> 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
show "ocircline_eq_cmat (of_circline_cmat H1) (of_circline_cmat H2)"
proof (cases "k > 0")
case True
hence "pos_oriented_cmat H1 = pos_oriented_cmat H2"
using * pos_oriented_cmat_mult_positive[OF hh]
by blast
thus ?thesis
using hh * \<open>k > 0\<close>
apply (simp del: pos_oriented_cmat_def)
apply (rule conjI)
apply (rule impI)
apply (simp, rule_tac x=k in exI, simp)
apply (rule impI)
apply (simp, rule_tac x=k in exI, simp)
done
next
case False
hence "k < 0"
using *
by simp
hence "pos_oriented_cmat H1 \<longleftrightarrow> \<not> (pos_oriented_cmat H2)"
using * pos_oriented_cmat_mult_negative[OF hh]
by blast
thus ?thesis
using hh * \<open>k < 0\<close>
apply (simp del: pos_oriented_cmat_def)
apply (rule conjI)
apply (rule impI)
apply (simp, rule_tac x="-k" in exI, simp)
apply (rule impI)
apply (simp, rule_tac x="-k" in exI, simp)
done
qed
qed
lemma pos_oriented_of_circline [simp]:
shows "pos_oriented (of_circline H)"
using pos_oriented_opposite_ocircline_cmat
by (transfer, transfer, simp)
lemma of_ocircline_of_circline [simp]:
shows "of_ocircline (of_circline H) = H"
apply (transfer, auto simp add: of_circline_clmat_def')
apply (transfer, simp, rule_tac x="-1" in exI, simp)
done
lemma of_circline_of_ocircline_pos_oriented [simp]:
assumes "pos_oriented H"
shows "of_circline (of_ocircline H) = H"
using assms
by (transfer, transfer, simp, rule_tac x=1 in exI, simp)
lemma inj_of_circline:
assumes "of_circline H = of_circline H'"
shows "H = H'"
using assms
proof (transfer, transfer)
fix H H'
assume "ocircline_eq_cmat (of_circline_cmat H) (of_circline_cmat H')"
then obtain k where "k > 0" "of_circline_cmat H' = cor k *\<^sub>s\<^sub>m of_circline_cmat H"
by auto
thus "circline_eq_cmat H H'"
using mult_sm_inv_l[of "-1" "H'" "cor k *\<^sub>s\<^sub>m H"]
using mult_sm_inv_l[of "-1" "H'" "(- (cor k)) *\<^sub>s\<^sub>m H"]
apply (simp split: if_split_asm)
apply (rule_tac x="k" in exI, simp)
apply (rule_tac x="-k" in exI, simp)
apply (rule_tac x="-k" in exI, simp)
apply (rule_tac x="k" in exI, simp)
done
qed
lemma of_circline_of_ocircline:
shows "of_circline (of_ocircline H') = H' \<or>
of_circline (of_ocircline H') = opposite_ocircline H'"
proof (cases "pos_oriented H'")
case True
thus ?thesis
by auto
next
case False
hence "pos_oriented (opposite_ocircline H')"
using pos_oriented
by auto
thus ?thesis
using of_ocircline_opposite_ocircline[of H']
using of_circline_of_ocircline_pos_oriented [of "opposite_ocircline H'"]
by auto
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Set of points on oriented and unoriented circlines\<close>
(* -------------------------------------------------------------------------- *)
lemma ocircline_set_of_circline [simp]:
shows "ocircline_set (of_circline H) = circline_set H"
unfolding ocircline_set_def circline_set_def
proof (safe)
fix z
assume "on_ocircline (of_circline H) z"
thus "on_circline H z"
by (transfer, transfer, simp del: on_circline_cmat_cvec_def opposite_ocircline_cmat_def split: if_split_asm)
next
fix z
assume "on_circline H z"
thus "on_ocircline (of_circline H) z"
by (transfer, transfer, simp del: on_circline_cmat_cvec_def opposite_ocircline_cmat_def split: if_split_asm)
qed
(* ----------------------------------------------------------------- *)
subsection \<open>Some special oriented circlines and discs\<close>
(* ----------------------------------------------------------------- *)
lift_definition mk_ocircline :: "complex \<Rightarrow> complex \<Rightarrow> complex \<Rightarrow> complex \<Rightarrow> ocircline" is mk_circline_clmat
done
text \<open>oriented unit circle and unit disc\<close>
lift_definition ounit_circle :: "ocircline" is unit_circle_clmat
done
lemma pos_oriented_ounit_circle [simp]:
shows "pos_oriented ounit_circle"
by (transfer, transfer, simp)
lemma of_ocircline_ounit_circle [simp]:
shows "of_ocircline ounit_circle = unit_circle"
by (transfer, transfer, simp)
lemma of_circline_unit_circle [simp]:
shows "of_circline (unit_circle) = ounit_circle"
by (transfer, transfer, simp)
lemma ocircline_set_ounit_circle [simp]:
shows "ocircline_set ounit_circle = circline_set unit_circle"
apply (subst of_circline_unit_circle[symmetric])
apply (subst ocircline_set_of_circline)
apply simp
done
definition unit_disc :: "complex_homo set" where
"unit_disc = disc ounit_circle"
definition unit_disc_compl :: "complex_homo set" where
"unit_disc_compl = disc_compl ounit_circle"
definition unit_circle_set :: "complex_homo set" where
"unit_circle_set = circline_set unit_circle"
lemma zero_in_unit_disc [simp]:
shows "0\<^sub>h \<in> unit_disc"
unfolding unit_disc_def disc_def
by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def)
lemma one_notin_unit_dic [simp]:
shows "1\<^sub>h \<notin> unit_disc"
unfolding unit_disc_def disc_def
by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def)
lemma inf_notin_unit_disc [simp]:
shows "\<infinity>\<^sub>h \<notin> unit_disc"
unfolding unit_disc_def disc_def
by (simp, transfer, transfer) (simp add: Let_def vec_cnj_def)
lemma unit_disc_iff_cmod_lt_1 [simp]:
shows "of_complex c \<in> unit_disc \<longleftrightarrow> cmod c < 1"
unfolding unit_disc_def disc_def
by (simp, transfer, transfer, simp add: vec_cnj_def cmod_def power2_eq_square)
lemma unit_disc_cmod_square_lt_1 [simp]:
assumes "z \<in> unit_disc"
shows "(cmod (to_complex z))\<^sup>2 < 1"
using assms inf_or_of_complex[of z]
by (auto simp add: abs_square_less_1)
lemma unit_disc_to_complex_inj:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
assumes "to_complex u = to_complex v"
shows "u = v"
using assms
using inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
lemma inversion_unit_disc [simp]:
shows "inversion ` unit_disc = unit_disc_compl"
unfolding unit_disc_def unit_disc_compl_def disc_def disc_compl_def
proof safe
fix x
assume "in_ocircline ounit_circle x"
thus "out_ocircline ounit_circle (inversion x)"
unfolding inversion_def
by (transfer, transfer, auto simp add: vec_cnj_def)
next
fix x
assume *: "out_ocircline ounit_circle x"
show "x \<in> inversion ` Collect (in_ocircline ounit_circle)"
proof (rule image_eqI)
show "x = inversion (inversion x)"
by auto
next
show "inversion x \<in> Collect (in_ocircline ounit_circle)"
using *
unfolding inversion_def
by (simp, transfer, transfer, auto simp add: vec_cnj_def)
qed
qed
lemma inversion_unit_disc_compl [simp]:
shows "inversion ` unit_disc_compl = unit_disc"
proof-
have "inversion ` (inversion ` unit_disc) = unit_disc"
by (auto simp del: inversion_unit_disc simp add: image_iff)
thus ?thesis
by simp
qed
lemma inversion_noteq_unit_disc:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "inversion u \<noteq> v"
proof-
from assms
have "inversion u \<in> unit_disc_compl"
by (metis image_eqI inversion_unit_disc)
thus ?thesis
using assms
unfolding unit_disc_def unit_disc_compl_def
using disc_inter_disc_compl
by fastforce
qed
lemma in_ocircline_ounit_circle_conjugate [simp]:
assumes "in_ocircline ounit_circle z"
shows "in_ocircline ounit_circle (conjugate z)"
using assms
by (transfer, transfer, auto simp add: vec_cnj_def)
lemma conjugate_unit_disc [simp]:
shows "conjugate ` unit_disc = unit_disc"
unfolding unit_disc_def disc_def
apply (auto simp add: image_iff)
apply (rule_tac x="conjugate x" in exI, simp)
done
lemma conjugate_in_unit_disc [simp]:
assumes "z \<in> unit_disc"
shows "conjugate z \<in> unit_disc"
using conjugate_unit_disc
using assms
by blast
lemma out_ocircline_ounit_circle_conjugate [simp]:
assumes "out_ocircline ounit_circle z"
shows "out_ocircline ounit_circle (conjugate z)"
using assms
by (transfer, transfer, auto simp add: vec_cnj_def)
lemma conjugate_unit_disc_compl [simp]:
shows "conjugate ` unit_disc_compl = unit_disc_compl"
unfolding unit_disc_compl_def disc_compl_def
apply (auto simp add: image_iff)
apply (rule_tac x="conjugate x" in exI, simp)
done
lemma conjugate_in_unit_disc_compl [simp]:
assumes "z \<in> unit_disc_compl"
shows "conjugate z \<in> unit_disc_compl"
using conjugate_unit_disc_compl
using assms
by blast
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Oriented x axis and lower half plane\<close>
(* -------------------------------------------------------------------------- *)
lift_definition o_x_axis :: "ocircline" is x_axis_clmat
done
lemma o_x_axis_pos_oriented [simp]:
shows "pos_oriented o_x_axis"
by (transfer, transfer, simp)
lemma of_ocircline_o_x_axis [simp]:
shows "of_ocircline o_x_axis = x_axis"
by (transfer, transfer, simp)
lemma of_circline_x_axis [simp]:
shows "of_circline x_axis = o_x_axis"
using of_circline_of_ocircline_pos_oriented[of o_x_axis]
using o_x_axis_pos_oriented
by simp
lemma ocircline_set_circline_set_x_axis [simp]:
shows "ocircline_set o_x_axis = circline_set x_axis"
by (subst of_circline_x_axis[symmetric], subst ocircline_set_of_circline, simp)
lemma ii_in_disc_o_x_axis [simp]:
shows "ii\<^sub>h \<notin> disc o_x_axis"
unfolding disc_def
by simp (transfer, transfer, simp add: Let_def vec_cnj_def)
lemma ii_notin_disc_o_x_axis [simp]:
shows "ii\<^sub>h \<in> disc_compl o_x_axis"
unfolding disc_compl_def
by simp (transfer, transfer, simp add: Let_def vec_cnj_def)
lemma of_complex_in_o_x_axis_disc [simp]:
shows "of_complex z \<in> disc o_x_axis \<longleftrightarrow> Im z < 0"
unfolding disc_def
by auto (transfer, transfer, simp add: vec_cnj_def)+
lemma inf_notin_disc_o_x_axis [simp]:
shows "\<infinity>\<^sub>h \<notin> disc o_x_axis"
unfolding disc_def
by simp (transfer, transfer, simp add: vec_cnj_def)
lemma disc_o_x_axis:
shows "disc o_x_axis = of_complex ` {z. Im z < 0}"
proof-
{
fix z
assume "z \<in> disc o_x_axis"
hence "\<exists> x. Im x < 0 \<and> z = of_complex x"
using inf_or_of_complex[of z]
by auto
}
thus ?thesis
by (auto simp add: image_iff)
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Oriented single point circline\<close>
(* -------------------------------------------------------------------------- *)
lift_definition o_circline_point_0 :: "ocircline" is circline_point_0_clmat
done
lemma of_ocircline_o_circline_point_0 [simp]:
shows "of_ocircline o_circline_point_0 = circline_point_0"
by (transfer, transfer, simp)
(* ----------------------------------------------------------------- *)
subsection \<open>Möbius action on oriented circlines and discs\<close>
(* ----------------------------------------------------------------- *)
text \<open>Möbius action on an oriented circline is the same as on to an unoriented circline.\<close>
lift_definition moebius_ocircline :: "moebius \<Rightarrow> ocircline \<Rightarrow> ocircline" is moebius_circline_mmat_clmat
apply (transfer, transfer)
apply simp
apply ((erule exE)+, (erule conjE)+)
apply (simp add: mat_inv_mult_sm)
apply (rule_tac x="ka / Re (k * cnj k)" in exI, auto simp add: complex_mult_cnj_cmod power2_eq_square)
done
text \<open>Möbius action on (unoriented) circlines could have been defined using the action on oriented
circlines, but not the other way around.\<close>
lemma moebius_circline_ocircline:
shows "moebius_circline M H = of_ocircline (moebius_ocircline M (of_circline H))"
apply (transfer, simp add: of_circline_clmat_def', safe)
apply (transfer, simp, rule_tac x="-1" in exI, simp)
done
lemma moebius_ocircline_circline:
shows "moebius_ocircline M H = of_circline (moebius_circline M (of_ocircline H)) \<or>
moebius_ocircline M H = opposite_ocircline (of_circline (moebius_circline M (of_ocircline H)))"
apply (transfer, simp add: of_circline_clmat_def', safe)
apply (transfer, simp, rule_tac x="1" in exI, simp)
apply (transfer, simp, erule_tac x="1" in allE, simp)
done
text \<open>Möbius action on oriented circlines have many nice properties as it was the case with
Möbius action on (unoriented) circlines. These transformations are injective and form group under
composition.\<close>
lemma inj_moebius_ocircline [simp]:
shows "inj (moebius_ocircline M)"
unfolding inj_on_def
proof (safe)
fix H H'
assume "moebius_ocircline M H = moebius_ocircline M H'"
thus "H = H'"
proof (transfer, transfer)
fix M H H' :: complex_mat
assume "mat_det M \<noteq> 0"
let ?iM = "mat_inv M"
assume "ocircline_eq_cmat (moebius_circline_cmat_cmat M H) (moebius_circline_cmat_cmat M H')"
then obtain k where "congruence ?iM H' = congruence ?iM (cor k *\<^sub>s\<^sub>m H)" "k > 0"
by (auto simp del: congruence_def)
thus "ocircline_eq_cmat H H'"
using \<open>mat_det M \<noteq> 0\<close> inj_congruence[of ?iM H' "cor k *\<^sub>s\<^sub>m H"] mat_det_inv[of M]
by auto
qed
qed
lemma moebius_ocircline_id_moebius [simp]:
shows "moebius_ocircline id_moebius H = H"
by (transfer, transfer) (force simp add: mat_adj_def mat_cnj_def)
lemma moebius_ocircline_comp [simp]:
shows "moebius_ocircline (moebius_comp M1 M2) H = moebius_ocircline M1 (moebius_ocircline M2 H)"
by (transfer, transfer, simp, rule_tac x=1 in exI, simp add: mat_inv_mult_mm mult_mm_assoc)
lemma moebius_ocircline_comp_inv_left [simp]:
shows "moebius_ocircline (moebius_inv M) (moebius_ocircline M H) = H"
by (subst moebius_ocircline_comp[symmetric]) simp
lemma moebius_ocircline_comp_inv_right [simp]:
shows "moebius_ocircline M (moebius_ocircline (moebius_inv M) H) = H"
by (subst moebius_ocircline_comp[symmetric]) simp
lemma moebius_ocircline_opposite_ocircline [simp]:
shows "moebius_ocircline M (opposite_ocircline H) = opposite_ocircline (moebius_ocircline M H)"
by (transfer, transfer, simp, rule_tac x=1 in exI, simp)
text \<open>Möbius action on oriented circlines preserve the set of points of the circline.\<close>
lemma ocircline_set_moebius_ocircline [simp]:
shows "ocircline_set (moebius_ocircline M H) = moebius_pt M ` ocircline_set H" (is "?lhs = ?rhs")
proof-
have "?rhs = circline_set (moebius_circline M (of_ocircline H))"
by simp
thus ?thesis
using moebius_ocircline_circline[of M H]
by auto
qed
lemma ocircline_set_fix_iff_ocircline_fix:
assumes "ocircline_set H' \<noteq> {}"
shows "ocircline_set (moebius_ocircline M H) = ocircline_set H' \<longleftrightarrow>
moebius_ocircline M H = H' \<or> moebius_ocircline M H = opposite_ocircline H'"
using assms
using inj_ocircline_set[of "moebius_ocircline M H" H']
by (auto simp del: ocircline_set_moebius_ocircline)
lemma disc_moebius_ocircline [simp]:
shows "disc (moebius_ocircline M H) = moebius_pt M ` (disc H)"
proof (safe)
fix z
assume "z \<in> disc H"
thus "moebius_pt M z \<in> disc (moebius_ocircline M H)"
unfolding disc_def
proof (safe)
assume "in_ocircline H z"
thus "in_ocircline (moebius_ocircline M H) (moebius_pt M z)"
proof (transfer, transfer)
fix H M :: complex_mat and z :: complex_vec
assume "mat_det M \<noteq> 0"
assume "in_ocircline_cmat_cvec H z"
thus "in_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) (moebius_pt_cmat_cvec M z)"
using \<open>mat_det M \<noteq> 0\<close> quad_form_congruence[of M z]
by simp
qed
qed
next
fix z
assume "z \<in> disc (moebius_ocircline M H)"
thus "z \<in> moebius_pt M ` disc H"
unfolding disc_def
proof(safe)
assume "in_ocircline (moebius_ocircline M H) z"
show "z \<in> moebius_pt M ` Collect (in_ocircline H)"
proof
show "z = moebius_pt M (moebius_pt (moebius_inv M) z)"
by simp
next
show "moebius_pt (moebius_inv M) z \<in> Collect (in_ocircline H)"
using \<open>in_ocircline (moebius_ocircline M H) z\<close>
proof (safe, transfer, transfer)
fix M H :: complex_mat and z :: complex_vec
assume "mat_det M \<noteq> 0"
hence "congruence (mat_inv (mat_inv M)) (congruence (mat_inv M) H) = H"
by (simp del: congruence_def)
hence "quad_form z (congruence (mat_inv M) H) = quad_form (mat_inv M *\<^sub>m\<^sub>v z) H"
using quad_form_congruence[of "mat_inv M" "z" "congruence (mat_inv M) H"]
using \<open>mat_det M \<noteq> 0\<close> mat_det_inv[of "M"]
by simp
moreover
assume "in_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) z"
ultimately
show "in_ocircline_cmat_cvec H (moebius_pt_cmat_cvec (moebius_inv_cmat M) z)"
by simp
qed
qed
qed
qed
lemma disc_compl_moebius_ocircline [simp]:
shows "disc_compl (moebius_ocircline M H) = moebius_pt M ` (disc_compl H)"
proof (safe)
fix z
assume "z \<in> disc_compl H"
thus "moebius_pt M z \<in> disc_compl (moebius_ocircline M H)"
unfolding disc_compl_def
proof (safe)
assume "out_ocircline H z"
thus "out_ocircline (moebius_ocircline M H) (moebius_pt M z)"
proof (transfer, transfer)
fix H M :: complex_mat and z :: complex_vec
assume "mat_det M \<noteq> 0"
assume "out_ocircline_cmat_cvec H z"
thus "out_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) (moebius_pt_cmat_cvec M z)"
using \<open>mat_det M \<noteq> 0\<close> quad_form_congruence[of M z]
by simp
qed
qed
next
fix z
assume "z \<in> disc_compl (moebius_ocircline M H)"
thus "z \<in> moebius_pt M ` disc_compl H"
unfolding disc_compl_def
proof(safe)
assume "out_ocircline (moebius_ocircline M H) z"
show "z \<in> moebius_pt M ` Collect (out_ocircline H)"
proof
show "z = moebius_pt M (moebius_pt (moebius_inv M) z)"
by simp
next
show "moebius_pt (moebius_inv M) z \<in> Collect (out_ocircline H)"
using \<open>out_ocircline (moebius_ocircline M H) z\<close>
proof (safe, transfer, transfer)
fix M H :: complex_mat and z :: complex_vec
assume "mat_det M \<noteq> 0"
hence "congruence (mat_inv (mat_inv M)) (congruence (mat_inv M) H) = H"
by (simp del: congruence_def)
hence "quad_form z (congruence (mat_inv M) H) = quad_form (mat_inv M *\<^sub>m\<^sub>v z) H"
using quad_form_congruence[of "mat_inv M" "z" "congruence (mat_inv M) H"]
using \<open>mat_det M \<noteq> 0\<close> mat_det_inv[of "M"]
by simp
moreover
assume "out_ocircline_cmat_cvec (moebius_circline_cmat_cmat M H) z"
ultimately
show "out_ocircline_cmat_cvec H (moebius_pt_cmat_cvec (moebius_inv_cmat M) z)"
by simp
qed
qed
qed
qed
(* ----------------------------------------------------------------- *)
subsection \<open>Orientation after Möbius transformations\<close>
(* ----------------------------------------------------------------- *)
text \<open>All Euclidean similarities preserve circline orientation.\<close>
lemma moebius_similarity_oriented_lines_to_oriented_lines:
assumes "a \<noteq> 0"
shows "\<infinity>\<^sub>h \<in> ocircline_set H \<longleftrightarrow> \<infinity>\<^sub>h \<in> ocircline_set (moebius_ocircline (moebius_similarity a b) H)"
using moebius_similarity_lines_to_lines[OF \<open>a \<noteq> 0\<close>, of b "of_ocircline H"]
by simp
lemma moebius_similarity_preserve_orientation':
assumes "a \<noteq> 0" and "\<infinity>\<^sub>h \<notin> ocircline_set H" and "pos_oriented H"
shows "pos_oriented (moebius_ocircline (moebius_similarity a b) H)"
proof-
let ?M = "moebius_similarity a b"
let ?H = "moebius_ocircline ?M H"
have "\<infinity>\<^sub>h \<notin> ocircline_set ?H"
using \<open>\<infinity>\<^sub>h \<notin> ocircline_set H\<close> moebius_similarity_oriented_lines_to_oriented_lines[OF \<open>a \<noteq> 0\<close>]
by simp
have "\<infinity>\<^sub>h \<in> disc_compl H"
using \<open>\<infinity>\<^sub>h \<notin> ocircline_set H\<close> \<open>pos_oriented H\<close> pos_oriented_circle_inf[of H] in_on_out
unfolding disc_def disc_compl_def ocircline_set_def
by auto
hence "\<infinity>\<^sub>h \<in> disc_compl ?H"
using moebius_similarity_inf[OF \<open>a \<noteq> 0\<close>, of b]
by force
thus "pos_oriented ?H"
using pos_oriented_circle_inf[of ?H] disc_inter_disc_compl[of ?H] \<open>\<infinity>\<^sub>h \<notin> ocircline_set ?H\<close>
by auto
qed
lemma moebius_similarity_preserve_orientation:
assumes "a \<noteq> 0" and "\<infinity>\<^sub>h \<notin> ocircline_set H"
shows "pos_oriented H \<longleftrightarrow> pos_oriented(moebius_ocircline (moebius_similarity a b) H)"
proof-
let ?M = "moebius_similarity a b"
let ?H = "moebius_ocircline ?M H"
have "\<infinity>\<^sub>h \<notin> ocircline_set ?H"
using \<open>\<infinity>\<^sub>h \<notin> ocircline_set H\<close> moebius_similarity_oriented_lines_to_oriented_lines[OF \<open>a \<noteq> 0\<close>]
by simp
have *: "H = moebius_ocircline (- moebius_similarity a b) ?H"
by simp
show ?thesis
using \<open>a \<noteq> 0\<close>
using moebius_similarity_preserve_orientation' [OF \<open>a \<noteq> 0\<close> \<open>\<infinity>\<^sub>h \<notin> ocircline_set H\<close>]
using moebius_similarity_preserve_orientation'[OF _ \<open>\<infinity>\<^sub>h \<notin> ocircline_set ?H\<close>, of "1/a" "-b/a"]
using moebius_similarity_inv[of a b, OF \<open>a \<noteq> 0\<close>] *
by auto
qed
lemma reciprocal_preserve_orientation:
assumes "0\<^sub>h \<in> disc_compl H"
shows "pos_oriented (moebius_ocircline moebius_reciprocal H)"
proof-
have "\<infinity>\<^sub>h \<in> disc_compl (moebius_ocircline moebius_reciprocal H)"
using assms
by force
thus "pos_oriented (moebius_ocircline moebius_reciprocal H)"
using pos_oriented_circle_inf[of "moebius_ocircline moebius_reciprocal H"]
using disc_inter_disc_compl[of "moebius_ocircline moebius_reciprocal H"]
using disc_compl_inter_ocircline_set[of "moebius_ocircline moebius_reciprocal H"]
by auto
qed
lemma reciprocal_not_preserve_orientation:
assumes "0\<^sub>h \<in> disc H"
shows "\<not> pos_oriented (moebius_ocircline moebius_reciprocal H)"
proof-
let ?H = "moebius_ocircline moebius_reciprocal H"
have "\<infinity>\<^sub>h \<in> disc ?H"
using assms
by force
thus "\<not> pos_oriented ?H"
using pos_oriented_circle_inf[of ?H] disc_inter_ocircline_set[of ?H]
by auto
qed
text \<open>Orientation of the image of a given oriented circline $H$ under a given Möbius transformation
$M$ depends on whether the pole of $M$ (the point that $M$ maps to $\infty_{hc}$) lies in the disc
or in the disc complement of $H$ (if it is on the set of $H$, then it maps onto a line and we do not
discuss the orientation).\<close>
lemma pole_in_disc:
assumes "M = mk_moebius a b c d" and "c \<noteq> 0" and "a*d - b*c \<noteq> 0"
assumes "is_pole M z" "z \<in> disc H"
shows "\<not> pos_oriented (moebius_ocircline M H)"
proof-
let ?t1 = "moebius_translation (a / c)"
let ?rd = "moebius_rotation_dilatation ((b * c - a * d) / (c * c))"
let ?r = "moebius_reciprocal"
let ?t2 = "moebius_translation (d / c)"
have "0\<^sub>h = moebius_pt (moebius_translation (d/c)) z"
using pole_mk_moebius[of a b c d z] assms
by simp
have "z \<notin> ocircline_set H"
using \<open>z \<in> disc H\<close> disc_inter_ocircline_set[of H]
by blast
hence "0\<^sub>h \<notin> ocircline_set (moebius_ocircline ?t2 H)"
using \<open>0\<^sub>h = moebius_pt ?t2 z\<close>
using moebius_pt_neq_I[of z _ ?t2]
by force
hence *: "\<infinity>\<^sub>h \<notin> ocircline_set (moebius_ocircline (?r + ?t2) H)"
using \<open>0\<^sub>h = moebius_pt (moebius_translation (d / c)) z\<close>
by (metis circline_set_moebius_circline circline_set_moebius_circline_iff circline_set_of_ocircline moebius_pt_comp moebius_reciprocal ocircline_set_moebius_ocircline plus_moebius_def reciprocal_zero)
hence **: "\<infinity>\<^sub>h \<notin> ocircline_set (moebius_ocircline (?rd + ?r + ?t2) H)"
using \<open>a*d - b*c \<noteq> 0\<close> \<open>c \<noteq> 0\<close>
unfolding moebius_rotation_dilatation_def
using moebius_similarity_oriented_lines_to_oriented_lines[of _ "moebius_ocircline (?r + ?t2) H"]
by (metis divide_eq_0_iff divisors_zero moebius_ocircline_comp plus_moebius_def right_minus_eq)
have "\<not> pos_oriented (moebius_ocircline (?r + ?t2) H)"
using pole_mk_moebius[of a b c d z] assms
using reciprocal_not_preserve_orientation
by force
hence "\<not> pos_oriented (moebius_ocircline (?rd + ?r + ?t2) H)"
using *
using \<open>a*d - b*c \<noteq> 0\<close> \<open>c \<noteq> 0\<close>
using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?r + ?t2) H"]
unfolding moebius_rotation_dilatation_def
by simp
hence "\<not> pos_oriented (moebius_ocircline (?t1 + ?rd + ?r + ?t2) H)"
using **
using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?rd + ?r + ?t2) H"]
unfolding moebius_translation_def
by simp
thus ?thesis
using assms
by simp (subst moebius_decomposition, simp_all)
qed
lemma pole_in_disc_compl:
assumes "M = mk_moebius a b c d" and "c \<noteq> 0" and "a*d - b*c \<noteq> 0"
assumes "is_pole M z" and "z \<in> disc_compl H"
shows "pos_oriented (moebius_ocircline M H)"
proof-
let ?t1 = "moebius_translation (a / c)"
let ?rd = "moebius_rotation_dilatation ((b * c - a * d) / (c * c))"
let ?r = "moebius_reciprocal"
let ?t2 = "moebius_translation (d / c)"
have "0\<^sub>h = moebius_pt (moebius_translation (d/c)) z"
using pole_mk_moebius[of a b c d z] assms
by simp
have "z \<notin> ocircline_set H"
using \<open>z \<in> disc_compl H\<close> disc_compl_inter_ocircline_set[of H]
by blast
hence "0\<^sub>h \<notin> ocircline_set (moebius_ocircline ?t2 H)"
using \<open>0\<^sub>h = moebius_pt ?t2 z\<close>
using moebius_pt_neq_I[of z _ ?t2]
by force
hence *: "\<infinity>\<^sub>h \<notin> ocircline_set (moebius_ocircline (?r + ?t2) H)"
using \<open>0\<^sub>h = moebius_pt (moebius_translation (d / c)) z\<close>
by (metis circline_set_moebius_circline circline_set_moebius_circline_iff circline_set_of_ocircline moebius_pt_comp moebius_reciprocal ocircline_set_moebius_ocircline plus_moebius_def reciprocal_zero)
hence **: "\<infinity>\<^sub>h \<notin> ocircline_set (moebius_ocircline (?rd + ?r + ?t2) H)"
using \<open>a*d - b*c \<noteq> 0\<close> \<open>c \<noteq> 0\<close>
unfolding moebius_rotation_dilatation_def
using moebius_similarity_oriented_lines_to_oriented_lines[of _ "moebius_ocircline (?r + ?t2) H"]
by (metis divide_eq_0_iff divisors_zero moebius_ocircline_comp plus_moebius_def right_minus_eq)
have "pos_oriented (moebius_ocircline (?r + ?t2) H)"
using pole_mk_moebius[of a b c d z] assms
using reciprocal_preserve_orientation
by force
hence "pos_oriented (moebius_ocircline (?rd + ?r + ?t2) H)"
using *
using \<open>a*d - b*c \<noteq> 0\<close> \<open>c \<noteq> 0\<close>
using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?r + ?t2) H"]
unfolding moebius_rotation_dilatation_def
by simp
hence "pos_oriented (moebius_ocircline (?t1 + ?rd + ?r + ?t2) H)"
using **
using moebius_similarity_preserve_orientation[of _ "moebius_ocircline (?rd + ?r + ?t2) H"]
unfolding moebius_translation_def
by simp
thus ?thesis
using assms
by simp (subst moebius_decomposition, simp_all)
qed
(* ----------------------------------------------------------------- *)
subsection \<open>Oriented circlines uniqueness\<close>
(* ----------------------------------------------------------------- *)
lemma ocircline_01inf:
assumes "0\<^sub>h \<in> ocircline_set H \<and> 1\<^sub>h \<in> ocircline_set H \<and> \<infinity>\<^sub>h \<in> ocircline_set H"
shows "H = o_x_axis \<or> H = opposite_ocircline o_x_axis"
proof-
have "0\<^sub>h \<in> circline_set (of_ocircline H) \<and> 1\<^sub>h \<in> circline_set (of_ocircline H) \<and> \<infinity>\<^sub>h \<in> circline_set (of_ocircline H)"
using assms
by simp
hence "of_ocircline H = x_axis"
using unique_circline_01inf'
by auto
thus "H = o_x_axis \<or> H = opposite_ocircline o_x_axis"
by (metis inj_of_ocircline of_ocircline_o_x_axis)
qed
lemma unique_ocircline_01inf:
shows "\<exists>! H. 0\<^sub>h \<in> ocircline_set H \<and> 1\<^sub>h \<in> ocircline_set H \<and> \<infinity>\<^sub>h \<in> ocircline_set H \<and> ii\<^sub>h \<notin> disc H"
proof
show "0\<^sub>h \<in> ocircline_set o_x_axis \<and> 1\<^sub>h \<in> ocircline_set o_x_axis \<and> \<infinity>\<^sub>h \<in> ocircline_set o_x_axis \<and> ii\<^sub>h \<notin> disc o_x_axis"
by simp
next
fix H
assume "0\<^sub>h \<in> ocircline_set H \<and> 1\<^sub>h \<in> ocircline_set H \<and> \<infinity>\<^sub>h \<in> ocircline_set H \<and> ii\<^sub>h \<notin> disc H"
hence "0\<^sub>h \<in> ocircline_set H \<and> 1\<^sub>h \<in> ocircline_set H \<and> \<infinity>\<^sub>h \<in> ocircline_set H" "ii\<^sub>h \<notin> disc H"
by auto
hence "H = o_x_axis \<or> H = opposite_ocircline o_x_axis"
using ocircline_01inf
by simp
thus "H = o_x_axis"
using \<open>ii\<^sub>h \<notin> disc H\<close>
by auto
qed
lemma unique_ocircline_set:
assumes "A \<noteq> B" and "A \<noteq> C" and "B \<noteq> C"
shows "\<exists>! H. pos_oriented H \<and> (A \<in> ocircline_set H \<and> B \<in> ocircline_set H \<and> C \<in> ocircline_set H)"
proof-
obtain M where *: "moebius_pt M A = 0\<^sub>h" "moebius_pt M B = 1\<^sub>h" "moebius_pt M C = \<infinity>\<^sub>h"
using ex_moebius_01inf[OF assms]
by auto
let ?iM = "moebius_pt (moebius_inv M)"
have **: "?iM 0\<^sub>h = A" "?iM 1\<^sub>h = B" "?iM \<infinity>\<^sub>h = C"
using *
by (auto simp add: moebius_pt_invert)
let ?H = "moebius_ocircline (moebius_inv M) o_x_axis"
have 1: "A \<in> ocircline_set ?H" "B \<in> ocircline_set ?H" "C \<in> ocircline_set ?H"
using **
by auto
have 2: "\<And> H'. A \<in> ocircline_set H' \<and> B \<in> ocircline_set H' \<and> C \<in> ocircline_set H' \<Longrightarrow> H' = ?H \<or> H' = opposite_ocircline ?H"
proof-
fix H'
let ?H' = "ocircline_set H'" and ?H'' = "ocircline_set (moebius_ocircline M H')"
assume "A \<in> ocircline_set H' \<and> B \<in> ocircline_set H' \<and> C \<in> ocircline_set H'"
hence "moebius_pt M A \<in> ?H''" "moebius_pt M B \<in> ?H''" "moebius_pt M C \<in> ?H''"
by auto
hence "0\<^sub>h \<in> ?H''" "1\<^sub>h \<in> ?H''" "\<infinity>\<^sub>h \<in> ?H''"
using *
by auto
hence "moebius_ocircline M H' = o_x_axis \<or> moebius_ocircline M H' = opposite_ocircline o_x_axis"
using ocircline_01inf
by auto
hence "o_x_axis = moebius_ocircline M H' \<or> o_x_axis = opposite_ocircline (moebius_ocircline M H')"
by auto
thus "H' = ?H \<or> H' = opposite_ocircline ?H"
proof
assume *: "o_x_axis = moebius_ocircline M H'"
show "H' = moebius_ocircline (moebius_inv M) o_x_axis \<or> H' = opposite_ocircline (moebius_ocircline (moebius_inv M) o_x_axis)"
by (rule disjI1) (subst *, simp)
next
assume *: "o_x_axis = opposite_ocircline (moebius_ocircline M H')"
show "H' = moebius_ocircline (moebius_inv M) o_x_axis \<or> H' = opposite_ocircline (moebius_ocircline (moebius_inv M) o_x_axis)"
by (rule disjI2) (subst *, simp)
qed
qed
show ?thesis (is "\<exists>! x. ?P x")
proof (cases "pos_oriented ?H")
case True
show ?thesis
proof
show "?P ?H"
using 1 True
by auto
next
fix H
assume "?P H"
thus "H = ?H"
using 1 2[of H] True
by auto
qed
next
case False
let ?OH = "opposite_ocircline ?H"
show ?thesis
proof
show "?P ?OH"
using 1 False
by auto
next
fix H
assume "?P H"
thus "H = ?OH"
using False 2[of H]
by auto
qed
qed
qed
lemma ocircline_set_0h:
assumes "ocircline_set H = {0\<^sub>h}"
shows "H = o_circline_point_0 \<or> H = opposite_ocircline (o_circline_point_0)"
proof-
have "of_ocircline H = circline_point_0"
using assms
using unique_circline_type_zero_0' card_eq1_circline_type_zero[of "of_ocircline H"]
by auto
thus ?thesis
by (metis inj_of_ocircline of_ocircline_o_circline_point_0)
qed
end
diff --git a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
--- a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
+++ b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
@@ -1,1202 +1,1202 @@
(* ---------------------------------------------------------------------------- *)
section \<open>Unit circle preserving Möbius transformations\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>In this section we shall examine Möbius transformations that map the unit circle onto itself.
We shall say that they fix or preserve the unit circle (although, they do not need to fix each of
its points).\<close>
theory Unit_Circle_Preserving_Moebius
imports Unitary11_Matrices Moebius Oriented_Circlines
begin
(* ---------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the unit circle\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>We define Möbius transformations that preserve unit circle as transformations represented by
generalized unitary matrices with the $1-1$ signature (elements of the gruop $GU_{1,1}(2,
\mathbb{C})$, defined earlier in the theory Unitary11Matrices).\<close>
lift_definition unit_circle_fix_mmat :: "moebius_mat \<Rightarrow> bool" is unitary11_gen
done
lift_definition unit_circle_fix :: "moebius \<Rightarrow> bool" is unit_circle_fix_mmat
apply transfer
apply (auto simp del: mult_sm.simps)
apply (simp del: mult_sm.simps add: unitary11_gen_mult_sm)
apply (simp del: mult_sm.simps add: unitary11_gen_div_sm)
done
text \<open>Our algebraic characterisation (by matrices) is geometrically correct.\<close>
lemma unit_circle_fix_iff:
shows "unit_circle_fix M \<longleftrightarrow>
moebius_circline M unit_circle = unit_circle" (is "?rhs = ?lhs")
proof
assume ?lhs
thus ?rhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
assume "circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
then obtain k where "k \<noteq> 0" "(1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) (1, 0, 0, -1)"
by auto
hence "(1/cor k, 0, 0, -1/cor k) = congruence (mat_inv M) (1, 0, 0, -1)"
using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) (1, 0, 0, -1)" ]
by simp
hence "congruence M (1/cor k, 0, 0, -1/cor k) = (1, 0, 0, -1)"
using \<open>mat_det M \<noteq> 0\<close> mat_det_inv[of M]
using congruence_inv[of "mat_inv M" "(1, 0, 0, -1)" "(1/cor k, 0, 0, -1/cor k)"]
by simp
hence "congruence M (1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)"
using congruence_scale_m[of "M" "1/cor k" "(1, 0, 0, -1)"]
using mult_sm_inv_l[of "1/ cor k" "congruence M (1, 0, 0, -1)" "(1, 0, 0, -1)"] \<open>k \<noteq> 0\<close>
by simp
thus "unitary11_gen M"
using \<open>k \<noteq> 0\<close>
unfolding unitary11_gen_def
by simp
qed
next
assume ?rhs
thus ?lhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
assume "unitary11_gen M"
hence "unitary11_gen (mat_inv M)"
using \<open>mat_det M \<noteq> 0\<close>
using unitary11_gen_mat_inv
by simp
thus " circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
unfolding unitary11_gen_real
by auto (rule_tac x="1/k" in exI, simp)
qed
qed
lemma circline_set_fix_iff_circline_fix:
assumes "circline_set H' \<noteq> {}"
shows "circline_set (moebius_circline M H) = circline_set H' \<longleftrightarrow>
moebius_circline M H = H'"
using assms
by auto (rule inj_circline_set, auto)
lemma unit_circle_fix_iff_unit_circle_set:
shows "unit_circle_fix M \<longleftrightarrow> moebius_pt M ` unit_circle_set = unit_circle_set"
proof-
have "circline_set unit_circle \<noteq> {}"
using one_in_unit_circle_set
by auto
thus ?thesis
using unit_circle_fix_iff[of M] circline_set_fix_iff_circline_fix[of unit_circle M unit_circle]
by (simp add: unit_circle_set_def)
qed
text \<open>Unit circle preserving Möbius transformations form a group. \<close>
lemma unit_circle_fix_id_moebius [simp]:
shows "unit_circle_fix id_moebius"
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma unit_circle_fix_moebius_add [simp]:
assumes "unit_circle_fix M1" and "unit_circle_fix M2"
shows "unit_circle_fix (M1 + M2)"
using assms
unfolding unit_circle_fix_iff
by auto
lemma unit_circle_fix_moebius_comp [simp]:
assumes "unit_circle_fix M1" and "unit_circle_fix M2"
shows "unit_circle_fix (moebius_comp M1 M2)"
using unit_circle_fix_moebius_add[OF assms]
by simp
lemma unit_circle_fix_moebius_uminus [simp]:
assumes "unit_circle_fix M"
shows "unit_circle_fix (-M)"
using assms
unfolding unit_circle_fix_iff
by (metis moebius_circline_comp_inv_left uminus_moebius_def)
lemma unit_circle_fix_moebius_inv [simp]:
assumes "unit_circle_fix M"
shows "unit_circle_fix (moebius_inv M)"
using unit_circle_fix_moebius_uminus[OF assms]
by simp
text \<open>Unit circle fixing transforms preserve inverse points.\<close>
lemma unit_circle_fix_moebius_pt_inversion [simp]:
assumes "unit_circle_fix M"
shows "moebius_pt M (inversion z) = inversion (moebius_pt M z)"
using assms
using symmetry_principle[of z "inversion z" unit_circle M]
using unit_circle_fix_iff[of M, symmetric]
using circline_symmetric_inv_homo_disc[of z]
using circline_symmetric_inv_homo_disc'[of "moebius_pt M z" "moebius_pt M (inversion z)"]
by metis
(* -------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the imaginary unit circle\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Only for completeness we show that Möbius transformations that preserve the imaginary unit
circle are exactly those characterised by generalized unitary matrices (with the (2, 0) signature).\<close>
lemma imag_unit_circle_fixed_iff_unitary_gen:
assumes "mat_det (A, B, C, D) \<noteq> 0"
shows "moebius_circline (mk_moebius A B C D) imag_unit_circle = imag_unit_circle \<longleftrightarrow>
unitary_gen (A, B, C, D)" (is "?lhs = ?rhs")
proof
assume ?lhs
thus ?rhs
using assms
proof (transfer, transfer)
fix A B C D :: complex
let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)"
assume "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat"
"mat_det ?M \<noteq> 0"
then obtain k where "k \<noteq> 0" "?E = cor k *\<^sub>s\<^sub>m congruence (mat_inv ?M) ?E"
by auto
hence "unitary_gen (mat_inv ?M)"
using mult_sm_inv_l[of "cor k" "congruence (mat_inv ?M) ?E" "?E"]
unfolding unitary_gen_def
by (metis congruence_def divide_eq_0_iff eye_def mat_eye_r of_real_eq_0_iff one_neq_zero)
thus "unitary_gen ?M"
using unitary_gen_inv[of "mat_inv ?M"] \<open>mat_det ?M \<noteq> 0\<close>
by (simp del: mat_inv.simps)
qed
next
assume ?rhs
thus ?lhs
using assms
proof (transfer, transfer)
fix A B C D :: complex
let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)"
assume "unitary_gen ?M" "mat_det ?M \<noteq> 0"
hence "unitary_gen (mat_inv ?M)"
using unitary_gen_inv[of ?M]
by simp
then obtain k where "k \<noteq> 0" "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M) = cor k *\<^sub>s\<^sub>m eye"
using unitary_gen_real[of "mat_inv ?M"] mat_det_inv[of ?M]
by auto
hence *: "?E = (1 / cor k) *\<^sub>s\<^sub>m (mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M))"
using mult_sm_inv_l[of "cor k" eye "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M)"]
by simp
have "\<exists>k. k \<noteq> 0 \<and>
(1, 0, 0, 1) = cor k *\<^sub>s\<^sub>m (mat_adj (mat_inv (A, B, C, D)) *\<^sub>m\<^sub>m (1, 0, 0, 1) *\<^sub>m\<^sub>m mat_inv (A, B, C, D))"
using \<open>mat_det ?M \<noteq> 0\<close> \<open>k \<noteq> 0\<close>
by (metis "*" Im_complex_of_real Re_complex_of_real \<open>mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m mat_inv ?M = cor k *\<^sub>s\<^sub>m eye\<close> complex_of_real_Re eye_def mat_eye_l mult_mm_assoc mult_mm_sm mult_sm_eye_mm of_real_1 of_real_divide of_real_eq_1_iff zero_eq_1_divide_iff)
thus "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat"
using \<open>mat_det ?M \<noteq> 0\<close> \<open>k \<noteq> 0\<close>
by (simp del: mat_inv.simps)
qed
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the oriented unit circle and the unit disc\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Möbius transformations that fix the unit circle either map the unit disc onto itself or
exchange it with its exterior. The transformations that fix the unit disc can be recognized from
their matrices -- they have the form as before, but additionally it must hold that $|a|^2 > |b|^2$.\<close>
definition unit_disc_fix_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "unit_disc_fix_cmat M \<longleftrightarrow>
(let (A, B, C, D) = M
in unitary11_gen (A, B, C, D) \<and> (B = 0 \<or> Re ((A*D)/(B*C)) > 1))"
lift_definition unit_disc_fix_mmat :: "moebius_mat \<Rightarrow> bool" is unit_disc_fix_cmat
done
lift_definition unit_disc_fix :: "moebius \<Rightarrow> bool" is unit_disc_fix_mmat
proof transfer
fix M M' :: complex_mat
assume det: "mat_det M \<noteq> 0" "mat_det M' \<noteq> 0"
assume "moebius_cmat_eq M M'"
then obtain k where *: "k \<noteq> 0" "M' = k *\<^sub>s\<^sub>m M"
by auto
hence **: "unitary11_gen M \<longleftrightarrow> unitary11_gen M'"
using unitary11_gen_mult_sm[of k M] unitary11_gen_div_sm[of k M]
by auto
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
obtain A' B' C' D' where MM': "(A', B', C', D') = M'"
by (cases M') auto
show "unit_disc_fix_cmat M = unit_disc_fix_cmat M'"
using * ** MM MM'
by auto
qed
text \<open>Transformations that fix the unit disc also fix the unit circle.\<close>
lemma unit_disc_fix_unit_circle_fix [simp]:
assumes "unit_disc_fix M"
shows "unit_circle_fix M"
using assms
by (transfer, transfer, auto)
text \<open>Transformations that preserve the unit disc preserve the orientation of the unit circle.\<close>
lemma unit_disc_fix_iff_ounit_circle:
shows "unit_disc_fix M \<longleftrightarrow>
moebius_ocircline M ounit_circle = ounit_circle" (is "?rhs \<longleftrightarrow> ?lhs")
proof
assume *: ?lhs
have "moebius_circline M unit_circle = unit_circle"
apply (subst moebius_circline_ocircline[of M unit_circle])
apply (subst of_circline_unit_circle)
apply (subst *)
by simp
hence "unit_circle_fix M"
by (simp add: unit_circle_fix_iff)
thus ?rhs
using *
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
let ?H = "(1, 0, 0, -1)"
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
assume "unitary11_gen M" "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
then obtain k where "0 < k" "?H = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) ?H"
by auto
hence "congruence M ?H = cor k *\<^sub>s\<^sub>m ?H"
using congruence_inv[of "mat_inv M" "?H" "(1/cor k) *\<^sub>s\<^sub>m ?H"] \<open>mat_det M \<noteq> 0\<close>
using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) ?H" "?H"]
using mult_sm_inv_l[of "1/cor k" "congruence M ?H"]
using congruence_scale_m[of M "1/cor k" "?H"]
using \<open>\<And>B. \<lbrakk>1 / cor k \<noteq> 0; (1 / cor k) *\<^sub>s\<^sub>m congruence M (1, 0, 0, - 1) = B\<rbrakk> \<Longrightarrow> congruence M (1, 0, 0, - 1) = (1 / (1 / cor k)) *\<^sub>s\<^sub>m B\<close>
by (auto simp add: mat_det_inv)
then obtain a b k' where "k' \<noteq> 0" "M = k' *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" "sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1"
using unitary11_sgn_det_orientation'[of M k] \<open>k > 0\<close>
by auto
moreover
have "mat_det (a, b, cnj b, cnj a) \<noteq> 0"
using \<open>sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1\<close>
by (smt sgn_0 zero_complex.simps(1))
ultimately
show "unit_disc_fix_cmat M"
using unitary11_sgn_det[of k' a b M A B C D]
using MM[symmetric] \<open>k > 0\<close> \<open>unitary11_gen M\<close>
by (simp add: sgn_1_pos split: if_split_asm)
qed
next
assume ?rhs
thus ?lhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
assume "unit_disc_fix_cmat M"
hence "unitary11_gen M" "B = 0 \<or> 1 < Re (A * D / (B * C))"
using MM[symmetric]
by auto
have "sgn (if B = 0 then 1 else sgn (Re (A * D / (B * C)) - 1)) = 1"
using \<open>B = 0 \<or> 1 < Re (A * D / (B * C))\<close>
by auto
then obtain k' where "k' > 0" "congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)"
using unitary11_orientation[OF \<open>unitary11_gen M\<close> MM[symmetric]]
by (auto simp add: sgn_1_pos)
thus "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
using congruence_inv[of M "(1, 0, 0, -1)" "cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)"] \<open>mat_det M \<noteq> 0\<close>
using congruence_scale_m[of "mat_inv M" "cor k'" "(1, 0, 0, -1)"]
by auto
qed
qed
text \<open>Our algebraic characterisation (by matrices) is geometrically correct.\<close>
lemma unit_disc_fix_iff [simp]:
assumes "unit_disc_fix M"
shows "moebius_pt M ` unit_disc = unit_disc"
using assms
using unit_disc_fix_iff_ounit_circle[of M]
unfolding unit_disc_def
by (subst disc_moebius_ocircline[symmetric], simp)
lemma unit_disc_fix_discI [simp]:
assumes "unit_disc_fix M" and "u \<in> unit_disc"
shows "moebius_pt M u \<in> unit_disc"
using unit_disc_fix_iff assms
by blast
text \<open>Unit disc preserving transformations form a group.\<close>
lemma unit_disc_fix_id_moebius [simp]:
shows "unit_disc_fix id_moebius"
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma unit_disc_fix_moebius_add [simp]:
assumes "unit_disc_fix M1" and "unit_disc_fix M2"
shows "unit_disc_fix (M1 + M2)"
using assms
unfolding unit_disc_fix_iff_ounit_circle
by auto
lemma unit_disc_fix_moebius_comp [simp]:
assumes "unit_disc_fix M1" and "unit_disc_fix M2"
shows "unit_disc_fix (moebius_comp M1 M2)"
using unit_disc_fix_moebius_add[OF assms]
by simp
lemma unit_disc_fix_moebius_uminus [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (-M)"
using assms
unfolding unit_disc_fix_iff_ounit_circle
by (metis moebius_ocircline_comp_inv_left uminus_moebius_def)
lemma unit_disc_fix_moebius_inv [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (moebius_inv M)"
using unit_disc_fix_moebius_uminus[OF assms]
by simp
(* -------------------------------------------------------------------------- *)
subsection \<open>Rotations are unit disc preserving transformations\<close>
(* -------------------------------------------------------------------------- *)
lemma unit_disc_fix_rotation [simp]:
shows "unit_disc_fix (moebius_rotation \<phi>)"
unfolding moebius_rotation_def moebius_similarity_def
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def cis_mult)
lemma moebius_rotation_unit_circle_fix [simp]:
shows "moebius_pt (moebius_rotation \<phi>) u \<in> unit_circle_set \<longleftrightarrow> u \<in> unit_circle_set"
using moebius_pt_moebius_inv_in_set unit_circle_fix_iff_unit_circle_set
by fastforce
lemma ex_rotation_mapping_u_to_positive_x_axis:
assumes "u \<noteq> 0\<^sub>h" and "u \<noteq> \<infinity>\<^sub>h"
shows "\<exists> \<phi>. moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
proof-
from assms obtain c where *: "u = of_complex c"
using inf_or_of_complex
by blast
have "is_real (cis (- arg c) * c)" "Re (cis (-arg c) * c) > 0"
using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff
by blast+
thus ?thesis
using *
by (rule_tac x="-arg c" in exI) (simp add: positive_x_axis_def circline_set_x_axis)
qed
lemma ex_rotation_mapping_u_to_positive_y_axis:
assumes "u \<noteq> 0\<^sub>h" and "u \<noteq> \<infinity>\<^sub>h"
shows "\<exists> \<phi>. moebius_pt (moebius_rotation \<phi>) u \<in> positive_y_axis"
proof-
from assms obtain c where *: "u = of_complex c"
using inf_or_of_complex
by blast
have "is_imag (cis (pi/2 - arg c) * c)" "Im (cis (pi/2 - arg c) * c) > 0"
using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff
by - (simp, simp, simp add: field_simps)
thus ?thesis
using *
by (rule_tac x="pi/2-arg c" in exI) (simp add: positive_y_axis_def circline_set_y_axis)
qed
lemma wlog_rotation_to_positive_x_axis:
assumes in_disc: "u \<in> unit_disc" and not_zero: "u \<noteq> 0\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<in> unit_disc; u \<noteq> 0\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes x_axis: "\<And>x. \<lbrakk>is_real x; 0 < Re x; Re x < 1\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from in_disc obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of u]
using inf_notin_unit_disc not_zero
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?x = "to_complex ?Mu"
have "?Mu \<in> unit_disc" "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<in> unit_disc\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_real (to_complex ?Mu)" "0 < Re ?x" "Re ?x < 1"
using *
unfolding positive_x_axis_def circline_set_x_axis
by (auto simp add: cmod_eq_Re)
thus ?thesis
using x_axis[of ?x] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF in_disc] not_zero
by simp
qed
lemma wlog_rotation_to_positive_x_axis':
assumes not_zero: "u \<noteq> 0\<^sub>h" and not_inf: "u \<noteq> \<infinity>\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<noteq> 0\<^sub>h; u \<noteq> \<infinity>\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes x_axis: "\<And>x. \<lbrakk>is_real x; 0 < Re x\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from not_zero and not_inf obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of u]
using inf_notin_unit_disc
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?x = "to_complex ?Mu"
have "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<noteq> \<infinity>\<^sub>h\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_real (to_complex ?Mu)" "0 < Re ?x"
using *
unfolding positive_x_axis_def circline_set_x_axis
by (auto simp add: cmod_eq_Re)
thus ?thesis
using x_axis[of ?x] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF not_zero not_inf]
by simp
qed
lemma wlog_rotation_to_positive_y_axis:
assumes in_disc: "u \<in> unit_disc" and not_zero: "u \<noteq> 0\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<in> unit_disc; u \<noteq> 0\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes y_axis: "\<And>x. \<lbrakk>is_imag x; 0 < Im x; Im x < 1\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from in_disc and not_zero obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_y_axis"
using ex_rotation_mapping_u_to_positive_y_axis[of u]
using inf_notin_unit_disc
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?y = "to_complex ?Mu"
have "?Mu \<in> unit_disc" "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<in> unit_disc\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_imag (to_complex ?Mu)" "0 < Im ?y" "Im ?y < 1"
using *
unfolding positive_y_axis_def circline_set_y_axis
by (auto simp add: cmod_eq_Im)
thus ?thesis
using y_axis[of ?y] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF in_disc not_zero]
by simp
qed
(* ---------------------------------------------------------------------------- *)
subsection \<open>Blaschke factors are unit disc preserving transformations\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>For a given point $a$, Blaschke factor transformations are of the form $k \cdot
\left(\begin{array}{cc}1 & -a\\ -\overline{a} & 1\end{array}\right)$. It is a disc preserving
Möbius transformation that maps the point $a$ to zero (by the symmetry principle, it must map the
inverse point of $a$ to infinity).\<close>
definition blaschke_cmat :: "complex \<Rightarrow> complex_mat" where
[simp]: "blaschke_cmat a = (if cmod a \<noteq> 1 then (1, -a, -cnj a, 1) else eye)"
lift_definition blaschke_mmat :: "complex \<Rightarrow> moebius_mat" is blaschke_cmat
by simp
lift_definition blaschke :: "complex \<Rightarrow> moebius" is blaschke_mmat
done
lemma blaschke_0_id [simp]: "blaschke 0 = id_moebius"
by (transfer, transfer, simp)
lemma blaschke_a_to_zero [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) (of_complex a) = 0\<^sub>h"
using assms
by (transfer, transfer, simp)
lemma blaschke_inv_a_inf [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) (inversion (of_complex a)) = \<infinity>\<^sub>h"
using assms
unfolding inversion_def
by (transfer, transfer) (simp add: vec_cnj_def, rule_tac x="1/(1 - a*cnj a)" in exI, simp)
lemma blaschke_inf [simp]:
assumes "cmod a < 1" and "a \<noteq> 0"
shows "moebius_pt (blaschke a) \<infinity>\<^sub>h = of_complex (- 1 / cnj a)"
using assms
by (transfer, transfer, simp add: complex_mod_sqrt_Re_mult_cnj)
lemma blaschke_0_minus_a [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) 0\<^sub>h = ~\<^sub>h (of_complex a)"
using assms
by (transfer, transfer, simp)
lemma blaschke_unit_circle_fix [simp]:
assumes "cmod a \<noteq> 1"
shows "unit_circle_fix (blaschke a)"
using assms
by (transfer, transfer) (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma blaschke_unit_disc_fix [simp]:
assumes "cmod a < 1"
shows "unit_disc_fix (blaschke a)"
using assms
proof (transfer, transfer)
fix a
assume *: "cmod a < 1"
show "unit_disc_fix_cmat (blaschke_cmat a)"
proof (cases "a = 0")
case True
thus ?thesis
by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
next
case False
hence "Re (a * cnj a) < 1"
using *
by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff)
hence "1 / Re (a * cnj a) > 1"
using False
by (smt complex_div_gt_0 less_divide_eq_1_pos one_complex.simps(1) right_inverse_eq)
hence "Re (1 / (a * cnj a)) > 1"
by (simp add: complex_is_Real_iff)
thus ?thesis
by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
qed
qed
lemma blaschke_unit_circle_fix':
assumes "cmod a \<noteq> 1"
shows "moebius_circline (blaschke a) unit_circle = unit_circle"
using assms
using blaschke_unit_circle_fix unit_circle_fix_iff
by simp
lemma blaschke_ounit_circle_fix':
assumes "cmod a < 1"
shows "moebius_ocircline (blaschke a) ounit_circle = ounit_circle"
proof-
have "Re (a * cnj a) < 1"
using assms
by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff)
thus ?thesis
using assms
using blaschke_unit_disc_fix unit_disc_fix_iff_ounit_circle
by simp
qed
lemma moebius_pt_blaschke [simp]:
assumes "cmod a \<noteq> 1" and "z \<noteq> 1 / cnj a"
shows "moebius_pt (blaschke a) (of_complex z) = of_complex ((z - a) / (1 - cnj a * z))"
using assms
proof (cases "a = 0")
case True
thus ?thesis
by auto
next
case False
thus ?thesis
using assms
apply (transfer, transfer)
apply (simp add: complex_mod_sqrt_Re_mult_cnj)
apply (rule_tac x="1 / (1 - cnj a * z)" in exI)
apply (simp add: field_simps)
done
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Blaschke factors for a real point $a$\<close>
(* -------------------------------------------------------------------------- *)
text \<open>If the point $a$ is real, the Blaschke factor preserve x-axis and the upper and the lower
halfplane.\<close>
lemma blaschke_real_preserve_x_axis [simp]:
assumes "is_real a" and "cmod a < 1"
shows "moebius_pt (blaschke a) z \<in> circline_set x_axis \<longleftrightarrow> z \<in> circline_set x_axis"
proof (cases "a = 0")
case True
thus ?thesis
by simp
next
case False
have "cmod a \<noteq> 1"
using assms
by linarith
let ?a = "of_complex a"
let ?i = "inversion ?a"
let ?M = "moebius_pt (blaschke a)"
have *: "?M ?a = 0\<^sub>h" "?M ?i = \<infinity>\<^sub>h" "?M 0\<^sub>h = of_complex (-a)"
using \<open>cmod a \<noteq> 1\<close> blaschke_a_to_zero[of a] blaschke_inv_a_inf[of a] blaschke_0_minus_a[of a]
by auto
let ?Mx = "moebius_circline (blaschke a) x_axis"
have "?a \<in> circline_set x_axis" "?i \<in> circline_set x_axis" "0\<^sub>h \<in> circline_set x_axis"
using \<open>is_real a\<close> \<open>a \<noteq> 0\<close> eq_cnj_iff_real[of a]
by auto
hence "0\<^sub>h \<in> circline_set ?Mx" "\<infinity>\<^sub>h \<in> circline_set ?Mx" "of_complex (-a) \<in> circline_set ?Mx"
using *
apply -
apply (force simp add: image_iff)+
apply (simp add: image_iff, rule_tac x="0\<^sub>h" in bexI, simp_all)
done
moreover
have "0\<^sub>h \<in> circline_set x_axis" "\<infinity>\<^sub>h \<in> circline_set x_axis" "of_complex (-a) \<in> circline_set x_axis"
using \<open>is_real a\<close>
by auto
moreover
have "of_complex (-a) \<noteq> 0\<^sub>h"
using \<open>a \<noteq> 0\<close>
by simp
hence "0\<^sub>h \<noteq> of_complex (-a)"
by metis
hence "\<exists>!H. 0\<^sub>h \<in> circline_set H \<and> \<infinity>\<^sub>h \<in> circline_set H \<and> of_complex (- a) \<in> circline_set H"
using unique_circline_set[of "0\<^sub>h" "\<infinity>\<^sub>h" "of_complex (-a)"] \<open>a \<noteq> 0\<close>
by simp
ultimately
have "moebius_circline (blaschke a) x_axis = x_axis"
by auto
thus ?thesis
by (metis circline_set_moebius_circline_iff)
qed
lemma blaschke_real_preserve_sgn_Im [simp]:
assumes "is_real a" and "cmod a < 1" and "z \<noteq> \<infinity>\<^sub>h" and "z \<noteq> inversion (of_complex a)"
shows "sgn (Im (to_complex (moebius_pt (blaschke a) z))) = sgn (Im (to_complex z))"
proof (cases "a = 0")
case True
thus ?thesis
by simp
next
case False
obtain z' where z': "z = of_complex z'"
using inf_or_of_complex[of z] \<open>z \<noteq> \<infinity>\<^sub>h\<close>
by auto
have "z' \<noteq> 1 / cnj a"
using assms z' \<open>a \<noteq> 0\<close>
by (auto simp add: of_complex_inj)
moreover
have "a * cnj a \<noteq> 1"
using \<open>cmod a < 1\<close>
by auto (simp add: complex_mod_sqrt_Re_mult_cnj)
moreover
have "sgn (Im ((z' - a) / (1 - a * z'))) = sgn (Im z')"
proof-
have "a * z' \<noteq> 1"
using \<open>is_real a\<close> \<open>z' \<noteq> 1 / cnj a\<close> \<open>a \<noteq> 0\<close> eq_cnj_iff_real[of a]
by (simp add: field_simps)
moreover
have "Re (1 - a\<^sup>2) > 0"
using \<open>is_real a\<close> \<open>cmod a < 1\<close>
by (smt Re_power2 minus_complex.simps(1) norm_complex_def one_complex.simps(1) power2_less_0 real_sqrt_lt_1_iff)
moreover
have "Im ((z' - a) / (1 - a * z')) = Re (((1 - a\<^sup>2) * Im z') / (cmod (1 - a*z'))\<^sup>2)"
proof-
have "1 - a * cnj z' \<noteq> 0"
using \<open>z' \<noteq> 1 / cnj a\<close>
by (metis Im_complex_div_eq_0 complex_cnj_zero_iff diff_eq_diff_eq diff_numeral_special(9) eq_divide_imp is_real_div mult_not_zero one_complex.simps(2) zero_neq_one)
hence "Im ((z' - a) / (1 - a * z')) = Im (((z' - a) * (1 - a * cnj z')) / ((1 - a * z') * cnj (1 - a * z')))"
using \<open>is_real a\<close> eq_cnj_iff_real[of a]
by simp
also have "... = Im ((z' - a - a * z' * cnj z' + a\<^sup>2 * cnj z') / (cmod (1 - a*z'))\<^sup>2)"
unfolding complex_mult_cnj_cmod
by (simp add: power2_eq_square field_simps)
finally show ?thesis
using \<open>is_real a\<close>
by (simp add: field_simps)
qed
moreover
have "0 < (1 - (Re a)\<^sup>2) * Im z' / (cmod (1 - a * z'))\<^sup>2 \<Longrightarrow> Im z' > 0"
- using `is_real a` `0 < Re (1 - a\<^sup>2)`
+ using \<open>is_real a\<close> \<open>0 < Re (1 - a\<^sup>2)\<close>
by (smt Re_power_real divide_le_0_iff minus_complex.simps(1) not_sum_power2_lt_zero one_complex.simps(1) zero_less_mult_pos)
ultimately
show ?thesis
unfolding sgn_real_def
using \<open>cmod a < 1\<close> \<open>a * z' \<noteq> 1\<close> \<open>is_real a\<close>
by (auto simp add: cmod_eq_Re)
qed
ultimately
show ?thesis
using assms z' moebius_pt_blaschke[of a z'] \<open>is_real a\<close> eq_cnj_iff_real[of a]
by simp
qed
lemma blaschke_real_preserve_sgn_arg [simp]:
assumes "is_real a" and "cmod a < 1" and "z \<notin> circline_set x_axis"
shows "sgn (arg (to_complex (moebius_pt (blaschke a) z))) = sgn (arg (to_complex z))"
proof-
have "z \<noteq> \<infinity>\<^sub>h"
using assms
using special_points_on_x_axis''(3) by blast
moreover
have "z \<noteq> inversion (of_complex a)"
using assms
by (metis calculation circline_equation_x_axis circline_set_x_axis_I conjugate_of_complex inversion_of_complex inversion_sym is_real_inversion o_apply of_complex_zero reciprocal_zero to_complex_of_complex)
ultimately
show ?thesis
using blaschke_real_preserve_sgn_Im[OF assms(1) assms(2), of z]
by (smt arg_Im_sgn assms(3) circline_set_x_axis_I norm_sgn of_complex_to_complex)
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Inverse Blaschke transform\<close>
(* -------------------------------------------------------------------------- *)
definition inv_blaschke_cmat :: "complex \<Rightarrow> complex_mat" where
[simp]: "inv_blaschke_cmat a = (if cmod a \<noteq> 1 then (1, a, cnj a, 1) else eye)"
lift_definition inv_blaschke_mmat :: "complex \<Rightarrow> moebius_mat" is inv_blaschke_cmat
by simp
lift_definition inv_blaschke :: "complex \<Rightarrow> moebius" is inv_blaschke_mmat
done
lemma inv_blaschke_neg [simp]: "inv_blaschke a = blaschke (-a)"
by (transfer, transfer) simp
lemma inv_blaschke:
assumes "cmod a \<noteq> 1"
shows "blaschke a + inv_blaschke a = 0"
apply simp
apply (transfer, transfer)
by auto (rule_tac x="1/(1 - a*cnj a)" in exI, simp)
lemma ex_unit_disc_fix_mapping_u_to_zero:
assumes "u \<in> unit_disc"
shows "\<exists> M. unit_disc_fix M \<and> moebius_pt M u = 0\<^sub>h"
proof-
from assms obtain c where *: "u = of_complex c"
by (metis inf_notin_unit_disc inf_or_of_complex)
hence "cmod c < 1"
using assms unit_disc_iff_cmod_lt_1
by simp
thus ?thesis
using *
by (rule_tac x="blaschke c" in exI)
(smt blaschke_a_to_zero blaschke_ounit_circle_fix' unit_disc_fix_iff_ounit_circle)
qed
lemma wlog_zero:
assumes in_disc: "u \<in> unit_disc"
assumes preserving: "\<And> a u. \<lbrakk>u \<in> unit_disc; cmod a < 1; P (moebius_pt (blaschke a) u)\<rbrakk> \<Longrightarrow> P u"
assumes zero: "P 0\<^sub>h"
shows "P u"
proof-
have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h"
by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1)
thus ?thesis
using preserving[of u "to_complex u"] in_disc zero
using inf_or_of_complex[of u]
by auto
qed
lemma wlog_real_zero:
assumes in_disc: "u \<in> unit_disc" and real: "is_real (to_complex u)"
assumes preserving: "\<And> a u. \<lbrakk>u \<in> unit_disc; is_real a; cmod a < 1; P (moebius_pt (blaschke a) u)\<rbrakk> \<Longrightarrow> P u"
assumes zero: "P 0\<^sub>h"
shows "P u"
proof-
have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h"
by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1)
thus ?thesis
using preserving[of u "to_complex u"] in_disc zero real
using inf_or_of_complex[of u]
by auto
qed
lemma unit_disc_fix_transitive:
assumes in_disc: "u \<in> unit_disc" and "u' \<in> unit_disc"
shows "\<exists> M. unit_disc_fix M \<and> moebius_pt M u = u'"
proof-
have "\<forall> u \<in> unit_disc. \<exists> M. unit_disc_fix M \<and> moebius_pt M u = u'" (is "?P u'")
proof (rule wlog_zero)
show "u' \<in> unit_disc" by fact
next
show "?P 0\<^sub>h"
by (simp add: ex_unit_disc_fix_mapping_u_to_zero)
next
fix a u
assume "cmod a < 1" and *: "?P (moebius_pt (blaschke a) u)"
show "?P u"
proof
fix u'
assume "u' \<in> unit_disc"
then obtain M' where "unit_disc_fix M'" "moebius_pt M' u' = moebius_pt (blaschke a) u"
using *
by auto
thus "\<exists>M. unit_disc_fix M \<and> moebius_pt M u' = u"
using \<open>cmod a < 1\<close> blaschke_unit_disc_fix[of a]
using unit_disc_fix_moebius_comp[of "- blaschke a" "M'"]
using unit_disc_fix_moebius_inv[of "blaschke a"]
by (rule_tac x="(- (blaschke a)) + M'" in exI, simp)
qed
qed
thus ?thesis
using assms
by auto
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>Decomposition of unit disc preserving Möbius transforms\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Each transformation preserving unit disc can be decomposed to a rotation around the origin and
a Blaschke factors that maps a point within the unit disc to zero.\<close>
lemma unit_disc_fix_decompose_blaschke_rotation:
assumes "unit_disc_fix M"
shows "\<exists> k \<phi>. cmod k < 1 \<and> M = moebius_rotation \<phi> + blaschke k"
using assms
unfolding moebius_rotation_def moebius_similarity_def
proof (simp, transfer, transfer)
fix M
assume *: "mat_det M \<noteq> 0" "unit_disc_fix_cmat M"
then obtain k a b :: complex where
**: "k \<noteq> 0" "mat_det (a, b, cnj b, cnj a) \<noteq> 0" "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)"
using unitary11_gen_iff[of M]
by auto
have "a \<noteq> 0"
using * **
by auto
then obtain a' k' \<phi>
where ***: "k' \<noteq> 0 \<and> a' * cnj a' \<noteq> 1 \<and> M = k' *\<^sub>s\<^sub>m (cis \<phi>, 0, 0, 1) *\<^sub>m\<^sub>m (1, - a', - cnj a', 1)"
using ** unitary11_gen_cis_blaschke[of k M a b]
by auto blast
have "a' = 0 \<or> 1 < 1 / (cmod a')\<^sup>2"
using * *** complex_mult_cnj_cmod[of a']
by simp
hence "cmod a' < 1"
by (smt less_divide_eq_1_pos norm_zero one_less_power one_power2 pos2)
thus "\<exists>k. cmod k < 1 \<and>
(\<exists>\<phi>. moebius_cmat_eq M (moebius_comp_cmat (mk_moebius_cmat (cis \<phi>) 0 0 1) (blaschke_cmat k)))"
using ***
apply (rule_tac x=a' in exI)
apply simp
apply (rule_tac x=\<phi> in exI)
apply simp
apply (rule_tac x="1/k'" in exI)
by auto
qed
lemma wlog_unit_disc_fix:
assumes "unit_disc_fix M"
assumes b: "\<And> k. cmod k < 1 \<Longrightarrow> P (blaschke k)"
assumes r: "\<And> \<phi>. P (moebius_rotation \<phi>)"
assumes comp: "\<And>M1 M2. \<lbrakk>unit_disc_fix M1; P M1; unit_disc_fix M2; P M2\<rbrakk> \<Longrightarrow> P (M1 + M2)"
shows "P M"
using assms
using unit_disc_fix_decompose_blaschke_rotation[OF assms(1)]
using blaschke_unit_disc_fix
by auto
lemma ex_unit_disc_fix_to_zero_positive_x_axis:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> v"
shows "\<exists> M. unit_disc_fix M \<and>
moebius_pt M u = 0\<^sub>h \<and> moebius_pt M v \<in> positive_x_axis"
proof-
from assms obtain B where
*: "unit_disc_fix B" "moebius_pt B u = 0\<^sub>h"
using ex_unit_disc_fix_mapping_u_to_zero
by blast
let ?v = "moebius_pt B v"
have "?v \<in> unit_disc"
using \<open>v \<in> unit_disc\<close> *
by auto
hence "?v \<noteq> \<infinity>\<^sub>h"
using inf_notin_unit_disc by auto
have "?v \<noteq> 0\<^sub>h"
using \<open>u \<noteq> v\<close> *
by (metis moebius_pt_invert)
obtain R where
"unit_disc_fix R"
"moebius_pt R 0\<^sub>h = 0\<^sub>h" "moebius_pt R ?v \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of ?v] \<open>?v \<noteq> 0\<^sub>h\<close> \<open>?v \<noteq> \<infinity>\<^sub>h\<close>
using moebius_pt_rotation_inf_iff moebius_pt_moebius_rotation_zero unit_disc_fix_rotation
by blast
thus ?thesis
using * moebius_comp[of R B, symmetric]
using unit_disc_fix_moebius_comp
by (rule_tac x="R + B" in exI) (simp add: comp_def)
qed
lemma wlog_x_axis:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc"
assumes preserved: "\<And> M u v. \<lbrakk>unit_disc_fix M; u \<in> unit_disc; v \<in> unit_disc; P (moebius_pt M u) (moebius_pt M v)\<rbrakk> \<Longrightarrow> P u v"
assumes axis: "\<And> x. \<lbrakk>is_real x; 0 \<le> Re x; Re x < 1\<rbrakk> \<Longrightarrow> P 0\<^sub>h (of_complex x)"
shows "P u v"
proof (cases "u = v")
case True
have "P u u" (is "?Q u")
proof (rule wlog_zero[where P="?Q"])
show "u \<in> unit_disc"
by fact
next
show "?Q 0\<^sub>h"
using axis[of 0]
by simp
next
fix a u
assume "u \<in> unit_disc" "cmod a < 1" "?Q (moebius_pt (blaschke a) u)"
thus "?Q u"
using preserved[of "blaschke a" u u]
using blaschke_unit_disc_fix[of a]
by simp
qed
thus ?thesis
using True
by simp
next
case False
from in_disc obtain M where
*: "unit_disc_fix M" "moebius_pt M u = 0\<^sub>h" "moebius_pt M v \<in> positive_x_axis"
using ex_unit_disc_fix_to_zero_positive_x_axis False
by auto
then obtain x where **: "moebius_pt M v = of_complex x" "is_real x"
unfolding positive_x_axis_def circline_set_x_axis
by auto
moreover
have "of_complex x \<in> unit_disc"
using \<open>unit_disc_fix M\<close> \<open>v \<in> unit_disc\<close> **
using unit_disc_fix_discI
by fastforce
hence "0 < Re x" "Re x < 1"
using \<open>moebius_pt M v \<in> positive_x_axis\<close> **
by (auto simp add: positive_x_axis_def cmod_eq_Re)
ultimately
have "P 0\<^sub>h (of_complex x)"
using \<open>is_real x\<close> axis
by auto
thus ?thesis
using preserved[OF *(1) assms(1-2)] *(2) **(1)
by simp
qed
lemma wlog_positive_x_axis:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes preserved: "\<And> M u v. \<lbrakk>unit_disc_fix M; u \<in> unit_disc; v \<in> unit_disc; u \<noteq> v; P (moebius_pt M u) (moebius_pt M v)\<rbrakk> \<Longrightarrow> P u v"
assumes axis: "\<And> x. \<lbrakk>is_real x; 0 < Re x; Re x < 1\<rbrakk> \<Longrightarrow> P 0\<^sub>h (of_complex x)"
shows "P u v"
proof-
have "u \<noteq> v \<longrightarrow> P u v" (is "?Q u v")
proof (rule wlog_x_axis)
show "u \<in> unit_disc" "v \<in> unit_disc"
by fact+
next
fix M u v
assume "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
"?Q (moebius_pt M u) (moebius_pt M v)"
thus "?Q u v"
using preserved[of M u v]
using moebius_pt_invert
by blast
next
fix x
assume "is_real x" "0 \<le> Re x" "Re x < 1"
thus "?Q 0\<^sub>h (of_complex x)"
using axis[of x] of_complex_zero_iff[of x] complex.expand[of x 0]
by fastforce
qed
thus ?thesis
using \<open>u \<noteq> v\<close>
by simp
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>All functions that fix the unit disc\<close>
(* -------------------------------------------------------------------------- *)
text \<open>It can be proved that continuous functions that fix the unit disc are either actions of
Möbius transformations that fix the unit disc (homographies), or are compositions of actions of
Möbius transformations that fix the unit disc and the conjugation (antihomographies). We postulate
this as a definition, but it this characterisation could also be formally shown (we do not need this
for our further applications).\<close>
definition unit_disc_fix_f where
"unit_disc_fix_f f \<longleftrightarrow>
(\<exists> M. unit_disc_fix M \<and> (f = moebius_pt M \<or> f = moebius_pt M \<circ> conjugate))"
text \<open>Unit disc fixing functions really fix unit disc.\<close>
lemma unit_disc_fix_f_unit_disc:
assumes "unit_disc_fix_f M"
shows "M ` unit_disc = unit_disc"
using assms
unfolding unit_disc_fix_f_def
using image_comp
by force
text \<open>Actions of unit disc fixing Möbius transformations (unit disc fixing homographies) are unit
disc fixing functions.\<close>
lemma unit_disc_fix_f_moebius_pt [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix_f (moebius_pt M)"
using assms
unfolding unit_disc_fix_f_def
by auto
text \<open>Compositions of unit disc fixing Möbius transformations and conjugation (unit disc fixing
antihomographies) are unit disc fixing functions.\<close>
lemma unit_disc_fix_conjugate_moebius [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (conjugate_moebius M)"
proof-
have "\<And>a aa ab b. \<lbrakk>1 < Re (a * b / (aa * ab)); \<not> 1 < Re (cnj a * cnj b / (cnj aa * cnj ab))\<rbrakk> \<Longrightarrow> aa = 0"
by (metis cnj.simps(1) complex_cnj_divide complex_cnj_mult)
thus ?thesis
using assms
by (transfer, transfer)
(auto simp add: mat_cnj_def unitary11_gen_def mat_adj_def field_simps)
qed
lemma unit_disc_fix_conjugate_comp_moebius [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix_f (conjugate \<circ> moebius_pt M)"
using assms
apply (subst conjugate_moebius)
apply (simp add: unit_disc_fix_f_def)
apply (rule_tac x="conjugate_moebius M" in exI, simp)
done
text \<open>Uniti disc fixing functions form a group under function composition.\<close>
lemma unit_disc_fix_f_comp [simp]:
assumes "unit_disc_fix_f f1" and "unit_disc_fix_f f2"
shows "unit_disc_fix_f (f1 \<circ> f2)"
using assms
apply (subst (asm) unit_disc_fix_f_def)
apply (subst (asm) unit_disc_fix_f_def)
proof safe
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f (moebius_pt M \<circ> moebius_pt M')"
unfolding unit_disc_fix_f_def
by (rule_tac x="M + M'" in exI) auto
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f (moebius_pt M \<circ> (moebius_pt M' \<circ> conjugate))"
unfolding unit_disc_fix_f_def
by (subst comp_assoc[symmetric])+
(rule_tac x="M + M'" in exI, auto)
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f ((moebius_pt M \<circ> conjugate) \<circ> moebius_pt M')"
unfolding unit_disc_fix_f_def
by (subst comp_assoc, subst conjugate_moebius, subst comp_assoc[symmetric])+
(rule_tac x="M + conjugate_moebius M'" in exI, auto)
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f ((moebius_pt M \<circ> conjugate) \<circ> (moebius_pt M' \<circ> conjugate))"
apply (subst comp_assoc[symmetric], subst comp_assoc)
apply (subst conjugate_moebius, subst comp_assoc, subst comp_assoc)
apply (simp add: unit_disc_fix_f_def)
apply (rule_tac x="M + conjugate_moebius M'" in exI, auto)
done
qed
lemma unit_disc_fix_f_inv:
assumes "unit_disc_fix_f M"
shows "unit_disc_fix_f (inv M)"
using assms
apply (subst (asm) unit_disc_fix_f_def)
proof safe
fix M
assume "unit_disc_fix M"
have "inv (moebius_pt M) = moebius_pt (-M)"
by (rule ext) (simp add: moebius_inv)
thus "unit_disc_fix_f (inv (moebius_pt M))"
using \<open>unit_disc_fix M\<close>
unfolding unit_disc_fix_f_def
by (rule_tac x="-M" in exI, simp)
next
fix M
assume "unit_disc_fix M"
have "inv (moebius_pt M \<circ> conjugate) = conjugate \<circ> inv (moebius_pt M)"
by (subst o_inv_distrib, simp_all)
also have "... = conjugate \<circ> (moebius_pt (-M))"
using moebius_inv
by auto
also have "... = moebius_pt (conjugate_moebius (-M)) \<circ> conjugate"
by (simp add: conjugate_moebius)
finally
show "unit_disc_fix_f (inv (moebius_pt M \<circ> conjugate))"
using \<open>unit_disc_fix M\<close>
unfolding unit_disc_fix_f_def
by (rule_tac x="conjugate_moebius (-M)" in exI, simp)
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Action of unit disc fixing functions on circlines\<close>
(* -------------------------------------------------------------------------- *)
definition unit_disc_fix_f_circline where
"unit_disc_fix_f_circline f H =
(if \<exists> M. unit_disc_fix M \<and> f = moebius_pt M then
moebius_circline (THE M. unit_disc_fix M \<and> f = moebius_pt M) H
else if \<exists> M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate then
(moebius_circline (THE M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate) \<circ> conjugate_circline) H
else
H)"
lemma unique_moebius_pt_conjugate:
assumes "moebius_pt M1 \<circ> conjugate = moebius_pt M2 \<circ> conjugate"
shows "M1 = M2"
proof-
from assms have "moebius_pt M1 = moebius_pt M2"
using conjugate_conjugate_comp rewriteL_comp_comp2 by fastforce
thus ?thesis
using unique_moebius_pt
by auto
qed
lemma unit_disc_fix_f_circline_direct:
assumes "unit_disc_fix M" and "f = moebius_pt M"
shows "unit_disc_fix_f_circline f H = moebius_circline M H"
proof-
have "M = (THE M. unit_disc_fix M \<and> f = moebius_pt M)"
using assms
using theI_unique[of "\<lambda> M. unit_disc_fix M \<and> f = moebius_pt M" M]
using unique_moebius_pt[of M]
by auto
thus ?thesis
using assms
unfolding unit_disc_fix_f_circline_def
by auto
qed
lemma unit_disc_fix_f_circline_indirect:
assumes "unit_disc_fix M" and "f = moebius_pt M \<circ> conjugate"
shows "unit_disc_fix_f_circline f H = ((moebius_circline M) \<circ> conjugate_circline) H"
proof-
have "\<not> (\<exists> M. unit_disc_fix M \<and> f = moebius_pt M)"
using assms homography_antihomography_exclusive[of f]
unfolding is_homography_def is_antihomography_def is_moebius_def
by auto
moreover
have "M = (THE M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate)"
using assms
using theI_unique[of "\<lambda> M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate" M]
using unique_moebius_pt_conjugate[of M]
by auto
ultimately
show ?thesis
using assms
unfolding unit_disc_fix_f_circline_def
by metis
qed
text \<open>Disc automorphisms - it would be nice to show that there are no disc automorphisms other than
unit disc fixing homographies and antihomographies, but this part of the theory is not yet
developed.\<close>
definition is_disc_aut where "is_disc_aut f \<longleftrightarrow> bij_betw f unit_disc unit_disc"
end
\ No newline at end of file
diff --git a/thys/Farkas/Farkas.thy b/thys/Farkas/Farkas.thy
--- a/thys/Farkas/Farkas.thy
+++ b/thys/Farkas/Farkas.thy
@@ -1,1633 +1,1633 @@
(* Authors: R. Bottesch, M. W. Haslbeck, R. Thiemann *)
section \<open>Farkas Coefficients via the Simplex Algorithm of Duterte and de~Moura\<close>
text \<open>Let $c_1,\dots,c_n$ be a finite list of linear inequalities.
Let $C$ be a list of pairs $(r_i,c_i)$ where $r_i$ is a rational number.
We say that $C$ is a list of \emph{Farkas coefficients} if
the sum of all products $r_i \cdot c_i$ results in an inequality that is
trivially unsatisfiable.
Farkas' Lemma
states that a finite set of non-strict linear inequalities is
unsatisfiable if and only if Farkas coefficients exist. We will prove this lemma
with the help of the simplex algorithm of Dutertre and de~Moura's.
Note that the simplex implementation works on four layers, and we will formulate and prove
a variant of Farkas' Lemma for each of these layers.\<close>
theory Farkas
imports Simplex.Simplex
begin
subsection \<open>Linear Inequalities\<close>
text \<open>Both Farkas' Lemma and Motzkin's Transposition Theorem require linear combinations
of inequalities. To this end we define one type that permits strict and non-strict
inequalities which are always of the form ``polynomial R constant'' where R is either
$\leq$ or $<$. On this type we can then define a commutative monoid.\<close>
text \<open>A type for the two relations: less-or-equal and less-than.\<close>
datatype le_rel = Leq_Rel | Lt_Rel
primrec rel_of :: "le_rel \<Rightarrow> 'a :: lrv \<Rightarrow> 'a \<Rightarrow> bool" where
"rel_of Leq_Rel = (\<le>)"
| "rel_of Lt_Rel = (<)"
instantiation le_rel :: comm_monoid_add begin
definition "zero_le_rel = Leq_Rel"
fun plus_le_rel where
"plus_le_rel Leq_Rel Leq_Rel = Leq_Rel"
| "plus_le_rel _ _ = Lt_Rel"
instance
proof
fix a b c :: le_rel
show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto)
show "a + b = b + a" by (cases a; cases b, auto)
show "0 + a = a" unfolding zero_le_rel_def by (cases a, auto)
qed
end
lemma Leq_Rel_0: "Leq_Rel = 0" unfolding zero_le_rel_def by simp
datatype 'a le_constraint = Le_Constraint (lec_rel: le_rel) (lec_poly: linear_poly) (lec_const: 'a)
abbreviation (input) "Leqc \<equiv> Le_Constraint Leq_Rel"
instantiation le_constraint :: (lrv) comm_monoid_add begin
fun plus_le_constraint :: "'a le_constraint \<Rightarrow> 'a le_constraint \<Rightarrow> 'a le_constraint" where
"plus_le_constraint (Le_Constraint r1 p1 c1) (Le_Constraint r2 p2 c2) =
(Le_Constraint (r1 + r2) (p1 + p2) (c1 + c2))"
definition zero_le_constraint :: "'a le_constraint" where
"zero_le_constraint = Leqc 0 0"
instance proof
fix a b c :: "'a le_constraint"
show "0 + a = a"
by (cases a, auto simp: zero_le_constraint_def Leq_Rel_0)
show "a + b = b + a" by (cases a; cases b, auto simp: ac_simps)
show "a + b + c = a + (b + c)" by (cases a; cases b; cases c, auto simp: ac_simps)
qed
end
primrec satisfiable_le_constraint :: "'a::lrv valuation \<Rightarrow> 'a le_constraint \<Rightarrow> bool" (infixl "\<Turnstile>\<^sub>l\<^sub>e" 100) where
"(v \<Turnstile>\<^sub>l\<^sub>e (Le_Constraint rel l r)) \<longleftrightarrow> (rel_of rel (l\<lbrace>v\<rbrace>) r)"
lemma satisfies_zero_le_constraint: "v \<Turnstile>\<^sub>l\<^sub>e 0"
by (simp add: valuate_zero zero_le_constraint_def)
lemma satisfies_sum_le_constraints:
assumes "v \<Turnstile>\<^sub>l\<^sub>e c" "v \<Turnstile>\<^sub>l\<^sub>e d"
shows "v \<Turnstile>\<^sub>l\<^sub>e (c + d)"
proof -
obtain lc rc ld rd rel1 rel2 where cd: "c = Le_Constraint rel1 lc rc" "d = Le_Constraint rel2 ld rd"
by (cases c; cases d, auto)
have 1: "rel_of rel1 (lc\<lbrace>v\<rbrace>) rc" using assms cd by auto
have 2: "rel_of rel2 (ld\<lbrace>v\<rbrace>) rd" using assms cd by auto
from 1 have le1: "lc\<lbrace>v\<rbrace> \<le> rc" by (cases rel1, auto)
from 2 have le2: "ld\<lbrace>v\<rbrace> \<le> rd" by (cases rel2, auto)
from 1 2 le1 le2 have "rel_of (rel1 + rel2) ((lc\<lbrace>v\<rbrace>) + (ld\<lbrace>v\<rbrace>)) (rc + rd)"
apply (cases rel1; cases rel2; simp add: add_mono)
by (metis add.commute le_less_trans order.strict_iff_order plus_less)+
thus ?thesis by (auto simp: cd valuate_add)
qed
lemma satisfies_sumlist_le_constraints:
assumes "\<And> c. c \<in> set (cs :: 'a :: lrv le_constraint list) \<Longrightarrow> v \<Turnstile>\<^sub>l\<^sub>e c"
shows "v \<Turnstile>\<^sub>l\<^sub>e sum_list cs"
using assms
by (induct cs, auto intro: satisfies_zero_le_constraint satisfies_sum_le_constraints)
lemma sum_list_lec:
"sum_list ls = Le_Constraint
(sum_list (map lec_rel ls))
(sum_list (map lec_poly ls))
(sum_list (map lec_const ls))"
proof (induct ls)
case Nil
show ?case by (auto simp: zero_le_constraint_def Leq_Rel_0)
next
case (Cons l ls)
show ?case by (cases l, auto simp: Cons)
qed
lemma sum_list_Leq_Rel: "((\<Sum>x\<leftarrow>C. lec_rel (f x)) = Leq_Rel) \<longleftrightarrow> (\<forall> x \<in> set C. lec_rel (f x) = Leq_Rel)"
proof (induct C)
case (Cons c C)
show ?case
proof (cases "lec_rel (f c)")
case Leq_Rel
show ?thesis using Cons by (simp add: Leq_Rel Leq_Rel_0)
qed simp
qed (simp add: Leq_Rel_0)
subsection \<open>Farkas' Lemma on Layer 4\<close>
text \<open>On layer 4 the algorithm works on a state containing a tableau, atoms (or bounds),
an assignment and a satisfiability flag. Only non-strict inequalities appear at this level.
In order to even state a variant of Farkas' Lemma on layer 4, we
need conversions from atoms to non-strict constraints and then further
to linear inequalities of type @{type le_constraint}.
The latter conversion is a partial operation, since non-strict constraints
of type @{type ns_constraint} permit greater-or-equal constraints, whereas @{type le_constraint}
allows only less-or-equal.\<close>
text \<open>The advantage of first going via @{type ns_constraint} is that this type permits a multiplication
with arbitrary rational numbers (the direction of the inequality must be flipped when
multiplying by a negative number, which is not possible with @{type le_constraint}).\<close>
instantiation ns_constraint :: (scaleRat) scaleRat
begin
fun scaleRat_ns_constraint :: "rat \<Rightarrow> 'a ns_constraint \<Rightarrow> 'a ns_constraint" where
"scaleRat_ns_constraint r (LEQ_ns p c) =
(if (r < 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))"
| "scaleRat_ns_constraint r (GEQ_ns p c) =
(if (r > 0) then GEQ_ns (r *R p) (r *R c) else LEQ_ns (r *R p) (r *R c))"
instance ..
end
lemma sat_scale_rat_ns: assumes "v \<Turnstile>\<^sub>n\<^sub>s ns"
shows "v \<Turnstile>\<^sub>n\<^sub>s (f *R ns)"
proof -
have "f < 0 | f = 0 | f > 0" by auto
then show ?thesis using assms by (cases ns, auto simp: valuate_scaleRat scaleRat_leq1 scaleRat_leq2)
qed
lemma scaleRat_scaleRat_ns_constraint: assumes "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0"
shows "a *R (b *R (c :: 'a :: lrv ns_constraint)) = (a * b) *R c"
proof -
have "b > 0 \<or> b < 0 \<or> b = 0" by linarith
moreover have "a > 0 \<or> a < 0 \<or> a = 0" by linarith
ultimately show ?thesis using assms
by (elim disjE; cases c, auto simp add: not_le not_less
mult_neg_pos mult_neg_neg mult_nonpos_nonneg mult_nonpos_nonpos mult_nonneg_nonpos mult_pos_neg)
qed
fun lec_of_nsc where
"lec_of_nsc (LEQ_ns p c) = (Leqc p c)"
fun is_leq_ns where
"is_leq_ns (LEQ_ns p c) = True"
| "is_leq_ns (GEQ_ns p c) = False"
lemma lec_of_nsc:
assumes "is_leq_ns c"
shows "(v \<Turnstile>\<^sub>l\<^sub>e lec_of_nsc c) \<longleftrightarrow> (v \<Turnstile>\<^sub>n\<^sub>s c)"
using assms by (cases c, auto)
fun nsc_of_atom where
"nsc_of_atom (Leq var b) = LEQ_ns (lp_monom 1 var) b"
| "nsc_of_atom (Geq var b) = GEQ_ns (lp_monom 1 var) b"
lemma nsc_of_atom: "v \<Turnstile>\<^sub>n\<^sub>s nsc_of_atom a \<longleftrightarrow> v \<Turnstile>\<^sub>a a"
by (cases a, auto)
text \<open>We say that $C$ is a list of Farkas coefficients \emph{for a given tableau $t$ and atom set $as$}, if
it is a list of pairs $(r,a)$ such that $a \in as$, $r$ is non-zero, $r \cdot a$ is a
`less-than-or-equal'-constraint, and the linear combination
of inequalities must result in an inequality of the form $p \leq c$, where $c < 0$ and $t \models
p = 0$.\<close>
definition farkas_coefficients_atoms_tableau where
"farkas_coefficients_atoms_tableau (as :: 'a :: lrv atom set) t C = (\<exists> p c.
(\<forall>(r,a) \<in> set C. a \<in> as \<and> is_leq_ns (r *R nsc_of_atom a) \<and> r \<noteq> 0) \<and>
(\<Sum>(r,a) \<leftarrow> C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c \<and>
c < 0 \<and>
(\<forall> v :: 'a valuation. v \<Turnstile>\<^sub>t t \<longrightarrow>(p\<lbrace>v\<rbrace> = 0)))"
text \<open>We first prove that if the check-function detects a conflict, then
Farkas coefficients do exist for the tableau and atom set for which the
conflict is detected.\<close>
definition bound_atoms :: "('i, 'a) state \<Rightarrow> 'a atom set" ("\<B>\<^sub>A") where
"bound_atoms s = (\<lambda>(v,x). Geq v x) ` (set_of_map (\<B>\<^sub>l s)) \<union>
(\<lambda>(v,x). Leq v x) ` (set_of_map (\<B>\<^sub>u s))"
context PivotUpdateMinVars
begin
lemma farkas_check:
assumes check: "check s' = s" and U: "\<U> s" "\<not> \<U> s'"
and inv: "\<nabla> s'" "\<triangle> (\<T> s')" "\<Turnstile>\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s'" "\<diamond> s'"
and index: "index_valid as s'"
shows "\<exists> C. farkas_coefficients_atoms_tableau (snd ` as) (\<T> s') C"
proof -
let ?Q = "\<lambda> s f p c C. set C \<subseteq> \<B>\<^sub>A s \<and>
distinct C \<and>
(\<forall>a \<in> set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \<and> f (atom_var a) \<noteq> 0) \<and>
(\<Sum>a \<leftarrow> C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc p c \<and>
c < 0 \<and>
(\<forall> v :: 'a valuation. v \<Turnstile>\<^sub>t \<T> s \<longrightarrow>(p\<lbrace>v\<rbrace> = 0))"
let ?P = "\<lambda> s. \<U> s \<longrightarrow> (\<exists> f p c C. ?Q s f p c C)"
have "?P (check s')"
proof (induct rule: check_induct''[OF inv, of ?P])
case (3 s x\<^sub>i dir I)
have dir: "dir = Positive \<or> dir = Negative" by fact
let ?eq = "(eq_for_lvar (\<T> s) x\<^sub>i)"
define X\<^sub>j where "X\<^sub>j = rvars_eq ?eq"
define XL\<^sub>j where "XL\<^sub>j = Abstract_Linear_Poly.vars_list (rhs ?eq)"
have [simp]: "set XL\<^sub>j = X\<^sub>j" unfolding XL\<^sub>j_def X\<^sub>j_def
using set_vars_list by blast
have XL\<^sub>j_distinct: "distinct XL\<^sub>j"
unfolding XL\<^sub>j_def using distinct_vars_list by simp
define A where "A = coeff (rhs ?eq)"
have bounds_id: "\<B>\<^sub>A (set_unsat I s) = \<B>\<^sub>A s" "\<B>\<^sub>u (set_unsat I s) = \<B>\<^sub>u s" "\<B>\<^sub>l (set_unsat I s) = \<B>\<^sub>l s"
by (auto simp: boundsl_def boundsu_def bound_atoms_def)
have t_id: "\<T> (set_unsat I s) = \<T> s" by simp
have u_id: "\<U> (set_unsat I s) = True" by simp
let ?p = "rhs ?eq - lp_monom 1 x\<^sub>i"
have p_eval_zero: "?p \<lbrace> v \<rbrace> = 0" if "v \<Turnstile>\<^sub>t \<T> s" for v :: "'a valuation"
proof -
have eqT: "?eq \<in> set (\<T> s)"
by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars)
have "v \<Turnstile>\<^sub>e ?eq" using that eqT satisfies_tableau_def by blast
also have "?eq = (lhs ?eq, rhs ?eq)" by (cases ?eq, auto)
also have "lhs ?eq = x\<^sub>i" by (simp add: 3(7) eq_for_lvar local.min_lvar_not_in_bounds_lvars)
finally have "v \<Turnstile>\<^sub>e (x\<^sub>i, rhs ?eq)" .
then show ?thesis by (auto simp: satisfies_eq_iff valuate_minus)
qed
have Xj_rvars: "X\<^sub>j \<subseteq> rvars (\<T> s)" unfolding X\<^sub>j_def
using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast
have xi_lvars: "x\<^sub>i \<in> lvars (\<T> s)"
using 3 min_lvar_not_in_bounds_lvars rvars_of_lvar_rvars by blast
have "lvars (\<T> s) \<inter> rvars (\<T> s) = {}"
using 3 normalized_tableau_def by auto
with xi_lvars Xj_rvars have xi_Xj: "x\<^sub>i \<notin> X\<^sub>j"
by blast
have rhs_eval_xi: "(rhs (eq_for_lvar (\<T> s) x\<^sub>i)) \<lbrace>\<langle>\<V> s\<rangle>\<rbrace> = \<langle>\<V> s\<rangle> x\<^sub>i"
proof -
have *: "(rhs eq) \<lbrace> v \<rbrace> = v (lhs eq)" if "v \<Turnstile>\<^sub>e eq" for v :: "'a valuation" and eq
using satisfies_eq_def that by metis
moreover have "\<langle>\<V> s\<rangle> \<Turnstile>\<^sub>e eq_for_lvar (\<T> s) x\<^sub>i"
using 3 satisfies_tableau_def eq_for_lvar curr_val_satisfies_no_lhs_def xi_lvars
by blast
ultimately show ?thesis
using eq_for_lvar xi_lvars by simp
qed
let ?\<B>\<^sub>l = "Direction.LB dir"
let ?\<B>\<^sub>u = "Direction.UB dir"
let ?lt = "Direction.lt dir"
let ?le = "Simplex.le ?lt"
let ?Geq = "Direction.GE dir"
let ?Leq = "Direction.LE dir"
have 0: "(if A x < 0 then ?\<B>\<^sub>l s x = Some (\<langle>\<V> s\<rangle> x) else ?\<B>\<^sub>u s x = Some (\<langle>\<V> s\<rangle> x)) \<and> A x \<noteq> 0"
if x: "x \<in> X\<^sub>j" for x
proof -
have "Some (\<langle>\<V> s\<rangle> x) = (?\<B>\<^sub>l s x)" if "A x < 0"
proof -
have cmp: "\<not> \<rhd>\<^sub>l\<^sub>b ?lt (\<langle>\<V> s\<rangle> x) (?\<B>\<^sub>l s x)"
using x that dir min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto
then obtain c where c: "?\<B>\<^sub>l s x = Some c"
by (cases "?\<B>\<^sub>l s x", auto simp: bound_compare_defs)
also have "c = \<langle>\<V> s\<rangle> x"
proof -
have "x \<in> rvars (\<T> s)" using that x Xj_rvars by blast
then have "x \<in> (- lvars (\<T> s))"
using 3 unfolding normalized_tableau_def by auto
moreover have "\<forall>x\<in>(- lvars (\<T> s)). in_bounds x \<langle>\<V> s\<rangle> (\<B>\<^sub>l s, \<B>\<^sub>u s)"
using 3 unfolding curr_val_satisfies_no_lhs_def
by (simp add: satisfies_bounds_set.simps)
ultimately have "in_bounds x \<langle>\<V> s\<rangle> (\<B>\<^sub>l s, \<B>\<^sub>u s)"
by blast
moreover have "?le (\<langle>\<V> s\<rangle> x) c"
using cmp c dir unfolding bound_compare_defs by auto
ultimately show ?thesis
using c dir by auto
qed
then show ?thesis
using c by simp
qed
moreover have "Some (\<langle>\<V> s\<rangle> x) = (?\<B>\<^sub>u s x)" if "0 < A x"
proof -
have cmp: "\<not> \<lhd>\<^sub>u\<^sub>b ?lt (\<langle>\<V> s\<rangle> x) (?\<B>\<^sub>u s x)"
using x that min_rvar_incdec_eq_None[OF 3(9)] unfolding X\<^sub>j_def A_def by auto
then obtain c where c: "?\<B>\<^sub>u s x = Some c"
by (cases "?\<B>\<^sub>u s x", auto simp: bound_compare_defs)
also have "c = \<langle>\<V> s\<rangle> x"
proof -
have "x \<in> rvars (\<T> s)" using that x Xj_rvars by blast
then have "x \<in> (- lvars (\<T> s))"
using 3 unfolding normalized_tableau_def by auto
moreover have "\<forall>x\<in>(- lvars (\<T> s)). in_bounds x \<langle>\<V> s\<rangle> (\<B>\<^sub>l s, \<B>\<^sub>u s)"
using 3 unfolding curr_val_satisfies_no_lhs_def
by (simp add: satisfies_bounds_set.simps)
ultimately have "in_bounds x \<langle>\<V> s\<rangle> (\<B>\<^sub>l s, \<B>\<^sub>u s)"
by blast
moreover have "?le c (\<langle>\<V> s\<rangle> x)"
using cmp c dir unfolding bound_compare_defs by auto
ultimately show ?thesis
using c dir by auto
qed
then show ?thesis
using c by simp
qed
moreover have "A x \<noteq> 0"
using that coeff_zero unfolding A_def X\<^sub>j_def by auto
ultimately show ?thesis
using that by auto
qed
have l_Ba: "l \<in> \<B>\<^sub>A s" if "l \<in> {?Geq x\<^sub>i (the (?\<B>\<^sub>l s x\<^sub>i))}" for l
proof -
from that have l: "l = ?Geq x\<^sub>i (the (?\<B>\<^sub>l s x\<^sub>i))" by simp
from 3(8) obtain c where bl': "?\<B>\<^sub>l s x\<^sub>i = Some c"
by (cases "?\<B>\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs)
hence bl: "(x\<^sub>i, c) \<in> set_of_map (?\<B>\<^sub>l s)" unfolding set_of_map_def by auto
show "l \<in> \<B>\<^sub>A s" unfolding l bound_atoms_def using bl bl' dir by auto
qed
let ?negA = "filter (\<lambda> x. A x < 0) XL\<^sub>j"
let ?posA = "filter (\<lambda> x. \<not> A x < 0) XL\<^sub>j"
define neg where "neg = (if dir = Positive then (\<lambda> x :: rat. x) else uminus)"
define negP where "negP = (if dir = Positive then (\<lambda> x :: linear_poly. x) else uminus)"
define nega where "nega = (if dir = Positive then (\<lambda> x :: 'a. x) else uminus)"
from dir have dirn: "dir = Positive \<and> neg = (\<lambda> x. x) \<and> negP = (\<lambda> x. x) \<and> nega = (\<lambda> x. x)
\<or> dir = Negative \<and> neg = uminus \<and> negP = uminus \<and> nega = uminus"
unfolding neg_def negP_def nega_def by auto
define C where "C = map (\<lambda>x. ?Geq x (the (?\<B>\<^sub>l s x))) ?negA
@ map (\<lambda> x. ?Leq x (the (?\<B>\<^sub>u s x))) ?posA
@ [?Geq x\<^sub>i (the (?\<B>\<^sub>l s x\<^sub>i))]"
define f where "f = (\<lambda>x. if x = x\<^sub>i then neg (-1) else neg (A x))"
define c where "c = (\<Sum>x\<leftarrow>C. lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))"
let ?q = "negP ?p"
show ?case unfolding bounds_id t_id u_id
proof (intro exI impI conjI allI)
show "v \<Turnstile>\<^sub>t \<T> s \<Longrightarrow> ?q \<lbrace> v \<rbrace> = 0" for v :: "'a valuation" using dirn p_eval_zero[of v]
by (auto simp: valuate_minus)
show "set C \<subseteq> \<B>\<^sub>A s"
unfolding C_def set_append set_map set_filter list.simps using 0 l_Ba dir
by (intro Un_least subsetI) (force simp: bound_atoms_def set_of_map_def)+
show is_leq: "\<forall>a\<in>set C. is_leq_ns (f (atom_var a) *R nsc_of_atom a) \<and> f (atom_var a) \<noteq> 0"
using dirn xi_Xj 0 unfolding C_def f_def
by (elim disjE, auto)
show "(\<Sum>a \<leftarrow> C. lec_of_nsc (f (atom_var a) *R nsc_of_atom a)) = Leqc ?q c"
unfolding sum_list_lec le_constraint.simps map_map o_def
proof (intro conjI)
define scale_poly :: "'a atom \<Rightarrow> linear_poly" where
"scale_poly = (\<lambda>x. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x)))"
have "(\<Sum>x\<leftarrow>C. scale_poly x) =
(\<Sum>x\<leftarrow>?negA. scale_poly (?Geq x (the (?\<B>\<^sub>l s x))))
+ (\<Sum>x\<leftarrow>?posA. scale_poly (?Leq x (the (?\<B>\<^sub>u s x))))
- negP (lp_monom 1 x\<^sub>i)"
unfolding C_def using dirn by (auto simp add: comp_def scale_poly_def f_def)
also have "(\<Sum>x\<leftarrow>?negA. scale_poly (?Geq x (the (?\<B>\<^sub>l s x))))
= (\<Sum>x\<leftarrow> ?negA. negP (A x *R lp_monom 1 x))"
unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto
also have "(\<Sum>x\<leftarrow>?posA. scale_poly (?Leq x (the (?\<B>\<^sub>u s x))))
= (\<Sum>x\<leftarrow> ?posA. negP (A x *R lp_monom 1 x))"
unfolding scale_poly_def f_def using dirn xi_Xj by (subst map_cong) auto
also have "(\<Sum>x\<leftarrow> ?negA. negP (A x *R lp_monom 1 x)) +
(\<Sum>x\<leftarrow> ?posA. negP (A x *R lp_monom 1 x))
= negP (rhs (eq_for_lvar (\<T> s) x\<^sub>i))"
using dirn XL\<^sub>j_distinct coeff_zero
by (elim disjE; intro poly_eqI, auto intro!: poly_eqI simp add: coeff_sum_list A_def X\<^sub>j_def
uminus_sum_list_map[unfolded o_def, symmetric])
finally show "(\<Sum>x\<leftarrow>C. lec_poly (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = ?q"
unfolding scale_poly_def using dirn by auto
show "(\<Sum>x\<leftarrow>C. lec_rel (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))) = Leq_Rel"
unfolding sum_list_Leq_Rel
proof
fix c
assume c: "c \<in> set C"
show "lec_rel (lec_of_nsc (f (atom_var c) *R nsc_of_atom c)) = Leq_Rel"
using is_leq[rule_format, OF c] by (cases "f (atom_var c) *R nsc_of_atom c", auto)
qed
qed (simp add: c_def)
show "c < 0"
proof -
define scale_const_f :: "'a atom \<Rightarrow> 'a" where
"scale_const_f x = lec_const (lec_of_nsc (f (atom_var x) *R nsc_of_atom x))" for x
obtain d where bl': "?\<B>\<^sub>l s x\<^sub>i = Some d"
using 3 by (cases "?\<B>\<^sub>l s x\<^sub>i", auto simp: bound_compare_defs)
have "c = (\<Sum>x\<leftarrow>map (\<lambda>x. ?Geq x (the (?\<B>\<^sub>l s x))) ?negA. scale_const_f x)
+ (\<Sum>x\<leftarrow> map (\<lambda>x. ?Leq x (the (?\<B>\<^sub>u s x))) ?posA. scale_const_f x)
- nega d"
unfolding c_def C_def f_def scale_const_f_def using dirn rhs_eval_xi bl' by auto
also have "(\<Sum>x\<leftarrow>map (\<lambda>x. ?Geq x (the (?\<B>\<^sub>l s x))) ?negA. scale_const_f x) =
(\<Sum>x\<leftarrow> ?negA. nega (A x *R the (?\<B>\<^sub>l s x)))"
using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def)
also have "\<dots> = (\<Sum>x\<leftarrow>?negA. nega (A x *R \<langle>\<V> s\<rangle> x))"
using 0 by (subst map_cong) auto
also have "(\<Sum>x\<leftarrow>map (\<lambda>x. ?Leq x (the (?\<B>\<^sub>u s x))) ?posA. scale_const_f x) =
(\<Sum>x\<leftarrow> ?posA. nega (A x *R the (?\<B>\<^sub>u s x)))"
using xi_Xj dirn by (subst map_cong) (auto simp add: f_def scale_const_f_def)
also have "\<dots> = (\<Sum>x\<leftarrow> ?posA. nega (A x *R \<langle>\<V> s\<rangle> x))"
using 0 by (subst map_cong) auto
also have "(\<Sum>x\<leftarrow>?negA. nega (A x *R \<langle>\<V> s\<rangle> x)) + (\<Sum>x\<leftarrow>?posA. nega (A x *R \<langle>\<V> s\<rangle> x))
= (\<Sum>x\<leftarrow>?negA @ ?posA. nega (A x *R \<langle>\<V> s\<rangle> x))"
by auto
also have "\<dots> = (\<Sum>x\<in> X\<^sub>j. nega (A x *R \<langle>\<V> s\<rangle> x))"
using XL\<^sub>j_distinct by (subst sum_list_distinct_conv_sum_set) (auto intro!: sum.cong)
also have "\<dots> = nega (\<Sum>x\<in> X\<^sub>j. (A x *R \<langle>\<V> s\<rangle> x))" using dirn by (auto simp: sum_negf)
also have "(\<Sum>x\<in> X\<^sub>j. (A x *R \<langle>\<V> s\<rangle> x)) = ((rhs ?eq) \<lbrace>\<langle>\<V> s\<rangle>\<rbrace>)"
unfolding A_def X\<^sub>j_def by (subst linear_poly_sum) (auto simp add: sum_negf)
also have "\<dots> = \<langle>\<V> s\<rangle> x\<^sub>i"
using rhs_eval_xi by blast
also have "nega (\<langle>\<V> s\<rangle> x\<^sub>i) - nega d < 0"
proof -
have "?lt (\<langle>\<V> s\<rangle> x\<^sub>i) d"
using dirn 3(2-) bl' by (elim disjE, auto simp: bound_compare_defs)
thus ?thesis using dirn unfolding minus_lt[symmetric] by auto
qed
finally show ?thesis .
qed
show "distinct C"
unfolding C_def using XL\<^sub>j_distinct xi_Xj dirn by (auto simp add: inj_on_def distinct_map)
qed
qed (insert U, blast+)
then obtain f p c C where Qs: "?Q s f p c C" using U unfolding check by blast
from index[folded check_tableau_index_valid[OF U(2) inv(3,4,2,1)]] check
have index: "index_valid as s" by auto
from check_tableau_equiv[OF U(2) inv(3,4,2,1), unfolded check]
have id: "v \<Turnstile>\<^sub>t \<T> s = v \<Turnstile>\<^sub>t \<T> s'" for v :: "'a valuation" by auto
let ?C = "map (\<lambda> a. (f (atom_var a), a)) C"
have "set C \<subseteq> \<B>\<^sub>A s" using Qs by blast
also have "\<dots> \<subseteq> snd ` as" using index
unfolding bound_atoms_def index_valid_def set_of_map_def boundsl_def boundsu_def o_def by force
finally have sub: "snd ` set ?C \<subseteq> snd ` as" by force
show ?thesis unfolding farkas_coefficients_atoms_tableau_def
by (intro exI[of _ p] exI[of _ c] exI[of _ ?C] conjI,
insert Qs[unfolded id] sub, (force simp: o_def)+)
qed
end
text \<open>Next, we show that a conflict found by the assert-bound function also gives rise to
Farkas coefficients.\<close>
context Update
begin
lemma farkas_assert_bound: assumes inv: "\<not> \<U> s" "\<Turnstile>\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s s" "\<triangle> (\<T> s)" "\<nabla> s" "\<diamond> s"
and index: "index_valid as s"
and U: "\<U> (assert_bound ia s)"
shows "\<exists> C. farkas_coefficients_atoms_tableau (snd ` (insert ia as)) (\<T> s) C"
proof -
obtain i a where ia[simp]: "ia = (i,a)" by force
let ?A = "snd ` insert ia as"
have "\<exists> x c d. Leq x c \<in> ?A \<and> Geq x d \<in> ?A \<and> c < d"
proof (cases a)
case (Geq x d)
let ?s = "update\<B>\<I> (Direction.UBI_upd (Direction (\<lambda>x y. y < x) \<B>\<^sub>i\<^sub>u \<B>\<^sub>i\<^sub>l \<B>\<^sub>u \<B>\<^sub>l \<I>\<^sub>u \<I>\<^sub>l \<B>\<^sub>i\<^sub>l_update Geq Leq (\<le>)))
i x d s"
have id: "\<U> ?s = \<U> s" by auto
have norm: "\<triangle> (\<T> ?s)" using inv by auto
have val: "\<nabla> ?s" using inv(4) unfolding tableau_valuated_def by simp
have idd: "x \<notin> lvars (\<T> ?s) \<Longrightarrow> \<U> (update x d ?s) = \<U> ?s"
by (rule update_unsat_id[OF norm val])
from U[unfolded ia Geq] inv(1) id idd
have "\<lhd>\<^sub>l\<^sub>b (\<lambda>x y. y < x) d (\<B>\<^sub>u s x)" by (auto split: if_splits simp: Let_def)
then obtain c where Bu: "\<B>\<^sub>u s x = Some c" and lt: "c < d"
by (cases "\<B>\<^sub>u s x", auto simp: bound_compare_defs)
from Bu obtain j where "Mapping.lookup (\<B>\<^sub>i\<^sub>u s) x = Some (j,c)"
unfolding boundsu_def by auto
with index[unfolded index_valid_def] have "(j, Leq x c) \<in> as" by auto
hence xc: "Leq x c \<in> ?A" by force
have xd: "Geq x d \<in> ?A" unfolding ia Geq by force
from xc xd lt show ?thesis by auto
next
case (Leq x c)
let ?s = "update\<B>\<I> (Direction.UBI_upd (Direction (<) \<B>\<^sub>i\<^sub>l \<B>\<^sub>i\<^sub>u \<B>\<^sub>l \<B>\<^sub>u \<I>\<^sub>l \<I>\<^sub>u \<B>\<^sub>i\<^sub>u_update Leq Geq (\<ge>))) i x c s"
have id: "\<U> ?s = \<U> s" by auto
have norm: "\<triangle> (\<T> ?s)" using inv by auto
have val: "\<nabla> ?s" using inv(4) unfolding tableau_valuated_def by simp
have idd: "x \<notin> lvars (\<T> ?s) \<Longrightarrow> \<U> (update x c ?s) = \<U> ?s"
by (rule update_unsat_id[OF norm val])
from U[unfolded ia Leq] inv(1) id idd
have "\<lhd>\<^sub>l\<^sub>b (<) c (\<B>\<^sub>l s x)" by (auto split: if_splits simp: Let_def)
then obtain d where Bl: "\<B>\<^sub>l s x = Some d" and lt: "c < d"
by (cases "\<B>\<^sub>l s x", auto simp: bound_compare_defs)
from Bl obtain j where "Mapping.lookup (\<B>\<^sub>i\<^sub>l s) x = Some (j,d)"
unfolding boundsl_def by auto
with index[unfolded index_valid_def] have "(j, Geq x d) \<in> as" by auto
hence xd: "Geq x d \<in> ?A" by force
have xc: "Leq x c \<in> ?A" unfolding ia Leq by force
from xc xd lt show ?thesis by auto
qed
then obtain x c d where c: "Leq x c \<in> ?A" and d: "Geq x d \<in> ?A" and cd: "c < d" by blast
show ?thesis unfolding farkas_coefficients_atoms_tableau_def
proof (intro exI conjI allI)
let ?C = "[(-1, Geq x d), (1,Leq x c)]"
show "\<forall>(r,a)\<in>set ?C. a \<in> ?A \<and> is_leq_ns (r *R nsc_of_atom a) \<and> r \<noteq> 0" using c d by auto
show "c - d < 0" using cd using minus_lt by auto
qed (auto simp: valuate_zero)
qed
end
text \<open>Moreover, we prove that all other steps of the simplex algorithm on layer~4, such as pivoting,
asserting bounds without conflict, etc., preserve Farkas coefficients.\<close>
lemma farkas_coefficients_atoms_tableau_mono: assumes "as \<subseteq> bs"
shows "farkas_coefficients_atoms_tableau as t C \<Longrightarrow> farkas_coefficients_atoms_tableau bs t C"
using assms unfolding farkas_coefficients_atoms_tableau_def by force
locale AssertAllState''' = AssertAllState'' init ass_bnd chk + Update update +
PivotUpdateMinVars eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update
for init and ass_bnd :: "'i \<times> 'a :: lrv atom \<Rightarrow> _" and chk :: "('i, 'a) state \<Rightarrow> ('i, 'a) state" and update :: "nat \<Rightarrow> 'a :: lrv \<Rightarrow> ('i, 'a) state \<Rightarrow> ('i, 'a) state"
and eq_idx_for_lvar :: "tableau \<Rightarrow> var \<Rightarrow> nat" and
min_lvar_not_in_bounds :: "('i,'a::lrv) state \<Rightarrow> var option" and
min_rvar_incdec_eq :: "('i,'a) Direction \<Rightarrow> ('i,'a) state \<Rightarrow> eq \<Rightarrow> 'i list + var" and
pivot_and_update :: "var \<Rightarrow> var \<Rightarrow> 'a \<Rightarrow> ('i,'a) state \<Rightarrow> ('i,'a) state"
+ assumes ass_bnd: "ass_bnd = Update.assert_bound update" and
chk: "chk = PivotUpdateMinVars.check eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update"
context AssertAllState'''
begin
lemma farkas_assert_bound_loop: assumes "\<U> (assert_bound_loop as (init t))"
and norm: "\<triangle> t"
shows "\<exists> C. farkas_coefficients_atoms_tableau (snd ` set as) t C"
proof -
let ?P = "\<lambda> as s. \<U> s \<longrightarrow> (\<exists> C. farkas_coefficients_atoms_tableau (snd ` as) (\<T> s) C)"
let ?s = "assert_bound_loop as (init t)"
have "\<not> \<U> (init t)" by (rule init_unsat_flag)
have "\<T> (assert_bound_loop as (init t)) = t \<and>
(\<U> (assert_bound_loop as (init t)) \<longrightarrow> (\<exists> C. farkas_coefficients_atoms_tableau (snd ` set as) (\<T> (init t)) C))"
proof (rule AssertAllState''Induct[OF norm], unfold ass_bnd, goal_cases)
case 1
have "\<not> \<U> (init t)" by (rule init_unsat_flag)
moreover have "\<T> (init t) = t" by (rule init_tableau_id)
ultimately show ?case by auto
next
case (2 as bs s)
hence "snd ` as \<subseteq> snd ` bs" by auto
from farkas_coefficients_atoms_tableau_mono[OF this] 2(2) show ?case by auto
next
case (3 s a ats)
let ?s = "assert_bound a s"
have tab: "\<T> ?s = \<T> s" unfolding ass_bnd by (rule assert_bound_nolhs_tableau_id, insert 3, auto)
have t: "t = \<T> s" using 3 by simp
show ?case unfolding t tab
proof (intro conjI impI refl)
assume "\<U> ?s"
from farkas_assert_bound[OF 3(1,3-6,8) this]
show "\<exists> C. farkas_coefficients_atoms_tableau (snd ` insert a (set ats)) (\<T> (init (\<T> s))) C"
unfolding t[symmetric] init_tableau_id .
qed
qed
thus ?thesis unfolding init_tableau_id using assms by blast
qed
text \<open>Now we get to the main result for layer~4: If the main algorithm returns unsat,
then there are Farkas coefficients for the tableau and atom set that were given as
input for this layer.\<close>
lemma farkas_assert_all_state: assumes U: "\<U> (assert_all_state t as)"
and norm: "\<triangle> t"
shows "\<exists> C. farkas_coefficients_atoms_tableau (snd ` set as) t C"
proof -
let ?s = "assert_bound_loop as (init t)"
show ?thesis
proof (cases "\<U> (assert_bound_loop as (init t))")
case True
from farkas_assert_bound_loop[OF this norm]
show ?thesis by auto
next
case False
from AssertAllState''_tableau_id[OF norm]
have T: "\<T> ?s = t" unfolding init_tableau_id .
from U have U: "\<U> (check ?s)" unfolding chk[symmetric] by simp
show ?thesis
proof (rule farkas_check[OF refl U False, unfolded T, OF _ norm])
from AssertAllState''_precond[OF norm, unfolded Let_def] False
show "\<Turnstile>\<^sub>n\<^sub>o\<^sub>l\<^sub>h\<^sub>s ?s" "\<diamond> ?s" "\<nabla> ?s" by blast+
from AssertAllState''_index_valid[OF norm]
show "index_valid (set as) ?s" .
qed
qed
qed
subsection \<open>Farkas' Lemma on Layer 3\<close>
text \<open>There is only a small difference between layers 3 and 4, namely that there is no
simplex algorithm (@{const assert_all_state}) on layer 3, but just a tableau and atoms.\<close>
text \<open>Hence, one task is to link the unsatisfiability flag
on layer 4 with unsatisfiability of the original tableau and atoms (layer 3). This can
be done via the existing soundness results of the simplex algorithm.
Moreover, we give an easy proof that the existence of Farkas coefficients for a tableau and
set of atoms implies unsatisfiability.\<close>
end
lemma farkas_coefficients_atoms_tableau_unsat:
assumes "farkas_coefficients_atoms_tableau as t C"
shows "\<nexists> v. v \<Turnstile>\<^sub>t t \<and> v \<Turnstile>\<^sub>a\<^sub>s as"
proof
assume "\<exists> v. v \<Turnstile>\<^sub>t t \<and> v \<Turnstile>\<^sub>a\<^sub>s as"
then obtain v where *: "v \<Turnstile>\<^sub>t t \<and> v \<Turnstile>\<^sub>a\<^sub>s as" by auto
then obtain p c where isleq: "(\<forall>(r,a) \<in> set C. a \<in> as \<and> is_leq_ns (r *R nsc_of_atom a) \<and> r \<noteq> 0)"
and leq: "(\<Sum>(r,a) \<leftarrow> C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c"
and cltz: "c < 0"
and p0: "p\<lbrace>v\<rbrace> = 0"
using assms farkas_coefficients_atoms_tableau_def by blast
have fa: "\<forall>(r,a) \<in> set C. v \<Turnstile>\<^sub>a a" using * isleq leq
satisfies_atom_set_def by force
{
fix r a
assume a: "(r,a) \<in> set C"
from a fa have va: "v \<Turnstile>\<^sub>a a" unfolding satisfies_atom_set_def by auto
hence v: "v \<Turnstile>\<^sub>n\<^sub>s (r *R nsc_of_atom a)" by (auto simp: nsc_of_atom sat_scale_rat_ns)
from a isleq have "is_leq_ns (r *R nsc_of_atom a)" by auto
from lec_of_nsc[OF this] v have "v \<Turnstile>\<^sub>l\<^sub>e lec_of_nsc (r *R nsc_of_atom a)" by blast
} note v = this
have "v \<Turnstile>\<^sub>l\<^sub>e Leqc p c" unfolding leq[symmetric]
by (rule satisfies_sumlist_le_constraints, insert v, auto)
then have "0 \<le> c" using p0 by auto
then show False using cltz by auto
qed
text \<open>Next is the main result for layer~3: a tableau and a finite set of atoms are unsatisfiable
if and only if there is a list of Farkas coefficients for the set of atoms and the tableau.\<close>
lemma farkas_coefficients_atoms_tableau: assumes norm: "\<triangle> t"
and fin: "finite as"
shows "(\<exists> C. farkas_coefficients_atoms_tableau as t C) \<longleftrightarrow> (\<nexists> v. v \<Turnstile>\<^sub>t t \<and> v \<Turnstile>\<^sub>a\<^sub>s as)"
proof
from finite_list[OF fin] obtain bs where as: "as = set bs" by auto
assume unsat: "\<nexists> v. v \<Turnstile>\<^sub>t t \<and> v \<Turnstile>\<^sub>a\<^sub>s as"
let ?as = "map (\<lambda> x. ((),x)) bs"
interpret AssertAllState''' init_state assert_bound_code check_code update_code
eq_idx_for_lvar min_lvar_not_in_bounds min_rvar_incdec_eq pivot_and_update_code
by (unfold_locales, auto simp: assert_bound_code_def check_code_def)
let ?call = "assert_all t ?as"
have id: "snd ` set ?as = as" unfolding as by force
from assert_all_sat[OF norm, of ?as, unfolded id] unsat
obtain I where "?call = Inl I" by (cases ?call, auto)
from this[unfolded assert_all_def Let_def]
have "\<U> (assert_all_state_code t ?as)"
by (auto split: if_splits simp: assert_all_state_code_def)
from farkas_assert_all_state[OF this[unfolded assert_all_state_code_def] norm, unfolded id]
show "\<exists> C. farkas_coefficients_atoms_tableau as t C" .
qed (insert farkas_coefficients_atoms_tableau_unsat, auto)
subsection \<open>Farkas' Lemma on Layer 2\<close>
text \<open>The main difference between layers 2 and 3 is the introduction of slack-variables in layer 3
via the preprocess-function. Our task here is to show that Farkas coefficients at layer 3 (where
slack-variables are used) can be converted into Farkas coefficients for layer 2 (before the
preprocessing).\<close>
text \<open>We also need to adapt the previos notion of Farkas coefficients, which was used in
@{const farkas_coefficients_atoms_tableau}, for layer~2. At layer 3, Farkas coefficients
are the coefficients in a linear combination of atoms that evaluates to an inequality of
the form $p \leq c$, where $p$ is a linear polynomial, $c < 0$, and $t \models p = 0$ holds.
At layer 2, the atoms are replaced by non-strict constraints where the left-hand side is a
polynomial in the original variables, but the corresponding linear combination (with Farkas
coefficients) evaluates directly to the inequality $0 \leq c$, with $c < 0$. The implication
$t \models p = 0$ is no longer possible in this layer, since there is no tableau $t$, nor is it
needed, since $p$ is $0$. Thus, the statement defining Farkas coefficients must be changed
accordingly.\<close>
definition farkas_coefficients_ns where
"farkas_coefficients_ns ns C = (\<exists> c.
(\<forall>(r, n) \<in> set C. n \<in> ns \<and> is_leq_ns (r *R n) \<and> r \<noteq> 0) \<and>
(\<Sum>(r, n) \<leftarrow> C. lec_of_nsc (r *R n)) = Leqc 0 c \<and>
c < 0)"
text \<open>The easy part is to prove that Farkas coefficients imply unsatisfiability.\<close>
lemma farkas_coefficients_ns_unsat:
assumes "farkas_coefficients_ns ns C"
shows "\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns"
proof
assume "\<exists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns"
then obtain v where *: "v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns" by auto
obtain c where
isleq: "(\<forall>(a,n) \<in> set C. n \<in> ns \<and> is_leq_ns (a *R n) \<and> a \<noteq> 0)" and
leq: "(\<Sum>(a,n) \<leftarrow> C. lec_of_nsc (a *R n)) = Leqc 0 c" and
cltz: "c < 0" using assms farkas_coefficients_ns_def by blast
{
fix a n
assume n: "(a,n) \<in> set C"
from n * isleq have "v \<Turnstile>\<^sub>n\<^sub>s n" by auto
hence v: "v \<Turnstile>\<^sub>n\<^sub>s (a *R n)" by (rule sat_scale_rat_ns)
from n isleq have "is_leq_ns (a *R n)" by auto
from lec_of_nsc[OF this] v
have "v \<Turnstile>\<^sub>l\<^sub>e lec_of_nsc (a *R n)" by blast
} note v = this
have "v \<Turnstile>\<^sub>l\<^sub>e Leqc 0 c" unfolding leq[symmetric]
by (rule satisfies_sumlist_le_constraints, insert v, auto)
then show False using cltz
by (metis leD satisfiable_le_constraint.simps valuate_zero rel_of.simps(1))
qed
text \<open>In order to eliminate the need for a tableau, we require the notion of an arbitrary substitution
on polynomials, where all variables can be replaced at once. The existing simplex formalization
provides only a function to replace one variable at a time.\<close>
definition subst_poly :: "(var \<Rightarrow> linear_poly) \<Rightarrow> linear_poly \<Rightarrow> linear_poly" where
"subst_poly \<sigma> p = (\<Sum> x \<in> vars p. coeff p x *R \<sigma> x)"
lemma subst_poly_0[simp]: "subst_poly \<sigma> 0 = 0" unfolding subst_poly_def by simp
lemma valuate_subst_poly: "(subst_poly \<sigma> p) \<lbrace> v \<rbrace> = (p \<lbrace> (\<lambda> x. ((\<sigma> x) \<lbrace> v \<rbrace>)) \<rbrace>)"
by (subst (2) linear_poly_sum, unfold subst_poly_def valuate_sum valuate_scaleRat, simp)
lemma subst_poly_add: "subst_poly \<sigma> (p + q) = subst_poly \<sigma> p + subst_poly \<sigma> q"
by (rule linear_poly_eqI, unfold valuate_add valuate_subst_poly, simp)
fun subst_poly_lec :: "(var \<Rightarrow> linear_poly) \<Rightarrow> 'a le_constraint \<Rightarrow> 'a le_constraint" where
"subst_poly_lec \<sigma> (Le_Constraint rel p c) = Le_Constraint rel (subst_poly \<sigma> p) c"
lemma subst_poly_lec_0[simp]: "subst_poly_lec \<sigma> 0 = 0" unfolding zero_le_constraint_def by simp
lemma subst_poly_lec_add: "subst_poly_lec \<sigma> (c1 + c2) = subst_poly_lec \<sigma> c1 + subst_poly_lec \<sigma> c2"
by (cases c1; cases c2, auto simp: subst_poly_add)
lemma subst_poly_lec_sum_list: "subst_poly_lec \<sigma> (sum_list ps) = sum_list (map (subst_poly_lec \<sigma>) ps)"
by (induct ps, auto simp: subst_poly_lec_add)
lemma subst_poly_lp_monom[simp]: "subst_poly \<sigma> (lp_monom r x) = r *R \<sigma> x"
unfolding subst_poly_def by (simp add: vars_lp_monom)
lemma subst_poly_scaleRat: "subst_poly \<sigma> (r *R p) = r *R (subst_poly \<sigma> p)"
by (rule linear_poly_eqI, unfold valuate_scaleRat valuate_subst_poly, simp)
text \<open>We need several auxiliary properties of the preprocess-function which are not present
in the simplex formalization.\<close>
lemma Tableau_is_monom_preprocess':
assumes "(x, p) \<in> set (Tableau (preprocess' cs start))"
shows "\<not> is_monom p"
using assms
by(induction cs start rule: preprocess'.induct)
(auto simp add: Let_def split: if_splits option.splits)
lemma preprocess'_atoms_to_constraints': assumes "preprocess' cs start = S"
shows "set (Atoms S) \<subseteq> {(i,qdelta_constraint_to_atom c v) | i c v. (i,c) \<in> set cs \<and>
(\<not> is_monom (poly c) \<longrightarrow> Poly_Mapping S (poly c) = Some v)}"
unfolding assms(1)[symmetric]
by (induct cs start rule: preprocess'.induct, auto simp: Let_def split: option.splits, force+)
lemma monom_of_atom_coeff:
assumes "is_monom (poly ns)" "a = qdelta_constraint_to_atom ns v"
shows "(monom_coeff (poly ns)) *R nsc_of_atom a = ns"
using assms is_monom_monom_coeff_not_zero
by(cases a; cases ns)
(auto split: atom.split ns_constraint.split simp add: monom_poly_assemble field_simps)
text \<open>The next lemma provides the functionality that is required to convert an
atom back to a non-strict constraint, i.e., it is a kind of inverse of the preprocess-function.\<close>
lemma preprocess'_atoms_to_constraints: assumes S: "preprocess' cs start = S"
and start: "start = start_fresh_variable cs"
and ns: "ns = (case a of Leq v c \<Rightarrow> LEQ_ns q c | Geq v c \<Rightarrow> GEQ_ns q c)"
and "a \<in> snd ` set (Atoms S)"
shows "(atom_var a \<notin> fst ` set (Tableau S) \<longrightarrow> (\<exists> r. r \<noteq> 0 \<and> r *R nsc_of_atom a \<in> snd ` set cs))
\<and> ((atom_var a, q) \<in> set (Tableau S) \<longrightarrow> ns \<in> snd ` set cs)"
proof -
let ?S = "preprocess' cs start"
from assms(4) obtain i where ia: "(i,a) \<in> set (Atoms S)" by auto
with preprocess'_atoms_to_constraints'[OF assms(1)] obtain c v
where a: "a = qdelta_constraint_to_atom c v" and c: "(i,c) \<in> set cs"
and nmonom: "\<not> is_monom (poly c) \<Longrightarrow> Poly_Mapping S (poly c) = Some v" by blast
hence c': "c \<in> snd ` set cs" by force
let ?p = "poly c"
show ?thesis
proof (cases "is_monom ?p")
case True
hence av: "atom_var a = monom_var ?p" unfolding a by (cases c, auto)
from Tableau_is_monom_preprocess'[of _ ?p cs start] True
have "(x, ?p) \<notin> set (Tableau ?S)" for x by blast
{
assume "(atom_var a, q) \<in> set (Tableau S)"
hence "(monom_var ?p, q) \<in> set (Tableau S)" unfolding av by simp
hence "monom_var ?p \<in> lvars (Tableau S)" unfolding lvars_def by force
from lvars_tableau_ge_start[rule_format, OF this[folded S]]
have "monom_var ?p \<ge> start" unfolding S .
moreover have "monom_var ?p \<in> vars_constraints (map snd cs)" using True c
by (auto intro!: bexI[of _ "(i,c)"] simp: monom_var_in_vars)
ultimately have False using start_fresh_variable_fresh[of cs, folded start] by force
}
moreover
from monom_of_atom_coeff[OF True a] True
have "\<exists>r. r \<noteq> 0 \<and> r *R nsc_of_atom a = c"
by (intro exI[of _ "monom_coeff ?p"], auto, cases a, auto)
ultimately show ?thesis using c' by auto
next
case False
hence av: "atom_var a = v" unfolding a by (cases c, auto)
from nmonom[OF False] have "Poly_Mapping S ?p = Some v" .
from preprocess'_Tableau_Poly_Mapping_Some[OF this[folded S]]
have tab: "(atom_var a, ?p) \<in> set (Tableau (preprocess' cs start))" unfolding av by simp
hence "atom_var a \<in> fst ` set (Tableau S)" unfolding S by force
moreover
{
assume "(atom_var a, q) \<in> set (Tableau S)"
from tab this have qp: "q = ?p" unfolding S using lvars_distinct[of cs start, unfolded S lhs_def]
by (simp add: case_prod_beta' eq_key_imp_eq_value)
have "ns = c" unfolding ns qp using av a False by (cases c, auto)
hence "ns \<in> snd ` set cs" using c' by blast
}
ultimately show ?thesis by blast
qed
qed
text \<open>Next follows the major technical lemma of this part,
namely that Farkas coefficients on layer~3 for preprocessed constraints can
be converted into Farkas coefficients on layer~2.\<close>
lemma farkas_coefficients_preprocess':
assumes pp: "preprocess' cs (start_fresh_variable cs) = S" and
ft: "farkas_coefficients_atoms_tableau (snd ` set (Atoms S)) (Tableau S) C"
shows "\<exists> C. farkas_coefficients_ns (snd ` set cs) C"
proof -
note ft[unfolded farkas_coefficients_atoms_tableau_def]
obtain p c where 0: "\<forall> (r,a) \<in> set C. a \<in> snd ` set (Atoms S) \<and> is_leq_ns (r *R nsc_of_atom a) \<and> r \<noteq> 0"
"(\<Sum>(r,a)\<leftarrow>C. lec_of_nsc (r *R nsc_of_atom a)) = Leqc p c"
"c < 0"
"\<And>v :: QDelta valuation. v \<Turnstile>\<^sub>t Tableau S \<Longrightarrow> p \<lbrace> v \<rbrace> = 0"
using ft unfolding farkas_coefficients_atoms_tableau_def by blast
note 0 = 0(1)[rule_format, of "(a, b)" for a b, unfolded split] 0(2-)
let ?T = "Tableau S"
define \<sigma> :: "var \<Rightarrow> linear_poly" where "\<sigma> = (\<lambda> x. case map_of ?T x of Some p \<Rightarrow> p | None \<Rightarrow> lp_monom 1 x)"
let ?P = "(\<lambda>r a s ns. ns \<in> (snd ` set cs) \<and> is_leq_ns (s *R ns) \<and> s \<noteq> 0 \<and>
subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a)) = lec_of_nsc (s *R ns))"
have "\<exists>s ns. ?P r a s ns" if ra: "(r,a) \<in> set C" for r a
proof -
have a: "a \<in> snd ` set (Atoms S)"
using ra 0 by force
from 0 ra have is_leq: "is_leq_ns (r *R nsc_of_atom a)" and r0: "r \<noteq> 0" by auto
let ?x = "atom_var a"
show ?thesis
proof (cases "map_of ?T ?x")
case (Some q)
hence \<sigma>: "\<sigma> ?x = q" unfolding \<sigma>_def by auto
from Some have xqT: "(?x, q) \<in> set ?T" by (rule map_of_SomeD)
define ns where "ns = (case a of Leq v c \<Rightarrow> LEQ_ns q c
| Geq v c \<Rightarrow> GEQ_ns q c)"
from preprocess'_atoms_to_constraints[OF pp refl ns_def a] xqT
have ns_mem: "ns \<in> snd ` set cs" by blast
have id: "subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a))
= lec_of_nsc (r *R ns)" using is_leq \<sigma>
by (cases a, auto simp: ns_def subst_poly_scaleRat)
from id is_leq \<sigma> have is_leq: "is_leq_ns (r *R ns)" by (cases a, auto simp: ns_def)
show ?thesis by (intro exI[of _ r] exI[of _ ns] conjI ns_mem id is_leq conjI r0)
next
case None
hence "?x \<notin> fst ` set ?T" by (meson map_of_eq_None_iff)
from preprocess'_atoms_to_constraints[OF pp refl refl a] this
obtain rr where rr: "rr *R nsc_of_atom a \<in> (snd ` set cs)" and rr0: "rr \<noteq> 0"
by blast
from None have \<sigma>: "\<sigma> ?x = lp_monom 1 ?x" unfolding \<sigma>_def by simp
define ns where "ns = rr *R nsc_of_atom a"
define s where "s = r / rr"
from rr0 r0 have s0: "s \<noteq> 0" unfolding s_def by auto
from is_leq \<sigma>
have "subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a))
= lec_of_nsc (r *R nsc_of_atom a)"
by (cases a, auto simp: subst_poly_scaleRat)
moreover have "r *R nsc_of_atom a = s *R ns" unfolding ns_def s_def
scaleRat_scaleRat_ns_constraint[OF rr0] using rr0 by simp
ultimately have "subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a))
= lec_of_nsc (s *R ns)" "is_leq_ns (s *R ns)" using is_leq by auto
then show ?thesis
unfolding ns_def using rr s0 by blast
qed
qed
hence "\<forall> ra. \<exists> s ns. (fst ra, snd ra) \<in> set C \<longrightarrow> ?P (fst ra) (snd ra) s ns" by blast
from choice[OF this] obtain s where s: "\<forall> ra. \<exists> ns. (fst ra, snd ra) \<in> set C \<longrightarrow> ?P (fst ra) (snd ra) (s ra) ns" by blast
from choice[OF this] obtain ns where ns: "\<And> r a. (r,a) \<in> set C \<Longrightarrow> ?P r a (s (r,a)) (ns (r,a))" by force
define NC where "NC = map (\<lambda>(r,a). (s (r,a), ns (r,a))) C"
have "(\<Sum>(s, ns)\<leftarrow>map (\<lambda>(r,a). (s (r,a), ns (r,a))) C'. lec_of_nsc (s *R ns)) =
(\<Sum>(r, a)\<leftarrow>C'. subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a)))"
if "set C' \<subseteq> set C" for C'
using that proof (induction C')
case Nil
then show ?case by simp
next
case (Cons a C')
have "(\<Sum>x\<leftarrow>a # C'. lec_of_nsc (s x *R ns x)) =
lec_of_nsc (s a *R ns a) + (\<Sum>x\<leftarrow>C'. lec_of_nsc (s x *R ns x))"
by simp
also have "(\<Sum>x\<leftarrow>C'. lec_of_nsc (s x *R ns x)) = (\<Sum>(r, a)\<leftarrow>C'. subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a)))"
using Cons by (auto simp add: case_prod_beta' comp_def)
also have "lec_of_nsc (s a *R ns a) = subst_poly_lec \<sigma> (lec_of_nsc (fst a *R nsc_of_atom (snd a)))"
proof -
have "a \<in> set C"
using Cons by simp
then show ?thesis
using ns by auto
qed
finally show ?case
by (auto simp add: case_prod_beta' comp_def)
qed
also have "(\<Sum>(r, a)\<leftarrow>C. subst_poly_lec \<sigma> (lec_of_nsc (r *R nsc_of_atom a)))
= subst_poly_lec \<sigma> (\<Sum>(r, a)\<leftarrow>C. (lec_of_nsc (r *R nsc_of_atom a)))"
by (auto simp add: subst_poly_lec_sum_list case_prod_beta' comp_def)
also have "(\<Sum>(r, a)\<leftarrow>C. (lec_of_nsc (r *R nsc_of_atom a))) = Leqc p c"
using 0 by blast
also have "subst_poly_lec \<sigma> (Leqc p c) = Leqc (subst_poly \<sigma> p) c" by simp
also have "subst_poly \<sigma> p = 0"
proof (rule all_valuate_zero)
fix v :: "QDelta valuation"
have "(subst_poly \<sigma> p) \<lbrace> v \<rbrace> = (p \<lbrace> \<lambda>x. ((\<sigma> x) \<lbrace> v \<rbrace>) \<rbrace>)" by (rule valuate_subst_poly)
also have "\<dots> = 0"
proof (rule 0(4))
have "(\<sigma> a) \<lbrace> v \<rbrace> = (q \<lbrace> \<lambda>x. ((\<sigma> x) \<lbrace> v \<rbrace>) \<rbrace>)" if "(a, q) \<in> set (Tableau S)" for a q
proof -
have "distinct (map fst ?T)"
using normalized_tableau_preprocess' assms unfolding normalized_tableau_def lhs_def
by (auto simp add: case_prod_beta')
then have 0: "\<sigma> a = q"
unfolding \<sigma>_def using that by auto
have "q \<lbrace> v \<rbrace> = (q \<lbrace> \<lambda>x. ((\<sigma> x) \<lbrace> v \<rbrace>) \<rbrace>)"
proof -
have "vars q \<subseteq> rvars ?T"
unfolding rvars_def using that by force
moreover have "(\<sigma> x) \<lbrace> v \<rbrace> = v x" if "x \<in> rvars ?T" for x
proof -
have "x \<notin> lvars (Tableau S)"
using that normalized_tableau_preprocess' assms
unfolding normalized_tableau_def by blast
then have "x \<notin> fst ` set (Tableau S)"
unfolding lvars_def by force
then have "map_of ?T x = None"
using map_of_eq_None_iff by metis
then have "\<sigma> x = lp_monom 1 x"
unfolding \<sigma>_def by auto
also have "(lp_monom 1 x) \<lbrace> v \<rbrace> = v x"
by auto
finally show ?thesis .
qed
ultimately show ?thesis
by (auto intro!: valuate_depend)
qed
then show ?thesis
using 0 by blast
qed
then show "(\<lambda>x. ((\<sigma> x) \<lbrace> v \<rbrace>)) \<Turnstile>\<^sub>t ?T"
using 0 by (auto simp add: satisfies_tableau_def satisfies_eq_def)
qed
finally show "(subst_poly \<sigma> p) \<lbrace> v \<rbrace> = 0" .
qed
finally have "(\<Sum>(s, n)\<leftarrow>NC. lec_of_nsc (s *R n)) = Le_Constraint Leq_Rel 0 c"
unfolding NC_def by blast
moreover have "ns (r,a) \<in> snd ` set cs" "is_leq_ns (s (r, a) *R ns (r, a))" "s (r, a) \<noteq> 0" if "(r, a) \<in> set C" for r a
using ns that by force+
ultimately have "farkas_coefficients_ns (snd ` set cs) NC"
unfolding farkas_coefficients_ns_def NC_def using 0 by force
then show ?thesis
by blast
qed
lemma preprocess'_unsat_indexD: "i \<in> set (UnsatIndices (preprocess' ns j)) \<Longrightarrow>
\<exists> c. poly c = 0 \<and> \<not> zero_satisfies c \<and> (i,c) \<in> set ns"
by (induct ns j rule: preprocess'.induct, auto simp: Let_def split: if_splits option.splits)
lemma preprocess'_unsat_index_farkas_coefficients_ns:
assumes "i \<in> set (UnsatIndices (preprocess' ns j))"
shows "\<exists> C. farkas_coefficients_ns (snd ` set ns) C"
proof -
from preprocess'_unsat_indexD[OF assms]
obtain c where contr: "poly c = 0" "\<not> zero_satisfies c" and mem: "(i,c) \<in> set ns" by auto
from mem have mem: "c \<in> snd ` set ns" by force
let ?c = "ns_constraint_const c"
define r where "r = (case c of LEQ_ns _ _ \<Rightarrow> 1 | _ \<Rightarrow> (-1 :: rat))"
define d where "d = (case c of LEQ_ns _ _ \<Rightarrow> ?c | _ \<Rightarrow> - ?c)"
have [simp]: "(- x < 0) = (0 < x)" for x :: QDelta using uminus_less_lrv[of _ 0] by simp
show ?thesis unfolding farkas_coefficients_ns_def
by (intro exI[of _ "[(r,c)]"] exI[of _ d], insert mem contr, cases "c",
auto simp: r_def d_def)
qed
text \<open>The combination of the previous results easily provides the main result of this section:
a finite set of non-strict constraints on layer~2 is unsatisfiable if and only if there are Farkas coefficients.
Again, here we use results from the simplex formalization, namely soundness of the preprocess-function.\<close>
lemma farkas_coefficients_ns: assumes "finite (ns :: QDelta ns_constraint set)"
shows "(\<exists> C. farkas_coefficients_ns ns C) \<longleftrightarrow> (\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns)"
proof
assume "\<exists> C. farkas_coefficients_ns ns C"
from farkas_coefficients_ns_unsat this show "\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns" by blast
next
assume unsat: "\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ns"
from finite_list[OF assms] obtain nsl where ns: "ns = set nsl" by auto
let ?cs = "map (Pair ()) nsl"
obtain I t ias where part1: "preprocess_part_1 ?cs = (t,ias,I)" by (cases "preprocess_part_1 ?cs", auto)
let ?as = "snd ` set ias"
let ?s = "start_fresh_variable ?cs"
have fin: "finite ?as" by auto
have id: "ias = Atoms (preprocess' ?cs ?s)" "t = Tableau (preprocess' ?cs ?s)"
"I = UnsatIndices (preprocess' ?cs ?s)"
using part1 unfolding preprocess_part_1_def Let_def by auto
have norm: "\<triangle> t" using normalized_tableau_preprocess'[of ?cs] unfolding id .
{
fix v
assume "v \<Turnstile>\<^sub>a\<^sub>s ?as" "v \<Turnstile>\<^sub>t t"
from preprocess'_sat[OF this[unfolded id], folded id] unsat[unfolded ns]
have "set I \<noteq> {}" by auto
then obtain i where "i \<in> set I" using all_not_in_conv by blast
from preprocess'_unsat_index_farkas_coefficients_ns[OF this[unfolded id]]
have "\<exists>C. farkas_coefficients_ns (snd ` set ?cs) C" by simp
}
with farkas_coefficients_atoms_tableau[OF norm fin]
obtain C where "farkas_coefficients_atoms_tableau ?as t C
\<or> (\<exists>C. farkas_coefficients_ns (snd ` set ?cs) C)" by blast
from farkas_coefficients_preprocess'[of ?cs, OF refl] this
have "\<exists> C. farkas_coefficients_ns (snd ` set ?cs) C"
using part1 unfolding preprocess_part_1_def Let_def by auto
also have "snd ` set ?cs = ns" unfolding ns by force
finally show "\<exists> C. farkas_coefficients_ns ns C" .
qed
subsection \<open>Farkas' Lemma on Layer 1\<close>
text \<open>The main difference of layers 1 and 2 is the restriction to non-strict constraints via delta-rationals.
Since we now work with another constraint type, @{type constraint}, we again need translations into
linear inequalities of type @{type le_constraint}. Moreover, we also need to define scaling of constraints
where flipping the comparison sign may be required.\<close>
fun is_le :: "constraint \<Rightarrow> bool" where
"is_le (LT _ _) = True"
| "is_le (LEQ _ _) = True"
| "is_le _ = False"
fun lec_of_constraint where
"lec_of_constraint (LEQ p c) = (Le_Constraint Leq_Rel p c)"
| "lec_of_constraint (LT p c) = (Le_Constraint Lt_Rel p c)"
lemma lec_of_constraint:
assumes "is_le c"
shows "(v \<Turnstile>\<^sub>l\<^sub>e (lec_of_constraint c)) \<longleftrightarrow> (v \<Turnstile>\<^sub>c c)"
using assms by (cases c, auto)
instantiation constraint :: scaleRat
begin
fun scaleRat_constraint :: "rat \<Rightarrow> constraint \<Rightarrow> constraint" where
"scaleRat_constraint r cc = (if r = 0 then LEQ 0 0 else
(case cc of
LEQ p c \<Rightarrow>
(if (r < 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c))
| LT p c \<Rightarrow>
(if (r < 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c))
| GEQ p c \<Rightarrow>
(if (r > 0) then GEQ (r *R p) (r *R c) else LEQ (r *R p) (r *R c))
| GT p c \<Rightarrow>
(if (r > 0) then GT (r *R p) (r *R c) else LT (r *R p) (r *R c))
| LTPP p q \<Rightarrow>
(if (r < 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0)
| LEQPP p q \<Rightarrow>
(if (r < 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0)
| GTPP p q \<Rightarrow>
(if (r > 0) then GT (r *R (p - q)) 0 else LT (r *R (p - q)) 0)
| GEQPP p q \<Rightarrow>
(if (r > 0) then GEQ (r *R (p - q)) 0 else LEQ (r *R (p - q)) 0)
| EQPP p q \<Rightarrow> LEQ (r *R (p - q)) 0 \<comment> \<open>We do not keep equality, since the aim is
to convert the scaled constraints into inequalities, which will then be summed up.\<close>
| EQ p c \<Rightarrow> LEQ (r *R p) (r *R c)
))"
instance ..
end
lemma sat_scale_rat: assumes "(v :: rat valuation) \<Turnstile>\<^sub>c c"
shows "v \<Turnstile>\<^sub>c (r *R c)"
proof -
have "r < 0 \<or> r = 0 \<or> r > 0" by auto
then show ?thesis using assms by (cases c, auto simp: right_diff_distrib
valuate_minus valuate_scaleRat scaleRat_leq1 scaleRat_leq2 valuate_zero)
qed
text \<open>In the following definition of Farkas coefficients (for layer 1), the main difference to
@{const farkas_coefficients_ns} is that the linear combination evaluates either to
a strict inequality where the constant must be non-positive, or to a non-strict inequality where
the constant must be negative.\<close>
definition farkas_coefficients where
"farkas_coefficients cs C = (\<exists> d rel.
(\<forall> (r,c) \<in> set C. c \<in> cs \<and> is_le (r *R c) \<and> r \<noteq> 0) \<and>
(\<Sum> (r,c) \<leftarrow> C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d \<and>
(rel = Leq_Rel \<and> d < 0 \<or> rel = Lt_Rel \<and> d \<le> 0))"
text \<open>Again, the existence Farkas coefficients immediately implies unsatisfiability.\<close>
lemma farkas_coefficients_unsat:
assumes "farkas_coefficients cs C"
shows "\<nexists> v. v \<Turnstile>\<^sub>c\<^sub>s cs"
proof
assume "\<exists> v. v \<Turnstile>\<^sub>c\<^sub>s cs"
then obtain v where *: "v \<Turnstile>\<^sub>c\<^sub>s cs" by auto
obtain d rel where
isleq: "(\<forall>(r,c) \<in> set C. c \<in> cs \<and> is_le (r *R c) \<and> r \<noteq> 0)" and
leq: "(\<Sum> (r,c) \<leftarrow> C. lec_of_constraint (r *R c)) = Le_Constraint rel 0 d" and
choice: "rel = Lt_Rel \<and> d \<le> 0 \<or> rel = Leq_Rel \<and> d < 0" using assms farkas_coefficients_def by blast
{
fix r c
assume c: "(r,c) \<in> set C"
from c * isleq have "v \<Turnstile>\<^sub>c c" by auto
hence v: "v \<Turnstile>\<^sub>c (r *R c)" by (rule sat_scale_rat)
from c isleq have "is_le (r *R c)" by auto
from lec_of_constraint[OF this] v
have "v \<Turnstile>\<^sub>l\<^sub>e lec_of_constraint (r *R c)" by blast
} note v = this
have "v \<Turnstile>\<^sub>l\<^sub>e Le_Constraint rel 0 d" unfolding leq[symmetric]
by (rule satisfies_sumlist_le_constraints, insert v, auto)
then show False using choice
by (cases rel, auto simp: valuate_zero)
qed
text \<open>Now follows the difficult implication.
The major part is proving that the translation @{const constraint_to_qdelta_constraint}
preserves the existence of Farkas coefficients via pointwise compatibility of the sum.
Here, compatibility links a strict or non-strict inequality from the input constraint to
a translated non-strict inequality over delta-rationals.\<close>
fun compatible_cs where
"compatible_cs (Le_Constraint Leq_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \<and> d = QDelta c 0)"
| "compatible_cs (Le_Constraint Lt_Rel p c) (Le_Constraint Leq_Rel q d) = (q = p \<and> qdfst d = c)"
| "compatible_cs _ _ = False"
lemma compatible_cs_0_0: "compatible_cs 0 0" by code_simp
lemma compatible_cs_plus: "compatible_cs c1 d1 \<Longrightarrow> compatible_cs c2 d2 \<Longrightarrow> compatible_cs (c1 + c2) (d1 + d2)"
by (cases c1; cases d1; cases c2; cases d2; cases "lec_rel c1"; cases "lec_rel d1"; cases "lec_rel c2";
cases "lec_rel d2"; auto simp: plus_QDelta_def)
lemma unsat_farkas_coefficients: assumes "\<nexists> v. v \<Turnstile>\<^sub>c\<^sub>s cs"
and fin: "finite cs"
shows "\<exists> C. farkas_coefficients cs C"
proof -
from finite_list[OF fin] obtain csl where cs: "cs = set csl" by blast
let ?csl = "map (Pair ()) csl"
let ?ns = "(snd ` set (to_ns ?csl))"
let ?nsl = "concat (map constraint_to_qdelta_constraint csl)"
have id: "snd ` set ?csl = cs" unfolding cs by force
have id2: "?ns = set ?nsl" unfolding to_ns_def set_concat by force
from SolveExec'Default.to_ns_sat[of ?csl, unfolded id] assms
have unsat: "\<nexists> v. \<langle>v\<rangle> \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns" by metis
have fin: "finite ?ns" by auto
have "\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns"
proof
assume "\<exists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns"
then obtain v where model: "v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns" by blast
let ?v = "Mapping.Mapping (\<lambda> x. Some (v x))"
have "v = \<langle>?v\<rangle>" by (intro ext, auto simp: map2fun_def Mapping.lookup.abs_eq)
from model this unsat show False by metis
qed
from farkas_coefficients_ns[OF fin] this id2 obtain N where
farkas: "farkas_coefficients_ns (set ?nsl) N" by metis
from this[unfolded farkas_coefficients_ns_def]
obtain d where
is_leq: "\<And> a n. (a,n) \<in> set N \<Longrightarrow> n \<in> set ?nsl \<and> is_leq_ns (a *R n) \<and> a \<noteq> 0" and
sum: "(\<Sum>(a,n)\<leftarrow>N. lec_of_nsc (a *R n)) = Le_Constraint Leq_Rel 0 d" and
d0: "d < 0" by blast
let ?prop = "\<lambda> NN C. (\<forall> (a,c) \<in> set C. c \<in> cs \<and> is_le (a *R c) \<and> a \<noteq> 0)
\<and> compatible_cs (\<Sum> (a,c) \<leftarrow> C. lec_of_constraint (a *R c))
(\<Sum>(a,n)\<leftarrow>NN. lec_of_nsc (a *R n))"
have "set NN \<subseteq> set N \<Longrightarrow> \<exists> C. ?prop NN C" for NN
proof (induct NN)
case Nil
have "?prop Nil Nil" by (simp add: compatible_cs_0_0)
thus ?case by blast
next
case (Cons an NN)
obtain a n where an: "an = (a,n)" by force
from Cons an obtain C where IH: "?prop NN C" and n: "(a,n) \<in> set N" by auto
have compat_CN: "compatible_cs (\<Sum>(f, c)\<leftarrow>C. lec_of_constraint (f *R c))
(\<Sum>(a,n)\<leftarrow>NN. lec_of_nsc (a *R n))"
using IH by blast
from n is_leq obtain c where c: "c \<in> cs" and nc: "n \<in> set (constraint_to_qdelta_constraint c)"
unfolding cs by force
from is_leq[OF n] have is_leq: "is_leq_ns (a *R n) \<and> a \<noteq> 0" by blast
have is_less: "is_le (a *R c)" and
a0: "a \<noteq> 0" and
compat_cn: "compatible_cs (lec_of_constraint (a *R c)) (lec_of_nsc (a *R n))"
by (atomize(full), cases c, insert is_leq nc, auto simp: QDelta_0_0 scaleRat_QDelta_def qdsnd_0 qdfst_0)
let ?C = "Cons (a, c) C"
let ?N = "Cons (a, n) NN"
have "?prop ?N ?C" unfolding an
proof (intro conjI)
show "\<forall> (a,c) \<in> set ?C. c \<in> cs \<and> is_le (a *R c) \<and> a \<noteq> 0" using IH is_less a0 c by auto
show "compatible_cs (\<Sum>(a, c)\<leftarrow>?C. lec_of_constraint (a *R c)) (\<Sum>(a,n)\<leftarrow>?N. lec_of_nsc (a *R n))"
using compatible_cs_plus[OF compat_cn compat_CN] by simp
qed
thus ?case unfolding an by blast
qed
from this[OF subset_refl, unfolded sum]
obtain C where
is_less: "(\<forall>(a, c)\<in>set C. c \<in> cs \<and> is_le (a *R c) \<and> a \<noteq> 0)" and
compat: "compatible_cs (\<Sum>(f, c)\<leftarrow>C. lec_of_constraint (f *R c)) (Le_Constraint Leq_Rel 0 d)"
(is "compatible_cs ?sum _")
by blast
obtain rel p e where "?sum = Le_Constraint rel p e" by (cases ?sum)
with compat have sum: "?sum = Le_Constraint rel 0 e" by (cases rel, auto)
have e: "(rel = Leq_Rel \<and> e < 0 \<or> rel = Lt_Rel \<and> e \<le> 0)" using compat[unfolded sum] d0
by (cases rel, auto simp: less_QDelta_def qdfst_0 qdsnd_0)
show ?thesis unfolding farkas_coefficients_def
by (intro exI conjI, rule is_less, rule sum, insert e, auto)
qed
text \<open>Finally we can prove on layer 1 that a finite set of constraints is
unsatisfiable if and only if there are Farkas coefficients.\<close>
lemma farkas_coefficients: assumes "finite cs"
shows "(\<exists> C. farkas_coefficients cs C) \<longleftrightarrow> (\<nexists> v. v \<Turnstile>\<^sub>c\<^sub>s cs)"
using farkas_coefficients_unsat unsat_farkas_coefficients[OF _ assms] by blast
section \<open>Corollaries from the Literature\<close>
text \<open>In this section, we convert the previous variations of Farkas' Lemma into more
well-known forms of this result.
Moreover, instead of referring to the various constraint types of the simplex formalization, we
now speak solely about constraints of type @{type le_constraint}.\<close>
subsection \<open>Farkas' Lemma on Delta-Rationals\<close>
text \<open>We start with Lemma~2 of \cite{Bromberger2017}, a
variant of Farkas' Lemma for delta-rationals. To be more precise, it states
that a set of non-strict inequalities over delta-rationals is unsatisfiable
if and only if there is a linear combination of the inequalities that results
in a trivial unsatisfiable constraint $0 < const$ for some negative constant $const$.
We can easily prove this statement via the lemma @{thm [source] farkas_coefficients_ns}
and some conversions between the
different constraint types.\<close>
lemma Farkas'_Lemma_Delta_Rationals: fixes cs :: "QDelta le_constraint set"
assumes only_non_strict: "lec_rel ` cs \<subseteq> {Leq_Rel}"
and fin: "finite cs"
shows "(\<nexists> v. \<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c) \<longleftrightarrow>
(\<exists> C const. (\<forall> (r, c) \<in> set C. r > 0 \<and> c \<in> cs)
\<and> (\<Sum> (r,c) \<leftarrow> C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const
\<and> const < 0)"
(is "?lhs = ?rhs")
proof -
{
fix c
assume "c \<in> cs"
with only_non_strict have "lec_rel c = Leq_Rel" by auto
then have "\<exists> p const. c = Leqc p const" by (cases c, auto)
} note leqc = this
let ?to_ns = "\<lambda> c. LEQ_ns (lec_poly c) (lec_const c)"
let ?ns = "?to_ns ` cs"
from fin have fin: "finite ?ns" by auto
have "v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns \<longleftrightarrow> (\<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c)" for v using leqc by fastforce
hence "?lhs = (\<nexists> v. v \<Turnstile>\<^sub>n\<^sub>s\<^sub>s ?ns)" by simp
also have "\<dots> = (\<exists>C. farkas_coefficients_ns ?ns C)" unfolding farkas_coefficients_ns[OF fin] ..
also have "\<dots> = ?rhs"
proof
assume "\<exists> C. farkas_coefficients_ns ?ns C"
then obtain C const where is_leq: "\<forall> (s, n) \<in> set C. n \<in> ?ns \<and> is_leq_ns (s *R n) \<and> s \<noteq> 0"
and sum: "(\<Sum>(s, n)\<leftarrow>C. lec_of_nsc (s *R n)) = Leqc 0 const"
and c0: "const < 0" unfolding farkas_coefficients_ns_def by blast
let ?C = "map (\<lambda> (s,n). (s,lec_of_nsc n)) C"
show ?rhs
proof (intro exI[of _ ?C] exI[of _ const] conjI c0, unfold sum[symmetric] map_map o_def set_map,
intro ballI, clarify)
{
fix s n
assume sn: "(s, n) \<in> set C"
with is_leq
have n_ns: "n \<in> ?ns" and is_leq: "is_leq_ns (s *R n)" "s \<noteq> 0" by force+
from n_ns obtain c where c: "c \<in> cs" and n: "n = LEQ_ns (lec_poly c) (lec_const c)" by auto
with leqc[OF c] obtain p d where cs: "Leqc p d \<in> cs" and n: "n = LEQ_ns p d" by auto
from is_leq[unfolded n] have s0: "s > 0" by (auto split: if_splits)
let ?n = "lec_of_nsc n"
from cs n have mem: "?n \<in> cs" by auto
show "0 < s \<and> ?n \<in> cs" using s0 mem by blast
have "Leqc (s *R lec_poly ?n) (s *R lec_const ?n) = lec_of_nsc (s *R n)"
unfolding n using s0 by simp
} note id = this
show "(\<Sum>x\<leftarrow>C. case case x of (s, n) \<Rightarrow> (s, lec_of_nsc n) of
(r, c) \<Rightarrow> Leqc (r *R lec_poly c) (r *R lec_const c)) =
(\<Sum>(s, n)\<leftarrow>C. lec_of_nsc (s *R n))" (is "sum_list (map ?f C) = sum_list (map ?g C)")
proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl])
fix pair
assume mem: "pair \<in> set C"
then obtain s n where pair: "pair = (s,n)" by force
show "?f pair = ?g pair" unfolding pair split using id[OF mem[unfolded pair]] .
qed
qed
next
assume ?rhs
then obtain C const
where C: "\<And> r c. (r,c) \<in> set C \<Longrightarrow> 0 < r \<and> c \<in> cs"
and sum: "(\<Sum>(r, c)\<leftarrow>C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const"
and const: "const < 0"
by blast
let ?C = "map (\<lambda> (r,c). (r, ?to_ns c)) C"
show "\<exists> C. farkas_coefficients_ns ?ns C" unfolding farkas_coefficients_ns_def
proof (intro exI[of _ ?C] exI[of _ const] conjI const, unfold sum[symmetric])
show "\<forall>(s, n)\<in>set ?C. n \<in> ?ns \<and> is_leq_ns (s *R n) \<and> s \<noteq> 0" using C by fastforce
show "(\<Sum>(s, n)\<leftarrow>?C. lec_of_nsc (s *R n))
= (\<Sum>(r, c)\<leftarrow>C. Leqc (r *R lec_poly c) (r *R lec_const c))"
unfolding map_map o_def
by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force)
qed
qed
finally show ?thesis .
qed
subsection \<open>Motzkin's Transposition Theorem or the Kuhn-Fourier Theorem\<close>
text \<open>Next, we prove a generalization of Farkas' Lemma that permits arbitrary combinations
of strict and non-strict inequalities: Motzkin's Transposition Theorem
which is also known as the Kuhn--Fourier Theorem.
The proof is mainly based on the lemma @{thm [source] farkas_coefficients},
again requiring conversions between constraint types.\<close>
theorem Motzkin's_transposition_theorem: fixes cs :: "rat le_constraint set"
assumes fin: "finite cs"
shows "(\<nexists> v. \<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c) \<longleftrightarrow>
(\<exists> C const rel. (\<forall> (r, c) \<in> set C. r > 0 \<and> c \<in> cs)
\<and> (\<Sum> (r,c) \<leftarrow> C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))
= Le_Constraint rel 0 const
\<and> (rel = Leq_Rel \<and> const < 0 \<or> rel = Lt_Rel \<and> const \<le> 0))"
(is "?lhs = ?rhs")
proof -
let ?to_cs = "\<lambda> c. (case lec_rel c of Leq_Rel \<Rightarrow> LEQ | _ \<Rightarrow> LT) (lec_poly c) (lec_const c)"
have to_cs: "v \<Turnstile>\<^sub>c ?to_cs c \<longleftrightarrow> v \<Turnstile>\<^sub>l\<^sub>e c" for v c by (cases c, cases "lec_rel c", auto)
let ?cs = "?to_cs ` cs"
from fin have fin: "finite ?cs" by auto
have "v \<Turnstile>\<^sub>c\<^sub>s ?cs \<longleftrightarrow> (\<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c)" for v using to_cs by auto
hence "?lhs = (\<nexists> v. v \<Turnstile>\<^sub>c\<^sub>s ?cs)" by simp
also have "\<dots> = (\<exists>C. farkas_coefficients ?cs C)" unfolding farkas_coefficients[OF fin] ..
also have "\<dots> = ?rhs"
proof
assume "\<exists> C. farkas_coefficients ?cs C"
then obtain C const rel where is_leq: "\<forall> (s, n) \<in> set C. n \<in> ?cs \<and> is_le (s *R n) \<and> s \<noteq> 0"
and sum: "(\<Sum>(s, n)\<leftarrow>C. lec_of_constraint (s *R n)) = Le_Constraint rel 0 const"
and c0: "(rel = Leq_Rel \<and> const < 0 \<or> rel = Lt_Rel \<and> const \<le> 0)"
unfolding farkas_coefficients_def by blast
let ?C = "map (\<lambda> (s,n). (s,lec_of_constraint n)) C"
show ?rhs
proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI c0, unfold map_map o_def set_map sum[symmetric],
intro ballI, clarify)
{
fix s n
assume sn: "(s, n) \<in> set C"
with is_leq
have n_ns: "n \<in> ?cs" and is_leq: "is_le (s *R n)" and s0: "s \<noteq> 0" by force+
from n_ns obtain c where c: "c \<in> cs" and n: "n = ?to_cs c" by auto
from is_leq[unfolded n] have "s \<ge> 0" by (cases "lec_rel c", auto split: if_splits)
with s0 have s0: "s > 0" by auto
let ?c = "lec_of_constraint n"
from c n have mem: "?c \<in> cs" by (cases c, cases "lec_rel c", auto)
show "0 < s \<and> ?c \<in> cs" using s0 mem by blast
have "lec_of_constraint (s *R n) = Le_Constraint (lec_rel ?c) (s *R lec_poly ?c) (s *R lec_const ?c)"
unfolding n using s0 by (cases c, cases "lec_rel c", auto)
} note id = this
show "(\<Sum>x\<leftarrow>C. case case x of (s, n) \<Rightarrow> (s, lec_of_constraint n) of
(r, c) \<Rightarrow> Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) =
(\<Sum>(s, n)\<leftarrow>C. lec_of_constraint (s *R n))"
(is "sum_list (map ?f C) = sum_list (map ?g C)")
proof (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl])
fix pair
assume mem: "pair \<in> set C"
obtain r c where pair: "pair = (r,c)" by force
show "?f pair = ?g pair" unfolding pair split id[OF mem[unfolded pair]] ..
qed
qed
next
assume ?rhs
then obtain C const rel
where C: "\<And> r c. (r,c) \<in> set C \<Longrightarrow> 0 < r \<and> c \<in> cs"
and sum: "(\<Sum>(r, c)\<leftarrow>C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))
= Le_Constraint rel 0 const"
and const: "rel = Leq_Rel \<and> const < 0 \<or> rel = Lt_Rel \<and> const \<le> 0"
by blast
let ?C = "map (\<lambda> (r,c). (r, ?to_cs c)) C"
show "\<exists> C. farkas_coefficients ?cs C" unfolding farkas_coefficients_def
proof (intro exI[of _ ?C] exI[of _ const] exI[of _ rel] conjI const, unfold sum[symmetric])
show "\<forall>(s, n)\<in>set ?C. n \<in> ?cs \<and> is_le (s *R n) \<and> s \<noteq> 0" using C by (fastforce split: le_rel.splits)
show "(\<Sum>(s, n)\<leftarrow>?C. lec_of_constraint (s *R n))
= (\<Sum>(r, c)\<leftarrow>C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))"
unfolding map_map o_def
by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, fastforce split: le_rel.splits)
qed
qed
finally show ?thesis .
qed
subsection \<open>Farkas' Lemma\<close>
text \<open>Finally we derive the commonly used form of Farkas' Lemma,
which easily follows from @{thm [source] Motzkin's_transposition_theorem}.
It only permits non-strict inequalities and, as a result,
the sum of inequalities will always be non-strict.\<close>
lemma Farkas'_Lemma: fixes cs :: "rat le_constraint set"
assumes only_non_strict: "lec_rel ` cs \<subseteq> {Leq_Rel}"
and fin: "finite cs"
shows "(\<nexists> v. \<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c) \<longleftrightarrow>
(\<exists> C const. (\<forall> (r, c) \<in> set C. r > 0 \<and> c \<in> cs)
\<and> (\<Sum> (r,c) \<leftarrow> C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const
\<and> const < 0)"
(is "_ = ?rhs")
proof -
{
fix c
assume "c \<in> cs"
with only_non_strict have "lec_rel c = Leq_Rel" by auto
then have "\<exists> p const. c = Leqc p const" by (cases c, auto)
} note leqc = this
let ?lhs = "\<exists>C const rel.
(\<forall>(r, c)\<in>set C. 0 < r \<and> c \<in> cs) \<and>
(\<Sum>(r, c)\<leftarrow>C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))
= Le_Constraint rel 0 const \<and>
(rel = Leq_Rel \<and> const < 0 \<or> rel = Lt_Rel \<and> const \<le> 0)"
show ?thesis unfolding Motzkin's_transposition_theorem[OF fin]
proof
assume ?rhs
then obtain C const where C: "\<And> r c. (r, c)\<in>set C \<Longrightarrow> 0 < r \<and> c \<in> cs" and
sum: "(\<Sum>(r, c)\<leftarrow>C. Leqc (r *R lec_poly c) (r *R lec_const c)) = Leqc 0 const" and
const: "const < 0" by blast
show ?lhs
proof (intro exI[of _ C] exI[of _ const] exI[of _ Leq_Rel] conjI)
show "\<forall> (r,c) \<in> set C. 0 < r \<and> c \<in> cs" using C by force
show "(\<Sum>(r, c)\<leftarrow> C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) =
Leqc 0 const" unfolding sum[symmetric]
by (rule arg_cong[of _ _ sum_list], rule map_cong[OF refl], insert C, force dest!: leqc)
qed (insert const, auto)
next
assume ?lhs
then obtain C const rel where C: "\<And> r c. (r, c)\<in>set C \<Longrightarrow> 0 < r \<and> c \<in> cs" and
sum: "(\<Sum>(r, c)\<leftarrow>C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c))
= Le_Constraint rel 0 const" and
const: "rel = Leq_Rel \<and> const < 0 \<or> rel = Lt_Rel \<and> const \<le> 0" by blast
have id: "(\<Sum>(r, c)\<leftarrow>C. Le_Constraint (lec_rel c) (r *R lec_poly c) (r *R lec_const c)) =
(\<Sum>(r, c)\<leftarrow>C. Leqc (r *R lec_poly c) (r *R lec_const c))" (is "_ = ?sum")
by (rule arg_cong[of _ _ sum_list], rule map_cong, auto dest!: C leqc)
have "lec_rel ?sum = Leq_Rel" unfolding sum_list_lec by (auto simp add: sum_list_Leq_Rel o_def)
with sum[unfolded id] have rel: "rel = Leq_Rel" by auto
with const have const: "const < 0" by auto
show ?rhs
by (intro exI[of _ C] exI[of _ const] conjI const, insert sum id C rel, force+)
qed
qed
text \<open>We also present slightly modified versions\<close>
lemma sum_list_map_filter_sum: fixes f :: "'a \<Rightarrow> 'b :: comm_monoid_add"
shows "sum_list (map f (filter g xs)) + sum_list (map f (filter (Not o g) xs)) = sum_list (map f xs)"
by (induct xs, auto simp: ac_simps)
text \<open>A version where every constraint obtains exactly one coefficient and where 0 coefficients are allowed.\<close>
lemma Farkas'_Lemma_set_sum: fixes cs :: "rat le_constraint set"
assumes only_non_strict: "lec_rel ` cs \<subseteq> {Leq_Rel}"
and fin: "finite cs"
shows "(\<nexists> v. \<forall> c \<in> cs. v \<Turnstile>\<^sub>l\<^sub>e c) \<longleftrightarrow>
(\<exists> C const. (\<forall> c \<in> cs. C c \<ge> 0)
\<and> (\<Sum> c \<in> cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const
\<and> const < 0)"
unfolding Farkas'_Lemma[OF only_non_strict fin]
proof ((standard; elim exE conjE), goal_cases)
case (2 C const)
from finite_distinct_list[OF fin] obtain csl where csl: "set csl = cs" and dist: "distinct csl"
by auto
let ?list = "filter (\<lambda> c. C c \<noteq> 0) csl"
let ?C = "map (\<lambda> c. (C c, c)) ?list"
show ?case
proof (intro exI[of _ ?C] exI[of _ const] conjI)
have "(\<Sum>(r, c)\<leftarrow>?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))
= (\<Sum>(r, c)\<leftarrow>map (\<lambda>c. (C c, c)) csl. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))"
unfolding map_map
by (rule sum_list_map_filter, auto simp: zero_le_constraint_def)
also have "\<dots> = Le_Constraint Leq_Rel 0 const" unfolding 2(2)[symmetric] csl[symmetric]
unfolding sum.distinct_set_conv_list[OF dist] map_map o_def split ..
finally
show "(\<Sum>(r, c)\<leftarrow>?C. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) = Le_Constraint Leq_Rel 0 const"
by auto
show "const < 0" by fact
show "\<forall>(r, c)\<in>set ?C. 0 < r \<and> c \<in> cs" using 2(1) unfolding set_map set_filter csl by auto
qed
next
case (1 C const)
define CC where "CC = (\<lambda> c. sum_list (map fst (filter (\<lambda> rc. snd rc = c) C)))"
show "(\<exists> C const. (\<forall> c \<in> cs. C c \<ge> 0)
\<and> (\<Sum> c \<in> cs. Leqc ((C c) *R lec_poly c) ((C c) *R lec_const c)) = Leqc 0 const
\<and> const < 0)"
proof (intro exI[of _ CC] exI[of _ const] conjI)
show "\<forall>c\<in>cs. 0 \<le> CC c" unfolding CC_def using 1(1)
by (force intro!: sum_list_nonneg)
show "const < 0" by fact
from 1 have snd: "snd ` set C \<subseteq> cs" by auto
show "(\<Sum>c\<in>cs. Le_Constraint Leq_Rel (CC c *R lec_poly c) (CC c *R lec_const c)) = Le_Constraint Leq_Rel 0 const"
unfolding 1(2)[symmetric] using fin snd unfolding CC_def
proof (induct cs arbitrary: C rule: finite_induct)
case empty
hence C: "C = []" by auto
thus ?case by simp
next
case *: (insert c cs C)
let ?D = "filter (Not \<circ> (\<lambda>rc. snd rc = c)) C"
from * have "snd ` set ?D \<subseteq> cs" by auto
note IH = *(3)[OF this]
have id: "(\<Sum>a\<leftarrow> ?D. case a of (r, c) \<Rightarrow> Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) =
(\<Sum>(r, c)\<leftarrow>?D. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c))"
by (induct C, force+)
show ?case
unfolding sum.insert[OF *(1,2)]
unfolding sum_list_map_filter_sum[of _ "\<lambda> rc. snd rc = c" C, symmetric]
proof (rule arg_cong2[of _ _ _ _ "(+)"], goal_cases)
case 2
show ?case unfolding IH[symmetric]
by (rule sum.cong, insert *(2,1), auto intro!: arg_cong[of _ _ "\<lambda> xs. sum_list (map _ xs)"], (induct C, auto)+)
next
case 1
show ?case
proof (rule sym, induct C)
case (Cons rc C)
thus ?case by (cases "rc", cases "snd rc = c", auto simp: field_simps scaleRat_left_distrib)
qed (auto simp: zero_le_constraint_def)
qed
qed
qed
qed
text \<open>A version with indexed constraints, i.e., in particular where constraints may occur several
times.\<close>
lemma Farkas'_Lemma_indexed: fixes c :: "nat \<Rightarrow> rat le_constraint"
assumes only_non_strict: "lec_rel ` c ` Is \<subseteq> {Leq_Rel}"
and fin: "finite Is"
shows "(\<nexists> v. \<forall> i \<in> Is. v \<Turnstile>\<^sub>l\<^sub>e c i) \<longleftrightarrow>
(\<exists> C const. (\<forall> i \<in> Is. C i \<ge> 0)
\<and> (\<Sum> i \<in> Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const
\<and> const < 0)"
proof -
let ?C = "c ` Is"
have fin: "finite ?C" using fin by auto
have "(\<nexists> v. \<forall> i \<in> Is. v \<Turnstile>\<^sub>l\<^sub>e c i) = (\<nexists> v. \<forall> cc \<in> ?C. v \<Turnstile>\<^sub>l\<^sub>e cc)" by force
also have "\<dots> = (\<exists> C const. (\<forall> i \<in> Is. C i \<ge> 0)
\<and> (\<Sum> i \<in> Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const
\<and> const < 0)" (is "?l = ?r")
proof
assume ?r
then obtain C const where r: "(\<forall> i \<in> Is. C i \<ge> 0)"
and eq: "(\<Sum> i \<in> Is. Leqc ((C i) *R lec_poly (c i)) ((C i) *R lec_const (c i))) = Leqc 0 const"
and "const < 0" by auto
- from finite_distinct_list[OF `finite Is`]
+ from finite_distinct_list[OF \<open>finite Is\<close>]
obtain Isl where isl: "set Isl = Is" and dist: "distinct Isl" by auto
let ?CC = "filter (\<lambda> rc. fst rc \<noteq> 0) (map (\<lambda> i. (C i, c i)) Isl)"
show ?l unfolding Farkas'_Lemma[OF only_non_strict fin]
proof (intro exI[of _ ?CC] exI[of _ const] conjI)
show "const < 0" by fact
show "\<forall> (r, ca) \<in> set ?CC. 0 < r \<and> ca \<in> ?C" using r(1) isl by auto
show "(\<Sum>(r, c)\<leftarrow>?CC. Le_Constraint Leq_Rel (r *R lec_poly c) (r *R lec_const c)) =
Le_Constraint Leq_Rel 0 const" unfolding eq[symmetric]
by (subst sum_list_map_filter, force simp: zero_le_constraint_def,
unfold map_map o_def, subst sum_list_distinct_conv_sum_set[OF dist], rule sum.cong, auto simp: isl)
qed
next
assume ?l
from this[unfolded Farkas'_Lemma_set_sum[OF only_non_strict fin]]
obtain C const where nonneg: "(\<forall>c\<in> ?C. 0 \<le> C c)"
and sum: "(\<Sum>c\<in> ?C. Le_Constraint Leq_Rel (C c *R lec_poly c) (C c *R lec_const c)) =
Le_Constraint Leq_Rel 0 const"
and const: "const < 0"
by blast
define I where "I = (\<lambda> i. (C (c i) / rat_of_nat (card (Is \<inter> { j. c i = c j}))))"
show ?r
proof (intro exI[of _ I] exI[of _ const] conjI const)
show "\<forall>i \<in> Is. 0 \<le> I i" using nonneg unfolding I_def by auto
show "(\<Sum> i \<in> Is. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i))) =
Le_Constraint Leq_Rel 0 const" unfolding sum[symmetric]
unfolding sum.image_gen[OF \<open>finite Is\<close>, of _ c]
proof (rule sum.cong[OF refl], goal_cases)
case (1 cc)
define II where "II = (Is \<inter> {j. cc = c j})"
from 1 have "II \<noteq> {}" unfolding II_def by auto
moreover have finII: "finite II" using \<open>finite Is\<close> unfolding II_def by auto
ultimately have card: "card II \<noteq> 0" by auto
let ?C = "\<lambda> II. rat_of_nat (card II)"
define ii where "ii = C cc / rat_of_nat (card II)"
have "(\<Sum>i\<in>{x \<in> Is. c x = cc}. Le_Constraint Leq_Rel (I i *R lec_poly (c i)) (I i *R lec_const (c i)))
= (\<Sum> i\<in> II. Le_Constraint Leq_Rel (ii *R lec_poly cc) (ii *R lec_const cc))"
unfolding I_def ii_def II_def by (rule sum.cong, auto)
also have "\<dots> = Le_Constraint Leq_Rel ((?C II * ii) *R lec_poly cc) ((?C II * ii) *R lec_const cc)"
using finII by (induct II rule: finite_induct, auto simp: zero_le_constraint_def field_simps
scaleRat_left_distrib)
also have "?C II * ii = C cc" unfolding ii_def using card by auto
finally show ?case .
qed
qed
qed
finally show ?thesis .
qed
end
\ No newline at end of file
diff --git a/thys/Generalized_Counting_Sort/Sorting.thy b/thys/Generalized_Counting_Sort/Sorting.thy
--- a/thys/Generalized_Counting_Sort/Sorting.thy
+++ b/thys/Generalized_Counting_Sort/Sorting.thy
@@ -1,1268 +1,1268 @@
(* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Proof of objects' sorting"
theory Sorting
imports Conservation
begin
text \<open>
\null
In this section, it is formally proven that GCsort actually sorts objects.
Here below, steps 5, 6, and 7 of the proof method are accomplished. Predicate @{text sort_inv} is
satisfied just in case, for any bucket delimited by the input counters' list $ns$, the keys of the
corresponding objects within the input objects' list $xs$ are not larger than those of the objects,
if any, to the right of that bucket. The underlying idea is that this predicate:
\begin{itemize}
\item
is trivially satisfied by the output of function @{const gcsort_in}, which places all objects into a
single bucket, and
\item
implies that $xs$ is sorted if every bucket delimited by $ns$ has size one, as happens when function
@{const gcsort_aux} terminates.
\end{itemize}
\null
\<close>
fun sort_inv :: "('a \<Rightarrow> 'b::linorder) \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow> bool" where
"sort_inv key (u, ns, xs) =
(\<forall>i < length ns. \<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<length xs}.
key (xs ! j) \<le> key (xs ! k))"
lemma gcsort_sort_input:
"sort_inv key (0, [length xs], xs)"
by simp
lemma offs_nth:
assumes
A: "find (\<lambda>n. Suc 0 < n) ns = None" and
B: "foldl (+) 0 ns = n" and
C: "k < n"
shows "\<exists>i < length ns. offs ns 0 ! i = k"
proof (cases ns, insert B C, simp, cases "k = 0", rule exI [of _ 0], simp,
rule ccontr, simp (no_asm_use))
fix m ms
assume
D: "ns = m # ms" and
E: "0 < k" and
F: "\<forall>i < length ns. offs ns 0 ! i \<noteq> k"
have G: "\<forall>n \<in> set ns. n \<le> Suc 0"
using A by (auto simp add: find_None_iff)
let ?A = "{i. i < length ns \<and> offs ns 0 ! i < k}"
have H: "Max ?A \<in> ?A"
using D and E by (rule_tac Max_in, simp_all, rule_tac exI [of _ 0], simp)
hence I: "Max ?A < length ns"
by simp
hence J: "offs ns 0 ! Max ?A = foldl (+) 0 (take (Max ?A) ns)"
by (rule offs_add)
have "Max ?A < length ns - Suc 0 \<or> Max ?A = length ns - Suc 0"
(is "?P \<or> ?Q")
using H by (simp, arith)
moreover {
assume ?P
hence K: "Suc (Max ?A) < length ns" by simp
hence "offs ns 0 ! Suc (Max ?A) = foldl (+) 0 (take (Suc (Max ?A)) ns)"
by (rule offs_add)
moreover have "take (Suc (Max ?A)) ns = take (Max ?A) ns @ [ns ! Max ?A]"
using I by (rule take_Suc_conv_app_nth)
ultimately have "offs ns 0 ! Suc (Max ?A) =
offs ns 0 ! Max ?A + ns ! Max ?A"
using J by simp
moreover have "offs ns 0 ! Max ?A < k"
using H by simp
moreover have "ns ! Max ?A \<in> set ns"
using I by (rule nth_mem)
with G have "ns ! Max ?A \<le> Suc 0" ..
ultimately have "offs ns 0 ! Suc (Max ?A) \<le> k" by simp
moreover have "offs ns 0 ! Suc (Max ?A) \<noteq> k"
using F and K by simp
ultimately have "offs ns 0 ! Suc (Max ?A) < k" by simp
hence "Suc (Max ?A) \<in> ?A"
using K by simp
hence "Suc (Max ?A) \<le> Max ?A"
by (rule_tac Max_ge, simp)
hence False by simp
}
moreover {
assume ?Q
hence "offs ns 0 ! Max ?A = foldl (+) 0 (take (length ns - Suc 0) ns)"
using J by simp
moreover have K: "length ns - Suc 0 < length ns"
using D by simp
hence "take (Suc (length ns - Suc 0)) ns =
take (length ns - Suc 0) ns @ [ns ! (length ns - Suc 0)]"
by (rule take_Suc_conv_app_nth)
hence "foldl (+) 0 ns =
foldl (+) 0 (take (length ns - Suc 0) ns @ [ns ! (length ns - Suc 0)])"
by simp
ultimately have "n = offs ns 0 ! Max ?A + ns ! (length ns - Suc 0)"
using B by simp
moreover have "offs ns 0 ! Max ?A < k"
using H by simp
moreover have "ns ! (length ns - Suc 0) \<in> set ns"
using K by (rule nth_mem)
with G have "ns ! (length ns - Suc 0) \<le> Suc 0" ..
ultimately have "n \<le> k" by simp
with C have False by simp
}
ultimately show False ..
qed
lemma gcsort_sort_intro:
"\<lbrakk>sort_inv key t; add_inv n t; find (\<lambda>n. Suc 0 < n) (fst (snd t)) = None\<rbrakk> \<Longrightarrow>
sorted (map key (gcsort_out t))"
proof (cases t, simp add: sorted_iff_nth_mono_less gcsort_out_def,
erule conjE, (rule allI)+, (rule impI)+)
fix ns xs j k
assume "find (\<lambda>n. Suc 0 < n) ns = None" and "foldl (+) 0 ns = n"
moreover assume A: "k < n"
ultimately have "\<exists>i < length ns. offs ns 0 ! i = k"
by (rule offs_nth)
then obtain i where "i < length ns \<and> offs ns 0 ! i = k" ..
moreover assume "\<forall>i < length ns. \<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n}.
key (xs ! j) \<le> key (xs ! k)"
hence "i < length ns \<longrightarrow> j < offs ns 0 ! i \<longrightarrow> k \<in> {offs ns 0 ! i..<n} \<longrightarrow>
key (xs ! j) \<le> key (xs ! k)"
by simp
ultimately have "j < k \<longrightarrow> k < n \<longrightarrow> key (xs ! j) \<le> key (xs ! k)"
by simp
moreover assume "j < k"
ultimately show "key (xs ! j) \<le> key (xs ! k)"
using A by simp
qed
text \<open>
\null
As lemma @{thm [source] gcsort_sort_intro} comprises an additional assumption concerning the form of
the fixed points of function @{const gcsort_aux}, step 8 of the proof method is necessary this time
to prove that such assumption is satisfied.
\null
\<close>
lemma gcsort_sort_form:
"find (\<lambda>n. Suc 0 < n) (fst (snd (gcsort_aux index key p t))) = None"
by (induction index key p t rule: gcsort_aux.induct, simp)
text \<open>
\null
Here below, step 9 of the proof method is accomplished.
In the most significant case of the proof by recursion induction of lemma @{text round_sort_inv},
namely that of a bucket $B$ with size larger than two and distinct minimum and maximum keys, the
following line of reasoning is adopted. Let $x$ be an object contained in a finer-grained bucket
$B'$ resulting from $B$'s partition, and $y$ an object to the right of $B'$. Then:
\begin{itemize}
\item
If $y$ is contained in some other finer-grained bucket resulting from $B$'s partition, inequality
@{term "key x \<le> key y"} holds because predicate @{const sort_inv} is satisfied by a counters' list
generated by function @{const enum} and an objects' list generated by function @{const fill} in case
@{const fill}'s input offsets' list is computed via the composition of functions @{const offs} and
@{const enum}, as happens within function @{const round}.
\\This is proven beforehand in lemma @{text fill_sort_inv}.
\item
Otherwise, inequality @{term "key x \<le> key y"} holds as well because object $x$ was contained in $B$
by lemma @{thm [source] fill_offs_enum_count_item}, object $y$ occurred to the right of $B$ by lemma
@{thm [source] round_count_inv}, and by hypothesis, the key of any object in $B$ was not larger than
that of any object to the right of $B$.
\end{itemize}
Using lemma @{text round_sort_inv}, the invariance of predicate @{const sort_inv} over inductive set
@{const gcsort_set} is then proven in lemma @{text gcsort_sort_inv}.
\null
\<close>
lemma mini_maxi_keys_le:
"x \<in> set xs \<Longrightarrow> key (xs ! mini xs key) \<le> key (xs ! maxi xs key)"
by (frule mini_lb, drule maxi_ub, erule order_trans)
lemma mini_maxi_keys_eq [rule_format]:
"key (xs ! mini xs key) = key (xs ! maxi xs key) \<longrightarrow> x \<in> set xs \<longrightarrow>
key x = key (xs ! maxi xs key)"
by (induction xs, simp_all add: Let_def, (rule_tac [!] impI, (rule_tac [!] conjI)?)+,
rule_tac [2-4] impI, frule_tac [1-3] mini_maxi_keys_le [where key = key], simp_all)
lemma offs_suc:
"i < length ns \<Longrightarrow> offs ns (Suc k) ! i = Suc (offs ns k ! i)"
by (simp add: offs_add add_suc)
lemma offs_base_zero:
"i < length ns \<Longrightarrow> offs ns k ! i = offs ns 0 ! i + k"
by (simp add: offs_add, subst add_base_zero, simp)
lemma offs_append:
"offs (ms @ ns) k = offs ms k @ offs ns (foldl (+) k ms)"
by (induction ms arbitrary: k, simp_all)
lemma offs_enum_next_le [rule_format]:
assumes
A: "index_less index key" and
B: "i < j" and
C: "j < n" and
D: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "offs_next (offs (enum xs index key n mi ma) k) (length xs + k)
xs index key mi ma i \<le> offs (enum xs index key n mi ma) k ! j"
(is "_ \<le> offs ?ns _ ! _")
proof (rule ccontr, simp add: not_le)
assume E: "offs ?ns k ! j <
offs_next (offs ?ns k) (length xs + k) xs index key mi ma i"
from B have "offs_set_next (offs ?ns k) xs index key mi ma i =
offs_set_next (offs ?ns k) xs index key mi ma j"
proof (rule_tac set_eqI, rule_tac iffI, simp_all, rule_tac ccontr, simp add: not_less)
fix m
assume "m < length (offs ?ns k) \<and> i < m \<and>
0 < offs_num (length (offs ?ns k)) xs index key mi ma m"
hence F: "m \<in> offs_set_next (offs ?ns k) xs index key mi ma i"
(is "_ \<in> ?A")
by simp
hence "Min ?A \<le> m"
by (rule_tac Min_le, simp)
moreover assume "m \<le> j"
ultimately have "offs ?ns k ! Min ?A \<le> offs ?ns k ! j"
using C by (rule_tac offs_mono, simp_all add: enum_length)
hence "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \<le>
offs ?ns k ! j"
using F by (simp only: offs_next_def split: if_split, rule_tac conjI,
blast, simp)
thus False
using E by simp
qed
hence "offs ?ns k ! j <
offs_next (offs ?ns k) (length xs + k) xs index key mi ma j"
using E by (simp only: offs_next_def)
moreover have "offs_num n xs index key mi ma j = 0"
proof (rule ccontr, simp)
assume "0 < offs_num n xs index key mi ma j"
hence "j \<in> offs_set_next (offs ?ns k) xs index key mi ma i"
(is "_ \<in> ?A")
using B and C by (simp add: offs_length enum_length)
moreover from this have "Min ?A \<le> j"
by (rule_tac Min_le, simp)
hence "offs ?ns k ! Min ?A \<le> offs ?ns k ! j"
using C by (erule_tac offs_mono, simp add: enum_length)
ultimately have "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \<le>
offs ?ns k ! j"
by (simp only: offs_next_def split: if_split, rule_tac conjI, blast, simp)
thus False
using E by simp
qed
hence "offs ?ns k ! j =
offs_next (offs ?ns k) (length xs + k) xs index key mi ma j"
using A and C and D by (rule_tac offs_enum_zero, simp_all)
ultimately show False by simp
qed
lemma offs_pred_ub_less:
"\<lbrakk>offs_pred ns ub xs index key mi ma; i < length ns;
0 < offs_num (length ns) xs index key mi ma i\<rbrakk> \<Longrightarrow>
ns ! i < ub"
by (drule offs_pred_ub, assumption+, simp)
lemma fill_count_none [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
ns \<noteq> [] \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
length xs \<le> ub \<longrightarrow>
count (mset (fill xs ns index key ub mi ma)) None = ub - length xs"
using A
proof (induction xs arbitrary: ns, simp_all add: replicate_count Let_def,
(rule_tac impI)+, (erule_tac conjE)+, subst mset_update, simp add: fill_length,
erule_tac offs_pred_ub_less, simp_all add: index_less_def offs_num_cons del:
not_None_eq, subst conj_commute, rule_tac conjI, rule_tac [!] impI, rotate_tac,
rotate_tac, erule_tac contrapos_np, rule_tac fill_index_none, simp_all)
fix x xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?ns' = "ns[?i := Suc (ns ! ?i)]"
assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
count (mset (fill xs ns index key ub mi ma)) None = ub - length xs"
moreover assume B: "ns \<noteq> []"
moreover assume
"offs_pred ns ub (x # xs) index key mi ma" and
"mi \<le> key x" and
"key x \<le> ma"
hence "offs_pred ?ns' ub xs index key mi ma"
using A and B by (rule_tac offs_pred_cons, simp_all add: index_less_def)
ultimately show "count (mset (fill xs ?ns' index key ub mi ma)) None -
Suc 0 = ub - Suc (length xs)"
by simp
qed
lemma fill_offs_enum_count_none [rule_format]:
"\<lbrakk>index_less index key; \<forall>x \<in> set xs. key x \<in> {mi..ma}; 0 < n\<rbrakk> \<Longrightarrow>
count (mset (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma)) None = 0"
by (subst fill_count_none, simp_all, simp only: length_greater_0_conv [symmetric]
offs_length enum_length, insert offs_enum_pred [of index key xs mi ma n 0], simp)
lemma fill_index [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
i < length ns \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma i \<longrightarrow>
j \<in> {ns ! i..<offs_next ns ub xs index key mi ma i} \<longrightarrow>
fill xs ns index key ub mi ma ! j = Some x \<longrightarrow>
index key x (length ns) mi ma = i"
proof (induction xs arbitrary: ns, simp add: offs_num_def, simp add: Let_def,
(rule impI)+, (erule conjE)+, simp)
fix y xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?i' = "index key y (length ns) mi ma"
let ?ns' = "ns[?i' := Suc (ns ! ?i')]"
let ?P = "\<lambda>ns.
offs_pred ns ub xs index key mi ma \<longrightarrow>
i < length ns \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma i \<longrightarrow>
ns ! i \<le> j \<and> j < offs_next ns ub xs index key mi ma i \<longrightarrow>
fill xs ns index key ub mi ma ! j = Some x \<longrightarrow>
index key x (length ns) mi ma = i"
assume
B: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
C: "mi \<le> key y" and
D: "key y \<le> ma" and
E: "offs_pred ns ub (y # xs) index key mi ma" and
F: "i < length ns" and
G: "0 < offs_num (length ns) (y # xs) index key mi ma i" and
H: "ns ! i \<le> j" and
I: "j < offs_next ns ub (y # xs) index key mi ma i" and
J: "\<And>ns. ?P ns" and
K: "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some y] ! j = Some x"
have "0 < length ns"
using F by arith
hence L: "?i' < length ns"
using A and C and D by (simp add: index_less_def)
hence "ns ! ?i' + offs_num (length ns) (y # xs) index key mi ma ?i' \<le> ub"
by (rule_tac offs_pred_ub [OF E], simp_all add: offs_num_cons)
hence "ns ! ?i' < ub"
by (simp add: offs_num_cons)
with K show "?i = i"
proof (cases "j = ns ! ?i'", simp, subst (asm) nth_list_update_eq, simp_all
add: fill_length)
assume
M: "j = ns ! ?i" and
N: "y = x"
show ?thesis
proof (rule ccontr, erule neqE)
assume "?i < i"
hence "ns ! ?i + offs_num (length ns) (y # xs) index key mi ma ?i \<le> ns ! i"
using F and G and N by (rule_tac offs_pred_asc [OF E], simp_all
add: offs_num_cons)
hence "j < ns ! i"
using M and N by (simp add: offs_num_cons)
thus False
using H by simp
next
let ?A = "offs_set_next ns (y # xs) index key mi ma i"
assume "i < ?i"
hence O: "?i \<in> ?A"
using N and L by (simp add: offs_num_cons)
hence P: "Min ?A \<in> ?A"
by (rule_tac Min_in, simp, blast)
have "Min ?A \<le> ?i"
using O by (rule_tac Min_le, simp)
moreover have "?A \<noteq> {}"
using O by blast
ultimately have "offs_next ns ub (y # xs) index key mi ma i \<le> j"
using M proof (simp only: offs_next_def, simp, subst (asm) le_less,
erule_tac disjE, simp_all)
assume "Min ?A < ?i"
hence "ns ! Min ?A + offs_num (length ns) (y # xs) index key mi ma
(Min ?A) \<le> ns ! ?i"
using O and P by (rule_tac offs_pred_asc [OF E], simp_all)
thus "ns ! Min ?A \<le> ns ! ?i" by simp
qed
thus False
using I by simp
qed
next
assume
M: "j \<noteq> ns ! ?i'" and
N: "fill xs ?ns' index key ub mi ma ! j = Some x"
have "?P ?ns'" using J .
moreover from D and F have "offs_pred ?ns' ub xs index key mi ma"
using L by (rule_tac offs_pred_cons [OF E], simp)
moreover have "i < length ?ns'"
using F by simp
moreover have "0 < offs_num (length ?ns') xs index key mi ma i"
proof (rule ccontr, simp)
assume O: "offs_num (length ns) xs index key mi ma i = 0"
hence P: "offs_num (length ns) (y # xs) index key mi ma i = Suc 0"
using G by (simp add: offs_num_cons split: if_split_asm)
hence "i = ?i'"
using O by (simp add: offs_num_cons split: if_split_asm)
hence "ns ! i < j"
using H and M by simp
hence "ns ! i + offs_num (length ns) (y # xs) index key mi ma i \<le> j"
using P by simp
with F and G and I have "offs_none ns ub (y # xs) index key mi ma j"
by (simp add: offs_none_def, rule_tac disjI1, rule_tac exI
[where x = i], simp)
with B and C and D and E and F have
"fill (y # xs) ns index key ub mi ma ! j = None"
by (rule_tac fill_none [OF A], simp_all, erule_tac disjE, simp_all, auto)
thus False
using K by (simp add: Let_def)
qed
moreover have "?ns' ! i \<le> j"
using F and H and M by (cases "?i' = i", simp_all)
moreover have "offs_next ns ub (y # xs) index key mi ma i \<le>
offs_next ?ns' ub xs index key mi ma i"
using G and L by (rule_tac offs_pred_next_cons [OF E], simp)
hence "j < offs_next ?ns' ub xs index key mi ma i"
using I by simp
ultimately show ?thesis
using N by simp
qed
qed
lemma fill_offs_enum_index [rule_format]:
"index_less index key \<Longrightarrow>
\<forall>x \<in> set xs. key x \<in> {mi..ma} \<Longrightarrow>
i < n \<Longrightarrow>
0 < offs_num n xs index key mi ma i \<Longrightarrow>
j \<in> {offs (enum xs index key n mi ma) 0 ! i..<
offs_next (offs (enum xs index key n mi ma) 0) (length xs)
xs index key mi ma i} \<Longrightarrow>
fill xs (offs (enum xs index key n mi ma) 0) index key (length xs)
mi ma ! j = Some x \<Longrightarrow>
index key x n mi ma = i"
by (insert fill_index [of index key xs mi ma "offs (enum xs index key n mi ma) 0"
"length xs" i j x], insert offs_enum_pred [of index key xs mi ma n 0],
simp add: offs_length enum_length)
lemma fill_sort_inv [rule_format]:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "sort_inv key (u, enum xs index key n mi ma,
map the (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma))"
(is "sort_inv _ (_, ?ns, _)")
proof (simp, (rule allI, rule impI)+, rule ballI, simp add: enum_length fill_length,
erule conjE)
fix i j k
let ?A = "{i. i < n \<and> 0 < ?ns ! i}"
let ?B = "{i. i < n \<and> offs ?ns 0 ! i \<le> j \<and> 0 < offs_num n xs index key mi ma i}"
let ?C = "{i. i < n \<and> offs ?ns 0 ! i \<le> k \<and> 0 < offs_num n xs index key mi ma i}"
assume
D: "i < n" and
E: "j < offs ?ns 0 ! i" and
F: "offs ?ns 0 ! i \<le> k" and
G: "k < length xs"
have H: "\<forall>i < length xs.
\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! i = Some x"
proof (rule allI, rule impI, rule ccontr, simp)
fix m
assume "fill xs (offs ?ns 0) index key (length xs) mi ma ! m = None"
moreover assume "m < length xs"
hence "fill xs (offs ?ns 0) index key (length xs) mi ma ! m
\<in> set (fill xs (offs ?ns 0) index key (length xs) mi ma)"
by (rule_tac nth_mem, simp add: fill_length)
ultimately have "None \<in> set (fill xs (offs ?ns 0) index key (length xs) mi ma)"
by simp
moreover have
"count (mset (fill xs (offs ?ns 0) index key (length xs) mi ma)) None = 0"
using C and D by (rule_tac fill_offs_enum_count_none [OF A], simp_all)
ultimately show False by simp
qed
have "\<exists>y ys. xs = y # ys"
using G by (rule_tac list.exhaust [of xs], simp_all)
then obtain z and zs where I: "xs = z # zs" by blast
hence "index key z n mi ma < n"
using A and C and D by (simp add: index_less_def)
moreover from this have "0 < ?ns ! index key z n mi ma"
using I by (simp add: Let_def, subst nth_list_update_eq, simp_all add:
enum_length)
ultimately have "index key z n mi ma \<in> ?A" by simp
hence J: "Min ?A \<in> ?A"
by (rule_tac Min_in, simp, blast)
hence "offs ?ns 0 ! 0 = offs ?ns 0 ! Min ?A"
proof (cases "Min ?A = 0", simp_all, erule_tac offs_equal, simp_all add:
enum_length, rule_tac ccontr, erule_tac conjE, simp)
fix m
assume "m < Min ?A" and "Min ?A < n" and "0 < ?ns ! m"
moreover from this have "Min ?A \<le> m"
by (rule_tac Min_le, simp_all)
ultimately show False by simp
qed
moreover have "\<exists>m ms. ?ns = m # ms"
using D by (rule_tac list.exhaust [of ?ns], simp_all,
simp only: length_0_conv [symmetric] enum_length)
then obtain m and ms where "?ns = m # ms" by blast
ultimately have "offs ?ns 0 ! Min ?A = 0" by simp
hence "Min ?A \<in> ?B"
using J by (simp, subst enum_offs_num [symmetric], simp_all)
hence K: "Max ?B \<in> ?B"
by (rule_tac Max_in, simp, blast)
moreover have "j < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?B)"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule_tac [2] ccontr, simp_all only: not_less)
show "j < length xs"
using E and F and G by simp
next
assume "offs_set_next (offs ?ns 0) xs index key mi ma (Max ?B) \<noteq> {}"
(is "?Z \<noteq> _")
hence L: "Min ?Z \<in> ?Z"
by (rule_tac Min_in, simp)
moreover assume "offs ?ns 0 ! Min ?Z \<le> j"
ultimately have "Min ?Z \<in> ?B"
by (simp add: offs_length enum_length)
hence "Min ?Z \<le> Max ?B"
by (rule_tac Max_ge, simp)
thus False
using L by simp
qed
moreover have
"\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! j = Some x"
using E and F and G and H by simp
then obtain x where
L: "fill xs (offs ?ns 0) index key (length xs) mi ma ! j = Some x" ..
ultimately have M: "index key x n mi ma = Max ?B"
using C by (rule_tac fill_offs_enum_index [OF A], simp_all)
have N: "Max ?B \<in> ?C"
using E and F and K by simp
hence "Max ?C \<in> ?C"
by (rule_tac Max_in, simp, blast)
moreover have O: "k < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?C)"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule_tac [2] ccontr, simp_all only: not_less)
show "k < length xs" using G .
next
assume "offs_set_next (offs ?ns 0) xs index key mi ma (Max ?C) \<noteq> {}"
(is "?Z \<noteq> _")
hence P: "Min ?Z \<in> ?Z"
by (rule_tac Min_in, simp)
moreover assume "offs ?ns 0 ! Min ?Z \<le> k"
ultimately have "Min ?Z \<in> ?C"
by (simp add: offs_length enum_length)
hence "Min ?Z \<le> Max ?C"
by (rule_tac Max_ge, simp)
thus False
using P by simp
qed
moreover have
"\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! k = Some x"
using G and H by simp
then obtain y where
P: "fill xs (offs ?ns 0) index key (length xs) mi ma ! k = Some y" ..
ultimately have Q: "index key y n mi ma = Max ?C"
using C by (rule_tac fill_offs_enum_index [OF A], simp_all)
have "Max ?B \<le> Max ?C"
using N by (rule_tac Max_ge, simp)
hence "Max ?B < Max ?C"
proof (rule_tac ccontr, simp)
assume "Max ?B = Max ?C"
hence "offs ?ns 0 ! i < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?B)"
using F and O by simp
moreover have "offs ?ns 0 ! Max ?B < offs ?ns 0 ! i"
using E and K by simp
with K have "Max ?B < i"
by (erule_tac contrapos_pp, subst not_less, subst (asm) not_less,
erule_tac offs_mono, simp add: enum_length)
hence "offs_next (offs ?ns 0) (length xs + 0) xs index key mi ma (Max ?B) \<le>
offs ?ns 0 ! i"
using C and D by (rule_tac offs_enum_next_le [OF A], simp_all)
ultimately show False by simp
qed
hence R: "index key x n mi ma < index key y n mi ma"
using M and Q by simp
have "count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) x = count (mset xs) x"
using C and D by (rule_tac fill_offs_enum_count_item [OF A], simp_all)
moreover have S: "j < length (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))"
using E and F and G by (simp add: fill_length)
hence "map the (fill xs (offs ?ns 0) index key (length xs) mi ma) ! j
\<in> set (map the (fill xs (offs ?ns 0) index key (length xs) mi ma))"
by (rule nth_mem)
hence "0 < count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) x"
using L and S by simp
ultimately have T: "x \<in> set xs" by simp
have "count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) y = count (mset xs) y"
using C and D by (rule_tac fill_offs_enum_count_item [OF A], simp_all)
moreover have U: "k < length (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))"
using G by (simp add: fill_length)
hence "map the (fill xs (offs ?ns 0) index key (length xs) mi ma) ! k
\<in> set (map the (fill xs (offs ?ns 0) index key (length xs) mi ma))"
by (rule nth_mem)
hence "0 < count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) y"
using P and U by simp
ultimately have V: "y \<in> set xs" by simp
have "key y \<le> key x \<longrightarrow> index key y n mi ma \<le> index key x n mi ma"
using B and C and T and V by (simp add: index_mono_def)
hence "key x < key y"
by (rule_tac contrapos_pp [OF R], simp add: not_less)
thus "key (the (fill xs (offs ?ns 0) index key (length xs) mi ma ! j)) \<le>
key (the (fill xs (offs ?ns 0) index key (length xs) mi ma ! k))"
using L and P by simp
qed
lemma round_sort_inv [rule_format]:
"index_less index key \<longrightarrow> index_mono index key \<longrightarrow> bn_inv p q t \<longrightarrow>
add_inv n t \<longrightarrow> sort_inv key t \<longrightarrow> sort_inv key (round index key p q r t)"
proof (induction index key p q r t arbitrary: n rule: round.induct,
(rule_tac [!] impI)+, simp, simp_all only: simp_thms)
fix index p q r u ns xs n and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, xs)"
assume "\<And>n. bn_inv p q (u, ns, xs) \<longrightarrow> add_inv n (u, ns, xs) \<longrightarrow>
sort_inv key (u, ns, xs) \<longrightarrow> sort_inv key ?t"
hence "bn_inv p q (u, ns, xs) \<longrightarrow> add_inv n (u, ns, xs) \<longrightarrow>
sort_inv key (u, ns, xs) \<longrightarrow> sort_inv key ?t" .
moreover assume
"bn_inv p q (u, 0 # ns, xs)"
"add_inv n (u, 0 # ns, xs)" and
"sort_inv key (u, 0 # ns, xs)"
ultimately show "sort_inv key (round index key p q r (u, 0 # ns, xs))"
by auto
next
fix index p q r u ns xs n and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume
A: "index_less index key" and
B: "bn_inv p q (u, Suc 0 # ns, xs)"
moreover assume
"\<And>n. bn_inv p q (u, ns, tl xs) \<longrightarrow> add_inv n (u, ns, tl xs) \<longrightarrow>
sort_inv key (u, ns, tl xs) \<longrightarrow> sort_inv key ?t" and
"add_inv n (u, Suc 0 # ns, xs)" and
"sort_inv key (u, Suc 0 # ns, xs)"
ultimately show "sort_inv key (round index key p q r (u, Suc 0 # ns, xs))"
proof (cases ?t, cases xs, simp_all add: add_suc, (rule_tac allI, rule_tac impI)+,
rule_tac ballI, simp, (erule_tac conjE)+, simp)
fix i j k y ys u' ns' xs'
assume
C: "round index key p q r (u, ns, ys) = (u', ns', xs')" and
D: "Suc (foldl (+) 0 ns) = n" and
E: "Suc (length ys) = n" and
F: "\<forall>i < Suc (length ns).
\<forall>j < (0 # offs ns (Suc 0)) ! i. \<forall>k \<in> {(0 # offs ns (Suc 0)) ! i..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0))"
assume A': "j < (0 # offs ns' (Suc 0)) ! i"
hence "\<exists>i'. i = Suc i'"
by (rule_tac nat.exhaust [of i], simp_all)
then obtain i' where B': "i = Suc i'" ..
assume "i < Suc (length ns')"
hence G: "i' < length ns'"
using B' by simp
hence H: "j < Suc (offs ns' 0 ! i')"
using A' and B' by (simp add: offs_suc)
assume "(0 # offs ns' (Suc 0)) ! i \<le> k"
hence "Suc (offs ns' 0 ! i') \<le> k"
using B' and G by (simp add: offs_suc)
moreover from this have "\<exists>k'. k = Suc k'"
by (rule_tac nat.exhaust [of k], simp_all)
then obtain k' where I: "k = Suc k'" ..
ultimately have J: "offs ns' 0 ! i' \<le> k'" by simp
assume "k < Suc (length xs')"
hence K: "k' < length xs'"
using I by simp
let ?P = "\<lambda>n. foldl (+) 0 ns = n \<and> length ys = n"
let ?Q = "\<lambda>n. \<forall>i < length ns.
\<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n}.
key (ys ! j) \<le> key (ys ! k)"
let ?R = "\<forall>i < length ns'.
\<forall>j < offs ns' 0 ! i. \<forall>k \<in> {offs ns' 0 ! i..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
assume "\<And>n. ?P n \<longrightarrow> ?Q n \<longrightarrow> ?R"
hence "?P (n - Suc 0) \<longrightarrow> ?Q (n - Suc 0) \<longrightarrow> ?R" .
hence "?Q (n - Suc 0) \<longrightarrow> ?R"
using D and E by auto
moreover have "?Q (n - Suc 0)"
proof ((rule allI, rule impI)+, rule ballI, simp, erule conjE)
fix i j k
have "Suc i < Suc (length ns) \<longrightarrow> (\<forall>j < (0 # offs ns (Suc 0)) ! Suc i.
\<forall>k \<in> {(0 # offs ns (Suc 0)) ! Suc i..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0)))"
using F ..
moreover assume "i < length ns"
ultimately have "Suc j < Suc (offs ns 0 ! i) \<longrightarrow>
(\<forall>k \<in> {Suc (offs ns 0 ! i)..<n}.
key ((y # ys) ! Suc j) \<le> key (ys ! (k - Suc 0)))"
by (auto simp add: offs_suc)
moreover assume "j < offs ns 0 ! i"
ultimately have "\<forall>k \<in> {Suc (offs ns 0 ! i)..<n}.
key (ys ! j) \<le> key (ys ! (k - Suc 0))"
by simp
moreover assume "offs ns 0 ! i \<le> k" and "k < n - Suc 0"
hence "Suc k \<in> {Suc (offs ns 0 ! i)..<n}" by simp
ultimately have "key (ys ! j) \<le> key (ys ! (Suc k - Suc 0))" ..
thus "key (ys ! j) \<le> key (ys ! k)" by simp
qed
ultimately have ?R ..
from I show "key ((y # xs') ! j) \<le> key (xs' ! (k - Suc 0))"
proof (cases j, simp_all)
have "bn_inv p q (u, ns, ys)"
using B by simp
moreover have "add_inv (n - Suc 0) (u, ns, ys)"
using D and E by auto
moreover have "count_inv (\<lambda>x. count (mset ys) x) (u, ns, ys)" by simp
ultimately have "count_inv (\<lambda>x. count (mset ys) x)
(round index key p q r (u, ns, ys))"
by (rule round_count_inv [OF A])
hence "count (mset xs') (xs' ! k') = count (mset ys) (xs' ! k')"
using C by simp
moreover have "0 < count (mset xs') (xs' ! k')"
using K by simp
ultimately have "xs' ! k' \<in> set ys" by simp
hence "\<exists>m < length ys. ys ! m = xs' ! k'"
by (simp add: in_set_conv_nth)
then obtain m where L: "m < length ys \<and> ys ! m = xs' ! k'" ..
have "Suc 0 < Suc (length ns) \<longrightarrow> (\<forall>j < (0 # offs ns (Suc 0)) ! Suc 0.
\<forall>k \<in> {(0 # offs ns (Suc 0)) ! Suc 0..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0)))"
using F ..
moreover have M: "0 < length ns"
using D [symmetric] and E and L by (rule_tac ccontr, simp)
ultimately have "\<forall>k \<in> {Suc (offs ns 0 ! 0)..<n}.
key y \<le> key (ys ! (k - Suc 0))"
by (auto simp add: offs_suc)
hence "\<forall>k \<in> {Suc 0..<n}. key y \<le> key (ys ! (k - Suc 0))"
using M by (cases ns, simp_all)
moreover have "Suc m \<in> {Suc 0..<n}"
using E and L by simp
ultimately have "key y \<le> key (ys ! (Suc m - Suc 0))" ..
thus "key y \<le> key (xs' ! k')"
using L by simp
next
case (Suc j')
moreover have "i' < length ns' \<longrightarrow> (\<forall>j < offs ns' 0 ! i'.
\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}. key (xs' ! j) \<le> key (xs' ! k))"
- using `?R` ..
+ using \<open>?R\<close> ..
hence "j' < offs ns' 0 ! i' \<longrightarrow>
(\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}. key (xs' ! j') \<le> key (xs' ! k))"
using G by simp
ultimately have "\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}.
key (xs' ! j') \<le> key (xs' ! k)"
using H by simp
moreover have "k' \<in> {offs ns' 0 ! i'..<length xs'}"
using J and K by simp
ultimately show "key (xs' ! j') \<le> key (xs' ! k')" ..
qed
qed
next
fix index p q r u m ns n and key :: "'a \<Rightarrow> 'b" and xs :: "'a list"
let ?ws = "take (Suc (Suc m)) xs"
let ?ys = "drop (Suc (Suc m)) xs"
let ?r = "\<lambda>m'. round_suc_suc index key ?ws m m' u"
let ?t = "\<lambda>r' v. round index key p q r' (v, ns, ?ys)"
assume
A: "index_less index key" and
B: "index_mono index key" and
C: "bn_inv p q (u, Suc (Suc m) # ns, xs)"
assume
"\<And>ws a b c d e f g h n.
ws = ?ws \<Longrightarrow> a = bn_comp m p q r \<Longrightarrow> (b, c) = bn_comp m p q r \<Longrightarrow>
d = ?r b \<Longrightarrow> (e, f) = ?r b \<Longrightarrow> (g, h) = f \<Longrightarrow>
bn_inv p q (e, ns, ?ys) \<longrightarrow> add_inv n (e, ns, ?ys) \<longrightarrow>
sort_inv key (e, ns, ?ys) \<longrightarrow> sort_inv key (?t c e)" and
"add_inv n (u, Suc (Suc m) # ns, xs)" and
"sort_inv key (u, Suc (Suc m) # ns, xs)"
with C show "sort_inv key (round index key p q r (u, Suc (Suc m) # ns, xs))"
using [[simproc del: defined_all]]
proof (simp split: prod.split, ((rule_tac allI)+, (rule_tac impI)+)+,
rule_tac ballI, simp, (erule_tac conjE)+, subst (asm) (2) add_base_zero, simp)
fix m' r' v ms' ws' u' ns' xs' i j k
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case m of 0 \<Rightarrow> m | Suc 0 \<Rightarrow> m | Suc (Suc i) \<Rightarrow> u + m'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
let ?P = "\<lambda>n'. foldl (+) 0 ns = n' \<and> n - Suc (Suc m) = n'"
let ?Q = "\<lambda>n'. \<forall>i < length ns.
\<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n'}.
key (xs ! Suc (Suc (m + j))) \<le> key (xs ! Suc (Suc (m + k)))"
let ?R = "\<forall>i < length ns'.
\<forall>j < offs ns' 0 ! i. \<forall>k \<in> {offs ns' 0 ! i..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
assume
D: "?r m' = (v, ms', ws')" and
E: "?t r' v = (u', ns', xs')" and
F: "bn_comp m p q r = (m', r')" and
G: "bn_valid m p q" and
H: "Suc (Suc (foldl (+) 0 ns + m)) = n" and
I: "length xs = n" and
J: "\<forall>i < Suc (length ns). \<forall>j < (0 # offs ns (Suc (Suc m))) ! i.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! i..<n}.
key (xs ! j) \<le> key (xs ! k)" and
K: "i < length ms' + length ns'" and
L: "j < offs (ms' @ ns') 0 ! i" and
M: "offs (ms' @ ns') 0 ! i \<le> k" and
N: "k < length ws' + length xs'"
have O: "length ?ws = Suc (Suc m)"
using H and I by simp
with D [symmetric] have P: "length ws' = Suc (Suc m)"
by (simp add: round_suc_suc_def Let_def fill_length split: if_split_asm)
have Q: "\<And>i. i < m \<Longrightarrow>
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! i) \<in> set ?ws"
proof -
fix i
let ?x = "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! i)"
assume R: "i < m"
hence "0 < m" by simp
hence "0 < ?k"
by (drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
hence "count (mset (map the (fill ?zs (offs ?ms 0)
index key (length ?zs) ?mi ?ma))) ?x = count (mset ?zs) ?x"
by (rule_tac fill_offs_enum_count_item [OF A], simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
hence "count (mset (map the (fill ?zs (offs ?ms 0)
index key m ?mi ?ma))) ?x = count (mset ?zs) ?x"
using O by (simp add: mini_maxi_nths)
moreover have "0 < count (mset (map the (fill ?zs (offs ?ms 0)
index key m ?mi ?ma))) ?x"
using R by (simp add: fill_length)
ultimately have "?x \<in> set ?zs" by simp
thus "?x \<in> set ?ws"
by (rule in_set_nthsD)
qed
have R: "\<And>i. i < Suc (Suc m) \<Longrightarrow> Suc (Suc m) \<le> k \<Longrightarrow>
key (?ws ! i) \<le> key (xs' ! (k - Suc (Suc m)))"
proof -
fix i
have "bn_inv p q (v, ns, ?ys)"
using C by simp
moreover have "add_inv (n - Suc (Suc m)) (v, ns, ?ys)"
using H and I by simp
moreover have "count_inv (\<lambda>x. count (mset ?ys) x) (v, ns, ?ys)"
by simp
ultimately have "count_inv (\<lambda>x. count (mset ?ys) x) (?t r' v)"
by (rule round_count_inv [OF A])
hence S: "count (mset xs') (xs' ! (k - Suc (Suc m))) =
count (mset ?ys) (xs' ! (k - Suc (Suc m)))"
using E by simp
have "k < Suc (Suc (m + length xs'))"
using N and P by simp
moreover assume "Suc (Suc m) \<le> k"
ultimately have "0 < count (mset xs') (xs' ! (k - Suc (Suc m)))"
by simp
hence "xs' ! (k - Suc (Suc m)) \<in> set ?ys"
using S by simp
hence "\<exists>p < length ?ys. ?ys ! p = xs' ! (k - Suc (Suc m))"
by (simp add: in_set_conv_nth)
then obtain p where
T: "p < length ?ys \<and> ?ys ! p = xs' ! (k - Suc (Suc m))" ..
have "Suc 0 < Suc (length ns) \<longrightarrow>
(\<forall>j < (0 # offs ns (Suc (Suc m))) ! Suc 0.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! Suc 0..<n}.
key (xs ! j) \<le> key (xs ! k))"
using J ..
moreover have U: "0 < length ns"
using H [symmetric] and I and T by (rule_tac ccontr, simp)
ultimately have "\<forall>j < offs ns (Suc (Suc m)) ! 0.
\<forall>k \<in> {offs ns (Suc (Suc m)) ! 0..<n}. key (xs ! j) \<le> key (xs ! k)"
by simp
moreover assume "i < Suc (Suc m)"
moreover have "offs ns (Suc (Suc m)) ! 0 = Suc (Suc m)"
using U by (subst offs_base_zero, simp, cases ns, simp_all)
ultimately have "\<forall>k \<in> {Suc (Suc m)..<n}. key (?ws ! i) \<le> key (xs ! k)"
by simp
moreover have "Suc (Suc m) + p \<in> {Suc (Suc m)..<n}"
using I and T by (simp, arith)
ultimately have "key (?ws ! i) \<le> key (xs ! (Suc (Suc m) + p))" ..
thus "key (?ws ! i) \<le> key (xs' ! (k - Suc (Suc m)))"
using O and T by simp
qed
assume "\<And>ws a b c d e f g h n'.
ws = ?ws \<Longrightarrow> a = (m', r') \<Longrightarrow> b = m' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> f = (ms', ws') \<Longrightarrow> g = ms' \<and> h = ws' \<Longrightarrow>
?P n' \<longrightarrow> ?Q n' \<longrightarrow> ?R"
hence "?P (n - Suc (Suc m)) \<longrightarrow> ?Q (n - Suc (Suc m)) \<longrightarrow> ?R"
by simp
hence "?Q (n - Suc (Suc m)) \<longrightarrow> ?R"
using H by simp
moreover have "?Q (n - Suc (Suc m))"
proof ((rule allI, rule impI)+, rule ballI, simp, erule conjE,
subst (1 2) append_take_drop_id [of "Suc (Suc m)", symmetric],
simp only: nth_append O, simp)
fix i j k
have "Suc i < Suc (length ns) \<longrightarrow>
(\<forall>j < (0 # offs ns (Suc (Suc m))) ! Suc i.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! Suc i..<n}.
key ((xs) ! j) \<le> key (xs ! k))"
using J ..
moreover assume "i < length ns"
ultimately have "j < offs ns 0 ! i \<longrightarrow>
(\<forall>k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}.
key (xs ! (Suc (Suc m) + j)) \<le> key (xs ! k))"
by (simp, subst (asm) offs_base_zero, simp,
subst (asm) (2) offs_base_zero, simp_all)
moreover assume "j < offs ns 0 ! i"
ultimately have "\<forall>k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}.
key (?ys ! j) \<le> key (xs ! k)"
using O by simp
moreover assume "offs ns 0 ! i \<le> k" and "k < n - Suc (Suc m)"
hence "Suc (Suc m) + k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}"
by simp
ultimately have "key (?ys ! j) \<le> key (xs ! (Suc (Suc m) + k))" ..
thus "key (?ys ! j) \<le> key (?ys ! k)"
using O by simp
qed
ultimately have ?R ..
show "key ((ws' @ xs') ! j) \<le> key ((ws' @ xs') ! k)"
proof (simp add: nth_append not_less P, (rule_tac [!] conjI,
rule_tac [!] impI)+, rule_tac [3] FalseE)
assume
S: "j < Suc (Suc m)" and
T: "k < Suc (Suc m)"
from D [symmetric] show "key (ws' ! j) \<le> key (ws' ! k)"
proof (simp add: round_suc_suc_def Let_def split: if_split_asm,
(erule_tac [2] conjE)+, simp_all)
assume U: "?mi = ?ma"
have "j < length ?ws"
using S and O by simp
hence "?ws ! j \<in> set ?ws"
by (rule nth_mem)
with U have "key (?ws ! j) = ?ma"
by (rule mini_maxi_keys_eq)
moreover have "k < length ?ws"
using T and O by simp
hence "?ws ! k \<in> set ?ws"
by (rule nth_mem)
with U have "key (?ws ! k) = ?ma"
by (rule mini_maxi_keys_eq)
ultimately show "key (?ws ! j) \<le> key (?ws ! k)" by simp
next
assume U: "ms' = Suc 0 # ?ms @ [Suc 0]"
hence V: "j < (0 # offs (?ms @ Suc 0 # ns') (Suc 0)) ! i"
using L by simp
hence "\<exists>i'. i = Suc i'"
by (rule_tac nat.exhaust [of i], simp_all)
then obtain i' where W: "i = Suc i'" ..
have "i < Suc (Suc (?k + length ns'))"
using K and U by (simp add: enum_length)
hence X: "i' < Suc (?k + length ns')"
using W by simp
hence Y: "j < Suc (offs (?ms @ Suc 0 # ns') 0 ! i')"
using V and W by (simp add: enum_length offs_suc)
have "(0 # offs (?ms @ Suc 0 # ns') (Suc 0)) ! i \<le> k"
using M and U by simp
hence "Suc (offs (?ms @ Suc 0 # ns') 0 ! i') \<le> k"
using W and X by (simp add: enum_length offs_suc)
moreover from this have "\<exists>k'. k = Suc k'"
by (rule_tac nat.exhaust [of k], simp_all)
then obtain k' where Z: "k = Suc k'" ..
ultimately have AA: "offs (?ms @ Suc 0 # ns') 0 ! i' \<le> k'"
by simp
have AB: "j \<le> k'"
using Y and AA by simp
with Z show
"key ((?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @
[?xma]) ! j) \<le>
key ((?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @
[?xma]) ! k)"
proof (cases j, case_tac [2] "j < Suc m", simp_all add: nth_append
not_less fill_length, rule_tac [1-2] conjI, rule_tac [1-4] impI,
rule_tac [5] FalseE, simp_all)
assume "k' < m"
hence "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k') \<in> set ?ws"
by (rule Q)
thus "?mi \<le> key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k'))"
by (rule mini_lb)
next
assume "m \<le> k'"
hence "k' = m"
using T and Z by simp
thus "?mi \<le> key ([?xma] ! (k' - m))"
by (simp, rule_tac mini_maxi_keys_le, rule_tac nth_mem [of 0],
subst O, simp)
next
fix j'
have "sort_inv key (0, ?ms, map the (fill ?zs (offs ?ms 0) index key
(length ?zs) ?mi ?ma))"
by (rule fill_sort_inv [OF A B], simp, rule conjI,
((rule mini_lb | rule maxi_ub), erule in_set_nthsD)+)
hence "sort_inv key (0, ?ms, map the (fill ?zs (offs ?ms 0) index key
m ?mi ?ma))"
using O by (simp add: mini_maxi_nths)
hence "\<forall>i < ?k. \<forall>j < offs ?ms 0 ! i. \<forall>k \<in> {offs ?ms 0 ! i..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j)) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by (simp add: enum_length fill_length)
moreover assume AC: "k' < m"
hence AD: "offs (?ms @ Suc 0 # ns') 0 ! i' < m"
using AA by simp
have AE: "i' < ?k"
proof (rule contrapos_pp [OF AD], simp add: not_less offs_append
nth_append offs_length enum_length)
have "0 < m"
using AC by simp
hence "0 < ?k"
by (drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
with A have "foldl (+) 0 ?ms = length ?zs"
by (rule enum_add, simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
hence "foldl (+) 0 ?ms = m"
using O by (simp add: mini_maxi_nths)
with X show "m \<le> (foldl (+) 0 ?ms #
offs ns' (Suc (foldl (+) 0 ?ms))) ! (i' - ?k)"
by (cases "i' - ?k", simp_all, subst offs_base_zero, simp_all)
qed
ultimately have "\<forall>j < offs ?ms 0 ! i'. \<forall>k \<in> {offs ?ms 0 ! i'..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j)) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by simp
moreover assume "j = Suc j'"
with Y and AE have "j' < offs ?ms 0 ! i'"
by (simp add: offs_append nth_append offs_length enum_length)
ultimately have "\<forall>k \<in> {offs ?ms 0 ! i'..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by simp
moreover from AA and AE have "offs ?ms 0 ! i' \<le> k'"
by (simp add: offs_append nth_append offs_length enum_length)
hence "k' \<in> {offs ?ms 0 ! i'..<m}"
using AC by simp
ultimately show
"key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k'))" ..
next
fix j'
assume "j' < m"
hence "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j') \<in> set ?ws"
(is "?x \<in> _")
by (rule Q)
moreover assume "m \<le> k'"
hence "k' = m"
using T and Z by simp
ultimately show "key ?x \<le> key ([?xma] ! (k' - m))"
by (simp, rule_tac maxi_ub)
next
fix j'
assume "j = Suc j'" and "m \<le> j'"
hence "Suc m \<le> j" by simp
hence "Suc (Suc m) \<le> k"
using Z and AB by simp
thus False
using T by simp
qed
qed
next
assume
S: "j < Suc (Suc m)" and
T: "Suc (Suc m) \<le> k"
from D [symmetric] show "key (ws' ! j) \<le> key (xs' ! (k - Suc (Suc m)))"
proof (simp add: round_suc_suc_def Let_def split: if_split_asm,
rule_tac R [OF S T], cases j, case_tac [2] "j < Suc m",
simp_all add: nth_append not_less fill_length)
have "?nmi < length ?ws"
using O by (rule_tac mini_less, simp)
hence "?nmi < Suc (Suc m)"
using O by simp
thus "?mi \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule R)
next
fix j'
assume "j' < m"
hence U: "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j') \<in> set ?ws"
by (rule Q)
have "\<exists>p < Suc (Suc m). ?ws ! p =
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')"
using O and U by (simp add: in_set_conv_nth)
then obtain p where V: "p < Suc (Suc m) \<and> ?ws ! p =
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')" ..
hence "key (?ws ! p) \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule_tac R, simp)
thus "key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (xs' ! (k - Suc (Suc m)))"
using V by simp
next
fix j'
assume "j = Suc j'" and "m \<le> j'"
moreover from this have "Suc m \<le> j" by simp
hence "j = Suc m"
using S by simp
ultimately have "j' = m" by simp
thus "key ([?xma] ! (j' - m)) \<le> key (xs' ! (k - Suc (Suc m)))"
proof simp
have "?nma < length ?ws"
using O by (rule_tac maxi_less, simp)
hence "?nma < Suc (Suc m)"
using O by simp
thus "?ma \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule R)
qed
qed
next
assume "k < Suc (Suc m)" and "Suc (Suc m) \<le> j"
hence "k < j" by simp
moreover have "j < k"
using L and M by simp
ultimately show False by simp
next
assume
S: "Suc (Suc m) \<le> j" and
T: "Suc (Suc m) \<le> k"
have U: "0 < length ns' \<and> 0 < length xs'"
proof (rule ccontr, simp, erule disjE)
have "bn_inv p q (v, ns, ?ys)"
using C by simp
moreover have "add_inv (n - Suc (Suc m)) (v, ns, ?ys)"
using H and I by simp
ultimately have "add_inv (n - Suc (Suc m)) (?t r' v)"
by (rule round_add_inv [OF A])
hence "length xs' = foldl (+) 0 ns'"
using E by simp
moreover assume "ns' = []"
ultimately have "length xs' = 0" by simp
hence "k < Suc (Suc m)"
using N and P by simp
thus False
using T by simp
next
assume "xs' = []"
hence "k < Suc (Suc m)"
using N and P by simp
thus False
using T by simp
qed
hence V: "i - length ms' < length ns'"
using K by arith
hence W: "\<forall>j < offs ns' 0 ! (i - length ms').
\<forall>k \<in> {offs ns' 0 ! (i - length ms')..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
- using `?R` by simp
+ using \<open>?R\<close> by simp
from D [symmetric] have X: "foldl (+) 0 ms' = Suc (Suc m)"
proof (simp add: round_suc_suc_def Let_def add_replicate add_suc
split: if_split_asm, cases "m = 0")
case True
hence "?k = 0" by simp
hence "length ?ms = 0"
by (simp add: enum_length)
thus "foldl (+) 0 ?ms = m"
using True by simp
next
case False
hence "0 < ?k"
by (simp, drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
with A have "foldl (+) 0 ?ms = length ?zs"
by (rule enum_add, simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
thus "foldl (+) 0 ?ms = m"
using O by (simp add: mini_maxi_nths)
qed
have "length ms' < i"
proof (rule ccontr, simp add: not_less)
assume "i \<le> length ms'"
moreover have "length ms' < length (ms' @ ns')"
using U by simp
ultimately have "offs (ms' @ ns') 0 ! i \<le>
offs (ms' @ ns') 0 ! (length ms')"
by (rule offs_mono)
also have "\<dots> = Suc (Suc m)"
using U and X by (simp add: offs_append nth_append offs_length,
cases ns', simp_all)
finally have "offs (ms' @ ns') 0 ! i \<le> Suc (Suc m)" .
hence "j < Suc (Suc m)"
using L by simp
thus False
using S by simp
qed
hence Y: "offs (ms' @ ns') 0 ! i =
Suc (Suc m) + offs ns' 0 ! (i - length ms')"
using X by (simp add: offs_append nth_append offs_length,
subst offs_base_zero [OF V], simp)
hence "j < Suc (Suc m) + offs ns' 0 ! (i - length ms')"
using L by simp
moreover from this have "0 < offs ns' 0 ! (i - length ms')"
using S by (rule_tac ccontr, simp)
ultimately have "j - Suc (Suc m) < offs ns' 0 ! (i - length ms')"
by simp
hence Z: "\<forall>k \<in> {offs ns' 0 ! (i - length ms')..<length xs'}.
key (xs' ! (j - Suc (Suc m))) \<le> key (xs' ! k)"
using W by simp
have "offs ns' 0 ! (i - length ms') \<le> k - Suc (Suc m)"
using M and Y by simp
moreover have "k - Suc (Suc m) < length xs'"
using N and P and U by arith
ultimately have "k - Suc (Suc m) \<in>
{offs ns' 0 ! (i - length ms')..<length xs'}"
by simp
with Z show "key (xs' ! (j - Suc (Suc m))) \<le>
key (xs' ! (k - Suc (Suc m)))" ..
qed
qed
qed
lemma gcsort_sort_inv:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "add_inv n t" and
D: "n \<le> p"
shows "\<lbrakk>t' \<in> gcsort_set index key p t; sort_inv key t\<rbrakk> \<Longrightarrow>
sort_inv key t'"
by (erule gcsort_set.induct, simp, drule gcsort_add_inv [OF A _ C D],
rule round_sort_inv [OF A B], simp_all del: bn_inv.simps, erule conjE,
frule sym, erule subst, rule bn_inv_intro, insert D, simp)
text \<open>
\null
The only remaining task is to address step 10 of the proof method, which is done by proving theorem
@{text gcsort_sorted}. It holds under the conditions that the objects' number is not larger than the
counters' upper bound and function @{text index} satisfies both predicates @{const index_less} and
@{const index_mono}, and states that function @{const gcsort} is successful in sorting the input
objects' list.
\null
\<close>
theorem gcsort_sorted:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "length xs \<le> p"
shows "sorted (map key (gcsort index key p xs))"
proof -
let ?t = "(0, [length xs], xs)"
have "sort_inv key (gcsort_aux index key p ?t)"
by (rule gcsort_sort_inv [OF A B _ C], rule gcsort_add_input,
rule gcsort_aux_set, rule gcsort_sort_input)
moreover have "add_inv (length xs) (gcsort_aux index key p ?t)"
by (rule gcsort_add_inv [OF A _ _ C],
rule gcsort_aux_set, rule gcsort_add_input)
ultimately have "sorted (map key (gcsort_out (gcsort_aux index key p ?t)))"
by (rule gcsort_sort_intro, simp add: gcsort_sort_form)
moreover have "?t = gcsort_in xs"
by (simp add: gcsort_in_def)
ultimately show ?thesis
by (simp add: gcsort_def)
qed
end
diff --git a/thys/Goedel_Incompleteness/Abstract_First_Goedel_Rosser.thy b/thys/Goedel_Incompleteness/Abstract_First_Goedel_Rosser.thy
--- a/thys/Goedel_Incompleteness/Abstract_First_Goedel_Rosser.thy
+++ b/thys/Goedel_Incompleteness/Abstract_First_Goedel_Rosser.thy
@@ -1,348 +1,348 @@
chapter \<open>Abstract Formulations of Gödel-Rosser's First Incompleteness Theorem\<close>
(*<*)
theory Abstract_First_Goedel_Rosser
imports Rosser_Formula Standard_Model_More
begin
(*>*)
text \<open>The development here is very similar to that of Gödel First Incompleteness Theorem.
It lacks classical logical variants, since for them Rosser's trick does bring
any extra value.\<close>
section \<open>Proof-Theoretic Versions\<close>
context Rosser_Form_Proofs
begin
lemma NN_neg_unique_xx':
assumes [simp]:"\<phi> \<in> fmla" "Fvars \<phi> = {}"
shows
"bprv (imp (NN \<langle>\<phi>\<rangle> (Var xx'))
(eql \<langle>neg \<phi>\<rangle> (Var xx')))"
using B.prv_subst[of yy _ "Var xx'", OF _ _ _ NN_neg_unique[OF assms]] by fastforce
lemma NN_imp_xx':
assumes [simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}" "\<chi> \<in> fmla"
shows "bprv (imp (subst \<chi> \<langle>neg \<phi>\<rangle> xx')
(all xx' (imp (NN \<langle>\<phi>\<rangle> (Var xx')) \<chi>)))"
proof-
have 2: "bprv (imp (eql \<langle>neg \<phi>\<rangle> (Var xx')) (imp (subst \<chi> \<langle>neg \<phi>\<rangle> xx') \<chi>))"
using B.prv_eql_subst_trm[of xx' \<chi> "\<langle>neg \<phi>\<rangle>" "Var xx'", simplified] .
have 1: "bprv (imp (subst \<chi> \<langle>neg \<phi>\<rangle> xx') (imp (eql \<langle>neg \<phi>\<rangle> (Var xx')) \<chi>))"
by (simp add: "2" B.prv_imp_com)
have 0: "bprv (imp (subst \<chi> \<langle>neg \<phi>\<rangle> xx') (imp (NN \<langle>\<phi>\<rangle> (Var xx')) \<chi>))"
using 1
by (elim B.prv_prv_imp_trans[rotated 3])
(auto simp add: B.prv_imp_com B.prv_imp_monoR NN_neg_unique_xx')
show ?thesis by (rule B.prv_all_imp_gen) (auto simp: 0)
qed
lemma goedel_rosser_first_theEasyHalf:
assumes c: "consistent"
shows "\<not> prv \<phi>R"
proof
assume 0: "prv \<phi>R"
then obtain "prf" where [simp]: "prf \<in> proof" and "prfOf prf \<phi>R" using prv_prfOf by auto
hence 00: "bprv (PPf (encPf prf) \<langle>\<phi>R\<rangle>)" using prfOf_PPf by auto
from dwf_dwfd.d_dwf.bprv_prv'[OF _ 00, simplified] have b00: "prv (PPf (encPf prf) \<langle>\<phi>R\<rangle>)" .
have "\<not> prv (neg \<phi>R)" using 0 c unfolding consistent_def3 by auto
hence "\<forall> prf \<in> proof. \<not> prfOf prf (neg \<phi>R)" using 00 prv_prfOf by auto
hence "bprv (neg (PPf p \<langle>neg \<phi>R\<rangle>))" if "p \<in> num" for p using not_prfOf_PPf Clean_PPf_encPf that
by (cases "p \<in> encPf ` proof") auto
hence 1: "prv (all zz (imp (LLq (Var zz) (encPf prf)) (neg (PPf (Var zz) \<langle>neg \<phi>R\<rangle>))))"
(* here use locale assumption about the order-like relation: *)
by (intro LLq_num) auto
have 11: "prv (RR (encPf prf) \<langle>\<phi>R\<rangle>)"
using NN_imp_xx'[of \<phi>R "neg (PPf (Var zz) (Var xx'))", simplified]
by (auto simp add: RR_def2 R_def
intro!: prv_all_congW[OF _ _ _ _ 1] prv_imp_monoL[OF _ _ _ dwf_dwfd.d_dwf.bprv_prv'])
have 3: "prv (cnj (PPf (encPf prf) \<langle>\<phi>R\<rangle>) (RR (encPf prf) \<langle>\<phi>R\<rangle>))"
by (rule prv_cnjI[OF _ _ b00 11]) auto
have "prv ((PP' \<langle>\<phi>R\<rangle>))" unfolding PP'_def P'_def
using 3 by (auto intro: prv_exiI[of _ _ "encPf prf"])
moreover have "prv (neg (PP' \<langle>\<phi>R\<rangle>))"
using prv_eqv_prv[OF _ _ 0 prv_\<phi>R_eqv] by auto
ultimately show False using c unfolding consistent_def3 by auto
qed
lemma goedel_rosser_first_theHardHalf:
assumes c: "consistent"
shows "\<not> prv (neg \<phi>R)"
proof
assume 0: "prv (neg \<phi>R)"
then obtain "prf" where [simp,intro!]: "prf \<in> proof" and pr: "prfOf prf (neg \<phi>R)" using prv_prfOf by auto
define p where p: "p = encPf prf"
have [simp,intro!]: "p \<in> num" unfolding p by auto
have 11: "bprv (PPf p \<langle>neg \<phi>R\<rangle>)" using pr prfOf_PPf unfolding p by auto
have 1: "bprv (NN \<langle>\<phi>R\<rangle> \<langle>neg \<phi>R\<rangle>)" using NN_neg by simp
have "\<not> prv \<phi>R" using 0 c unfolding consistent_def3 by auto
from not_prv_prv_neg_PPf[OF _ _ this]
have b2: "\<forall> r \<in> num. bprv (neg (PPf r \<langle>\<phi>R\<rangle>))" by auto
hence 2: "\<forall> r \<in> num. prv (neg (PPf r \<langle>\<phi>R\<rangle>))"
by (auto intro: dwf_dwfd.d_dwf.bprv_prv')
obtain P where P[simp,intro!]: "P \<subseteq>num" "finite P"
and 3: "prv (dsj (sdsj {eql (Var yy) r |r. r \<in> P}) (LLq p (Var yy)))"
(* here use the other locale assumption about the order-like relation: *)
using LLq_num2 by auto
have "prv (imp (cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>)) fls)"
proof(rule prv_dsj_cases[OF _ _ _ 3])
{fix r assume r: "r \<in> P" hence rn[simp]: "r \<in> num" using P(1) by blast
have "prv (imp (cnj (PPf r \<langle>\<phi>R\<rangle>) (RR r \<langle>\<phi>R\<rangle>)) fls)"
using 2 unfolding neg_def
by (metis FvarsT_num PPf RR rn \<phi>R all_not_in_conv cnj enc fls imp in_num prv_imp_cnj3L prv_imp_mp)
hence "prv (imp (eql (Var yy) r)
(imp (cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>)) fls))"
using prv_eql_subst_trm_id[of yy "cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>)" r,simplified]
unfolding neg_def[symmetric]
by (intro prv_neg_imp_imp_trans) auto
}
thus "prv (imp (sdsj {eql (Var yy) r |r. r \<in> P})
(imp (cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>)) fls))"
using Var P(1) eql by (intro prv_sdsj_imp) (auto 0 0 simp: set_rev_mp)
next
let ?\<phi> = "all xx' (imp (NN \<langle>\<phi>R\<rangle> (Var xx')) (neg (PPf p (Var xx'))))"
have "bprv (neg ?\<phi>)"
using 1 11 by (intro B.prv_imp_neg_allWI[where t = "\<langle>neg \<phi>R\<rangle>"]) (auto intro: B.prv_prv_neg_imp_neg)
hence "prv (neg ?\<phi>)" by (auto intro: dwf_dwfd.d_dwf.bprv_prv')
hence 00: "prv (imp (LLq p (Var yy))
(imp (imp (LLq p (Var yy)) ?\<phi>) fls))"
unfolding neg_def[symmetric] by (intro prv_imp_neg_imp_neg_imp) auto
have "prv (imp (LLq p (Var yy))
(imp (RR (Var yy) \<langle>\<phi>R\<rangle>) fls))"
unfolding neg_def[symmetric] using 00[folded neg_def]
by (auto simp add: RR_def2 R_def intro!: prv_imp_neg_allI[where t = p])
thus "prv (imp (LLq p (Var yy))
(imp (cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>)) fls))"
unfolding neg_def[symmetric] by (intro prv_imp_neg_imp_cnjR) auto
qed(auto, insert Var P(1) eql, simp_all add: set_rev_mp)
hence "prv (neg (exi yy (cnj (PPf (Var yy) \<langle>\<phi>R\<rangle>) (RR (Var yy) \<langle>\<phi>R\<rangle>))))"
unfolding neg_def[symmetric] by (intro prv_neg_neg_exi) auto
hence "prv (neg (PP' \<langle>\<phi>R\<rangle>))" unfolding PP'_def P'_def by simp
hence "prv \<phi>R" using prv_\<phi>R_eqv by (meson PP' \<phi>R enc in_num neg prv_eqv_prv_rev)
- with `\<not> prv \<phi>R` show False using c unfolding consistent_def3 by auto
+ with \<open>\<not> prv \<phi>R\<close> show False using c unfolding consistent_def3 by auto
qed
theorem goedel_rosser_first:
assumes "consistent"
shows "\<not> prv \<phi>R \<and> \<not> prv (neg \<phi>R)"
using assms goedel_rosser_first_theEasyHalf goedel_rosser_first_theHardHalf by blast
theorem goedel_rosser_first_ex:
assumes "consistent"
shows "\<exists> \<phi>. \<phi> \<in> fmla \<and> \<not> prv \<phi> \<and> \<not> prv (neg \<phi>)"
using assms goedel_rosser_first by (intro exI[of _ \<phi>R]) blast
end \<comment> \<open>context @{locale Rosser_Form}\<close>
section \<open>Model-Theoretic Versions\<close>
subsection \<open>First model-theoretic version\<close>
locale Rosser_Form_Proofs_Minimal_Truth =
Rosser_Form_Proofs
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
fls
prv bprv
Lq
dsj
"proof" prfOf
enc
N S
encPf
Pf
+
Minimal_Truth_Soundness
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
bprv
isTrue
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and Lq
and prv bprv
and enc ("\<langle>_\<rangle>")
and N S P
and "proof" :: "'proof set" and prfOf encPf
and Pf
and isTrue
begin
lemma Fvars_PP'[simp]: "Fvars (PP' \<langle>\<phi>R\<rangle>) = {}" unfolding PP'_def
by (subst Fvars_subst) auto
lemma Fvars_RR'[simp]: "Fvars (RR (Var yy) \<langle>\<phi>R\<rangle>) = {yy}"
unfolding RR_def
by (subst Fvars_psubst) (fastforce intro!: exI[of _ "{yy}"])+
lemma isTrue_PPf_implies_\<phi>R:
assumes "isTrue (all yy (neg (PPf (Var yy) \<langle>\<phi>R\<rangle>)))"
(is "isTrue ?H")
shows "isTrue \<phi>R"
proof-
define F where "F \<equiv> RR (Var yy) \<langle>\<phi>R\<rangle>"
have [simp]: "F \<in> fmla" "Fvars F = {yy}"
unfolding F_def by auto
have [simp]: "exi yy (PPf (Var yy) \<langle>\<phi>R\<rangle>) \<in> fmla"
unfolding PPf_def by auto
have 1: "bprv
(imp (all yy (neg (PPf (Var yy) \<langle>\<phi>R\<rangle>)))
(neg (exi yy (PPf (Var yy) \<langle>\<phi>R\<rangle>))))"
(is "bprv (imp (all yy (neg ?G)) (neg (exi yy ?G)))")
using B.prv_all_neg_imp_neg_exi[of yy ?G] by auto
have 2: "bprv (imp (neg (exi yy ?G)) (neg (exi yy (cnj ?G F))))"
by (auto intro!: B.prv_imp_neg_rev B.prv_exi_cong B.prv_imp_cnjL)
have "bprv (imp (all yy (neg ?G)) (neg (exi yy (cnj ?G F))))"
using B.prv_prv_imp_trans[OF _ _ _ 1 2] by simp
hence "bprv (imp ?H (neg (PP' \<langle>\<phi>R\<rangle>)))"
unfolding PP'_def P'_def
by (simp add: F_def)
from B.prv_prv_imp_trans[OF _ _ _ this bprv_imp_\<phi>R]
have "bprv (imp ?H \<phi>R)" by auto
from prv_imp_implies_isTrue[OF _ _ _ _ this assms, simplified]
show ?thesis .
qed
theorem isTrue_\<phi>R:
assumes "consistent"
shows "isTrue \<phi>R"
proof-
have "\<forall> n \<in> num. bprv (neg (PPf n \<langle>\<phi>R\<rangle>))"
using not_prv_prv_neg_PPf[OF _ _ goedel_rosser_first_theEasyHalf[OF assms]]
by auto
hence "\<forall> n \<in> num. isTrue (neg (PPf n \<langle>\<phi>R\<rangle>))" by (auto intro: sound_isTrue)
hence "isTrue (all yy (neg (PPf (Var yy) \<langle>\<phi>R\<rangle>)))" by (auto intro: isTrue_all)
thus ?thesis using isTrue_PPf_implies_\<phi>R by auto
qed
theorem goedel_rosser_first_strong: "consistent \<Longrightarrow> \<not> prv \<phi>R \<and> \<not> prv (neg \<phi>R) \<and> isTrue \<phi>R"
using isTrue_\<phi>R goedel_rosser_first by blast
theorem goedel_rosser_first_strong_ex:
"consistent \<Longrightarrow> \<exists> \<phi>. \<phi> \<in> fmla \<and> \<not> prv \<phi> \<and> \<not> prv (neg \<phi>) \<and> isTrue \<phi>"
using goedel_rosser_first_strong by (intro exI[of _ \<phi>R]) blast
end \<comment> \<open>context @{locale Rosser_Form_Proofs_Minimal_Truth}\<close>
subsection \<open>Second model-theoretic version\<close>
context Rosser_Form
begin
print_context
end
locale Rosser_Form_Minimal_Truth_Soundness_HBL1iff_Compl_Pf =
Rosser_Form
var trm fmla Var
FvarsT substT Fvars subst
num
eql cnj imp all exi
fls
prv bprv
Lq
dsj
enc
N
S
P
+
Minimal_Truth_Soundness_HBL1iff_Compl_Pf
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
enc
prv bprv
P
isTrue
Pf
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and Lq
and enc ("\<langle>_\<rangle>")
and N S
and isTrue
and P Pf
locale Rosser_Form_Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf =
Rosser_Form_Minimal_Truth_Soundness_HBL1iff_Compl_Pf
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv bprv
Lq
enc
N S
isTrue
P
Pf
+
M : Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
enc
prv bprv
N
isTrue
Pf
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and Lq
and enc ("\<langle>_\<rangle>")
and N S P
and isTrue
and Pf
sublocale
Rosser_Form_Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf <
recover_proofs: Rosser_Form_Proofs_Minimal_Truth
where prfOf = prfOf and "proof" = "proof" and encPf = encPf
and prv = prv and bprv = bprv
by standard
context Rosser_Form_Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf
begin
thm recover_proofs.goedel_rosser_first_strong
end
(*<*)
end
(*>*)
diff --git a/thys/Goedel_Incompleteness/Deduction2.thy b/thys/Goedel_Incompleteness/Deduction2.thy
--- a/thys/Goedel_Incompleteness/Deduction2.thy
+++ b/thys/Goedel_Incompleteness/Deduction2.thy
@@ -1,280 +1,280 @@
chapter \<open>Deduction with Two Provability Relations\<close>
(*<*)
theory Deduction2 imports "Syntax_Independent_Logic.Deduction"
begin
(*>*)
text \<open>We work with two provability relations:
provability @{term prv} and basic provability @{term bprv}.\<close>
section \<open>From Deduction with One Provability Relation to Two\<close>
locale Deduct2 =
Deduct
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv
+
B: Deduct
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
bprv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and num
and eql cnj imp all exi
and prv bprv
+
assumes bprv_prv: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv \<phi> \<Longrightarrow> prv \<phi>"
begin
(* Removing the sentence (empty Fvars) hypothesis from bprv_prv *)
lemma bprv_prv':
assumes \<phi>: "\<phi> \<in> fmla" and b: "bprv \<phi>"
shows "prv \<phi>"
proof-
obtain V where V: "Fvars \<phi> = V" by blast
have VV: "V \<subseteq> var" using Fvars V \<phi> by blast
have f: "finite V" using V finite_Fvars[OF \<phi>] by auto
thus ?thesis using \<phi> b V VV
proof(induction V arbitrary: \<phi> rule: finite.induct)
case (emptyI \<phi>)
then show ?case by (simp add: bprv_prv)
next
case (insertI W v \<phi>)
show ?case proof(cases "v \<in> W")
case True
thus ?thesis
- using insertI.IH[OF `\<phi> \<in> fmla`] insertI.prems
+ using insertI.IH[OF \<open>\<phi> \<in> fmla\<close>] insertI.prems
by (simp add: insert_absorb)
next
case False
hence 1: "Fvars (all v \<phi>) = W"
using insertI.prems by auto
moreover have "bprv (all v \<phi>)"
using B.prv_all_gen insertI.prems by auto
ultimately have "prv (all v \<phi>)" using insertI by auto
thus ?thesis using allE_id insertI.prems by blast
qed
qed
qed
end \<comment> \<open>context @{locale Deduct2}\<close>
locale Deduct2_with_False =
Deduct_with_False
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
num
prv
+
B: Deduct_with_False
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
num
bprv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and num
and prv bprv
+
assumes bprv_prv: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv \<phi> \<Longrightarrow> prv \<phi>"
sublocale Deduct2_with_False < d_dwf: Deduct2
by standard (fact bprv_prv)
context Deduct2_with_False begin
lemma consistent_B_consistent: "consistent \<Longrightarrow> B.consistent"
using B.consistent_def bprv_prv consistent_def by blast
end \<comment> \<open>context @{locale Deduct2_with_False}\<close>
locale Deduct2_with_False_Disj =
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
+
B: Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
bprv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
+
assumes bprv_prv: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv \<phi> \<Longrightarrow> prv \<phi>"
sublocale Deduct2_with_False_Disj < dwf_dwfd: Deduct2_with_False
by standard (fact bprv_prv)
(* Factoring in a strict-order-like relation (not actually required to be an order): *)
locale Deduct2_with_PseudoOrder =
Deduct2_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv bprv
+
Syntax_PseudoOrder
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
Lq
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and Lq
+
assumes
(* We do not assume any ordering properties, but only these two axioms, Lq_num and Lq_num2,
which (interestingly) would be satisfied by both \<le> and < within a sufficiently strong
arithmetic such as Robinson's Q *)
(* For each formula \<phi>(z) and numeral q, if \<phi>(p) is provable for all p
then \<forall> z \<le> q. \<phi>(z) is provable.
(Note that a more natural property would assume \<phi>(p) is provable for all p\<le>q,
but we can get away with the stronger assumption (on the left of the implication). )
*)
Lq_num:
"let LLq = (\<lambda> t1 t2. psubst Lq [(t1,zz), (t2,yy)]) in
\<forall> \<phi> \<in> fmla. \<forall> q \<in> num. Fvars \<phi> = {zz} \<and> (\<forall> p \<in> num. bprv (subst \<phi> p zz))
\<longrightarrow> prv (all zz (imp (LLq (Var zz) q) \<phi>))"
and
(* For each numeral p, there exists a finite set P such that it is provable that
\<forall> y. (\<Or>p\<in>P. x = p) \<or> y \<le> p
(where we write \<le> instead of Lq, but could also think of <):
*)
Lq_num2:
"let LLq = (\<lambda> t1 t2. psubst Lq [(t1,zz), (t2,yy)]) in
\<forall> p \<in> num. \<exists> P \<subseteq> num. finite P \<and> prv (dsj (sdsj {eql (Var yy) r | r. r \<in> P}) (LLq p (Var yy)))"
begin
lemma LLq_num:
assumes "\<phi> \<in> fmla" "q \<in> num" "Fvars \<phi> = {zz}" "\<forall> p \<in> num. bprv (subst \<phi> p zz)"
shows "prv (all zz (imp (LLq (Var zz) q) \<phi>))"
using assms Lq_num unfolding LLq_def by auto
lemma LLq_num2:
assumes "p \<in> num"
shows "\<exists> P \<subseteq> num. finite P \<and> prv (dsj (sdsj {eql (Var yy) r | r. r \<in> P}) (LLq p (Var yy)))"
using assms Lq_num2 unfolding LLq_def by auto
end \<comment> \<open>context @{locale Deduct2_with_PseudoOrder}\<close>
section \<open>Factoring In Explicit Proofs\<close>
locale Deduct_with_Proofs =
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
+
fixes
"proof" :: "'proof set"
and
prfOf :: "'proof \<Rightarrow> 'fmla \<Rightarrow> bool"
assumes
\<comment> \<open>Provability means there exists a proof (only needed for sentences):\<close>
prv_prfOf: "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> prv \<phi> \<longleftrightarrow> (\<exists> prf \<in> proof. prfOf prf \<phi>)"
(* We consider proof structure only for prv, not for bprv *)
locale Deduct2_with_Proofs =
Deduct2_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv bprv
+
Deduct_with_Proofs
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
"proof" prfOf
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and "proof" :: "'proof set" and prfOf
locale Deduct2_with_Proofs_PseudoOrder =
Deduct2_with_Proofs
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv bprv
"proof" prfOf
+
Deduct2_with_PseudoOrder
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv bprv
Lq
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and "proof" :: "'proof set" and prfOf
and Lq
(*<*)
end
(*>*)
\ No newline at end of file
diff --git a/thys/Goedel_Incompleteness/Jeroslow_Original.thy b/thys/Goedel_Incompleteness/Jeroslow_Original.thy
--- a/thys/Goedel_Incompleteness/Jeroslow_Original.thy
+++ b/thys/Goedel_Incompleteness/Jeroslow_Original.thy
@@ -1,337 +1,337 @@
section \<open>A Formalization of Jeroslow's Original Argument\<close>
(*<*)
theory Jeroslow_Original imports
"Syntax_Independent_Logic.Pseudo_Term"
Abstract_Jeroslow_Encoding
begin
(*>*)
subsection \<open>Preliminaries\<close>
text \<open>The First Derivability Condition was stated using a formula
with free variable @{term xx}, whereas the pseudo-term theory employs a different variable,
inp. The distinction is of course immaterial, because we can perform
a change of variable in the instantiation:\<close>
context HBL1
begin
text \<open>Changing the variable (from @{term xx} to @{term inp}) in the provability predicate:\<close>
definition "Pinp \<equiv> subst P (Var inp) xx"
lemma PP_Pinp: "t \<in> trm \<Longrightarrow> PP t = instInp Pinp t"
unfolding PP_def Pinp_def instInp_def by auto
text \<open>A version of HBL1 that uses the @{term inp} variable:\<close>
lemma HBL1_inp:
"\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> prv \<phi> \<Longrightarrow> prv (instInp Pinp \<langle>\<phi>\<rangle>)"
unfolding Pinp_def instInp_def by (auto intro: HBL1)
end \<comment> \<open>context @{locale HBL1 }\<close>
subsection \<open>Jeroslow-style diagonalization\<close>
locale Jeroslow_Diagonalization =
Deduct_with_False_Disj_Rename
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
+
Encode
var trm fmla Var FvarsT substT Fvars subst
num
enc
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
and enc ("\<langle>_\<rangle>")
+
fixes F :: "('trm \<Rightarrow> 'trm) set"
and encF :: "('trm \<Rightarrow> 'trm) \<Rightarrow> 'fmla"
and N :: "'trm \<Rightarrow> 'trm"
and ssap :: "'fmla \<Rightarrow> 'trm \<Rightarrow> 'trm"
assumes
\<comment> \<open>For the members @{term f} of @{term F}, we will only care about their action on numerals,
and we assume that they send numerals to numerals.\<close>
F[simp,intro!]: "\<And> f n. f \<in> F \<Longrightarrow> n \<in> num \<Longrightarrow> f n \<in> num"
and
encF[simp,intro!]: "\<And> f. f \<in> F \<Longrightarrow> encF f \<in> ptrm (Suc 0)"
and
N[simp,intro!]: "N \<in> F"
and
ssap[simp]: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {inp} \<Longrightarrow> ssap \<phi> \<in> F"
and
ReprF: "\<And>f n. f \<in> F \<Longrightarrow> n \<in> num \<Longrightarrow> prveqlPT (instInp (encF f) n) (f n)"
and
CapN: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> N \<langle>\<phi>\<rangle> = \<langle>neg \<phi>\<rangle>"
and
CapSS: \<comment> \<open>We consider formulas @{term \<psi>} of one variable, called @{term inp}:\<close>
"\<And> \<psi> f. \<psi> \<in> fmla \<Longrightarrow> Fvars \<psi> = {inp} \<Longrightarrow> f \<in> F \<Longrightarrow>
ssap \<psi> \<langle>encF f\<rangle> = \<langle>instInpP \<psi> 0 (instInp (encF f) \<langle>encF f\<rangle>)\<rangle>"
begin
lemma encF_fmla[simp,intro!]: "\<And> f. f \<in> F \<Longrightarrow> encF f \<in> fmla"
by auto
lemma enc_trm: "\<phi> \<in> fmla \<Longrightarrow> \<langle>\<phi>\<rangle> \<in> trm"
by auto
definition \<tau>J :: "'fmla \<Rightarrow> 'fmla" where
"\<tau>J \<psi> \<equiv> instInp (encF (ssap \<psi>)) (\<langle>encF (ssap \<psi>)\<rangle>)"
definition \<phi>J :: "'fmla \<Rightarrow> 'fmla" where
"\<phi>J \<psi> \<equiv> instInpP \<psi> 0 (\<tau>J \<psi>)"
lemma \<tau>J[simp]:
assumes "\<psi> \<in> fmla" and "Fvars \<psi> = {inp}"
shows "\<tau>J \<psi> \<in> ptrm 0"
unfolding \<tau>J_def apply(rule instInp)
using assms by auto
lemma \<tau>J_fmla[simp]:
assumes "\<psi> \<in> fmla" and "Fvars \<psi> = {inp}"
shows "\<tau>J \<psi> \<in> fmla"
using \<tau>J[OF assms] unfolding ptrm_def by auto
lemma FvarsT_\<tau>J[simp]:
assumes "\<psi> \<in> fmla" and "Fvars \<psi> = {inp}"
shows "Fvars (\<tau>J \<psi>) = {out}"
using \<tau>J[OF assms] unfolding ptrm_def by auto
lemma \<phi>J[simp]:
assumes "\<psi> \<in> fmla" and "Fvars \<psi> = {inp}"
shows "\<phi>J \<psi> \<in> fmla"
unfolding \<phi>J_def using assms by (intro instInpP_fmla) auto
lemma Fvars_\<phi>J[simp]:
assumes "\<psi> \<in> fmla" and "Fvars \<psi> = {inp}"
shows "Fvars (\<phi>J \<psi>) = {}"
using assms unfolding \<phi>J_def by auto
lemma diagonalization:
assumes \<psi>[simp]: "\<psi> \<in> fmla" and [simp]: "Fvars \<psi> = {inp}"
shows "prveqlPT (\<tau>J \<psi>) \<langle>instInpP \<psi> 0 (\<tau>J \<psi>)\<rangle> \<and>
prv (eqv (\<phi>J \<psi>) (instInp \<psi> \<langle>\<phi>J \<psi>\<rangle>))"
proof
define f where "f \<equiv> ssap \<psi>"
have f[simp]: "f \<in> F" unfolding f_def using assms by auto
have ff: "f \<langle>encF f\<rangle> = \<langle>instInpP \<psi> 0 (\<tau>J \<psi>)\<rangle>"
using assms unfolding f_def \<tau>J_def by (intro CapSS) auto
show "prveqlPT (\<tau>J \<psi>) \<langle>instInpP \<psi> 0 (\<tau>J \<psi>)\<rangle>"
using ReprF[OF f, of "\<langle>encF f\<rangle>"]
unfolding \<tau>J_def[of \<psi>, unfolded f_def[symmetric],symmetric] ff[symmetric]
by auto
from prveqlPT_prv_instInp_eqv_instInpP[OF \<psi>, of "\<tau>J \<psi>", OF _ _ _ _ this,
unfolded \<phi>J_def[symmetric]]
show "prv (eqv (\<phi>J \<psi>) (instInp \<psi> \<langle>\<phi>J \<psi>\<rangle>))"
by auto
qed
end \<comment> \<open>context @{locale Jeroslow_Diagonalization}\<close>
subsection \<open>Jeroslow's Second Incompleteness Theorem\<close>
text \<open>We follow Jeroslow's pseudo-term-based development of the
Second Incompleteness Theorem and point out the location in the proof that
implicitly uses an unstated assumption: the fact that, for certain two provably
equivalent formulas @{term \<phi>} and @{term \<phi>'}, it is provable that
the provability of the encoding of @{term \<phi>'} implies
the provability of the encoding of @{term \<phi>}. \<close>
locale Jeroslow_Godel_Second =
Jeroslow_Diagonalization
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
enc
F encF N ssap
+
HBL1
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv prv
enc
P
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
and enc ("\<langle>_\<rangle>")
and P
and F encF N ssap
+
assumes
SHBL3: "\<And>\<tau>. \<tau> \<in> ptrm 0 \<Longrightarrow> prv (imp (instInpP Pinp 0 \<tau>) (instInp Pinp \<langle>instInpP Pinp 0 \<tau>\<rangle>))"
begin
text \<open>Consistency formula a la Jeroslow:\<close>
definition jcons :: "'fmla" where
"jcons \<equiv> all xx (neg (cnj (instInp Pinp (Var xx))
(instInpP Pinp 0 (instInp (encF N) (Var (xx))))))"
lemma prv_eql_subst_trm3:
"x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow>
prv (eql t1 t2) \<Longrightarrow> prv (subst \<phi> t1 x) \<Longrightarrow> prv (subst \<phi> t2 x)"
using prv_eql_subst_trm2
by (meson subst prv_imp_mp)
lemma Pinp[simp,intro!]: "Pinp \<in> fmla"
and Fvars_Pinp[simp]: "Fvars Pinp = {inp}"
unfolding Pinp_def by auto
lemma ReprF_combineWith_CapN:
assumes "\<phi> \<in> fmla" and "Fvars \<phi> = {}"
shows "prveqlPT (instInp (encF N) \<langle>\<phi>\<rangle>) \<langle>neg \<phi>\<rangle>"
using assms unfolding CapN[symmetric, OF assms] by (intro ReprF) auto
theorem jeroslow_godel_second:
assumes consistent
\<comment> \<open>Assumption that is not stated by Jeroslow, but seems to be needed:\<close>
assumes unstated:
"let \<psi> = instInpP Pinp (Suc 0) (encF N);
\<tau> = \<tau>J \<psi>;
\<phi> = instInpP (instInpP Pinp (Suc 0) (encF N)) 0 \<tau>;
\<phi>' = instInpP Pinp 0 (instInpP (encF N) 0 \<tau>)
in prv (imp (instInp Pinp \<langle>\<phi>'\<rangle>) (instInp Pinp \<langle>\<phi>\<rangle>))"
shows "\<not> prv jcons"
proof
assume *: "prv jcons"
define \<psi> where "\<psi> \<equiv> instInpP Pinp (Suc 0) (encF N)"
define \<tau> where "\<tau> \<equiv> \<tau>J \<psi>"
define \<phi> where "\<phi> \<equiv> \<phi>J \<psi>"
have \<psi>[simp,intro]: "\<psi> \<in> fmla" "Fvars \<psi> = {inp}"
unfolding \<psi>_def by auto
have \<tau>[simp,intro]: "\<tau> \<in> ptrm 0" "\<tau> \<in> fmla" "Fvars \<tau> = {out}"
unfolding \<tau>_def by auto
have [simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}" unfolding \<phi>_def by auto
define eN\<tau> where "eN\<tau> \<equiv> instInpP (encF N) 0 \<tau>"
have eN\<tau>[simp]: "eN\<tau> \<in> ptrm 0" "eN\<tau> \<in> fmla" "Fvars eN\<tau> = {out}"
unfolding eN\<tau>_def by auto
define \<phi>' where "\<phi>' \<equiv> instInpP Pinp 0 eN\<tau>"
have [simp]: "\<phi>' \<in> fmla" "Fvars \<phi>' = {}" unfolding \<phi>'_def by auto
have \<phi>\<phi>': "prv (imp \<phi> \<phi>')" and \<phi>'\<phi>: "prv (imp \<phi>' \<phi>)" and \<phi>e\<phi>': "prv (eqv \<phi> \<phi>')"
unfolding \<phi>_def \<phi>J_def \<phi>'_def eN\<tau>_def \<tau>_def[symmetric] unfolding \<psi>_def
using prv_instInpP_compose[of Pinp "encF N" \<tau>] by auto
from diagonalization[OF \<psi>]
have "prveqlPT \<tau> \<langle>instInpP \<psi> 0 \<tau>\<rangle>" and **: "prv (eqv \<phi> (instInp \<psi> \<langle>\<phi>\<rangle>))"
unfolding \<tau>_def[symmetric] \<phi>_def[symmetric] by auto
have "**1": "prv (imp \<phi> (instInp \<psi> \<langle>\<phi>\<rangle>))" "prv (imp (instInp \<psi> \<langle>\<phi>\<rangle>) \<phi>)"
using prv_imp_eqvEL[OF _ _ **] prv_imp_eqvER[OF _ _ **] by auto
from SHBL3[OF eN\<tau>(1)]
have "prv (imp (instInpP Pinp 0 eN\<tau>) (instInp Pinp \<langle>instInpP Pinp 0 eN\<tau>\<rangle>))" .
hence "prv (imp \<phi>' (instInp Pinp \<langle>\<phi>'\<rangle>))" unfolding \<phi>'_def .
from prv_prv_imp_trans[OF _ _ _ \<phi>\<phi>' this]
have 0: "prv (imp \<phi> (instInp Pinp \<langle>\<phi>'\<rangle>))" by auto
note unr = unstated[unfolded Let_def
\<phi>_def[unfolded \<phi>J_def \<tau>_def[symmetric], symmetric] \<psi>_def[symmetric]
\<tau>_def[symmetric] eN\<tau>_def[symmetric] \<phi>'_def[symmetric] \<phi>J_def]
have 1: "prv (imp \<phi> (instInp Pinp \<langle>\<phi>\<rangle>))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_addLemmaE[OF unr])
apply(nrule r: nprv_addImpLemmaE[OF 0])
apply(nrule r: nprv_clear3_3)
by (simp add: nprv_clear2_2 prv_nprv1I unr)
have 2: "prv (imp \<phi> (cnj (instInp Pinp \<langle>\<phi>\<rangle>)
(instInp \<psi> \<langle>\<phi>\<rangle>)))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_cnjI)
subgoal apply(nrule r: nprv_addImpLemmaE[OF 1]) .
subgoal apply(nrule r: nprv_addImpLemmaE[OF "**1"(1)]) . .
define z where "z \<equiv> Variable (Suc (Suc 0))"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<notin> Fvars Pinp"
"out \<noteq> z \<and> z \<noteq> out" "inp \<noteq> z \<and> z \<noteq> inp"
unfolding z_def by auto
have 30: "subst (instInpP Pinp 0 (instInp (encF N) (Var xx))) \<langle>\<phi>\<rangle> xx =
instInpP Pinp 0 (instInp (encF N) \<langle>\<phi>\<rangle>)"
unfolding z_def[symmetric] instInp_def instInpP_def Let_def
by (variousSubsts4 auto
s1: subst_compose_diff s2: subst_subst
s3: subst_notIn[of _ _ xx] s4: subst_compose_diff)
have 31: "subst (instInp Pinp (Var xx)) \<langle>\<phi>\<rangle> xx =
instInp Pinp \<langle>\<phi>\<rangle>" unfolding instInp_def by auto
have [simp]: "instInp (instInpP Pinp (Suc 0) (encF N)) \<langle>\<phi>\<rangle> =
instInpP Pinp 0 (instInp (encF N) \<langle>\<phi>\<rangle>)"
by (auto simp: instInp_instInpP \<psi>_def)
have 3: "prv (neg (cnj (instInp Pinp \<langle>\<phi>\<rangle>)
(instInp \<psi> \<langle>\<phi>\<rangle>)))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF *, unfolded jcons_def])
apply(rule nprv_allE0[of _ _ _ "\<langle>\<phi>\<rangle>"], auto)
unfolding 30 31
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_negI)
apply(nrule r: nprv_negE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_cnjE0)
apply(nrule r: nprv_clear3_3)
apply(nrule r: nprv_cnjI)
apply(nrule r: nprv_clear2_1)
unfolding \<psi>_def
apply(nrule r: nprv_hyp) .
have ***: "prv (neg \<phi>)"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_negI)
apply(nrule r: nprv_addImpLemmaE[OF 2])
apply(nrule r: nprv_addLemmaE[OF 3])
apply(nrule r: nprv_negE0) .
have 4: "prv (instInp Pinp \<langle>neg \<phi>\<rangle>)" using HBL1_inp[OF _ _ ***] by auto
have 5: "prveqlPT (instInp (encF N) \<langle>\<phi>\<rangle>) \<langle>neg \<phi>\<rangle>"
using ReprF_combineWith_CapN[of \<phi>] by auto
have [simp]: "instInp (encF N) \<langle>\<phi>\<rangle> \<in> ptrm 0" using instInp by auto
have 6: "prv (instInpP Pinp 0 (instInp (encF N) \<langle>\<phi>\<rangle>))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF 4])
apply(nrule r: prveqlPT_nprv_instInpP_instInp[OF _ _ _ _ _ 5]) .
note lem = "**1"(2)[unfolded \<psi>_def]
have "prv \<phi>"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF 6])
apply(nrule r: nprv_addImpLemmaE[OF lem]) .
- from this *** `consistent` show False unfolding consistent_def3 by auto
+ from this *** \<open>consistent\<close> show False unfolding consistent_def3 by auto
qed
end \<comment> \<open>context @{locale Jeroslow_Godel_Second}\<close>
(*<*)
end
(*>*)
diff --git a/thys/Goedel_Incompleteness/Standard_Model_More.thy b/thys/Goedel_Incompleteness/Standard_Model_More.thy
--- a/thys/Goedel_Incompleteness/Standard_Model_More.thy
+++ b/thys/Goedel_Incompleteness/Standard_Model_More.thy
@@ -1,469 +1,469 @@
chapter \<open>Standard Models with Two Provability Relations\<close>
(*<*)
(* Abstract notion of standard model and truth *)
theory Standard_Model_More
imports Derivability_Conditions "Syntax_Independent_Logic.Standard_Model"
begin
(*>*)
locale Minimal_Truth_Soundness_Proof_Repr =
CleanRepr_Proofs
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv bprv
enc
fls
dsj
"proof" prfOf
encPf
Pf
+ \<comment> \<open>The label "B" stands for "basic", as a reminder that soundness
refers to the basic provability relation:\<close>
B: Minimal_Truth_Soundness
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
bprv
isTrue
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv bprv
and isTrue
and enc ("\<langle>_\<rangle>")
and "proof" :: "'proof set" and prfOf
and encPf Pf
begin
lemmas prfOf_iff_PPf = B_consistent_prfOf_iff_PPf[OF B.consistent]
text \<open>The provability predicate is decided by basic provability on encodings:\<close>
lemma isTrue_prv_PPf_prf_or_neg:
"prf \<in> proof \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
bprv (PPf (encPf prf) \<langle>\<phi>\<rangle>) \<or> bprv (neg (PPf (encPf prf) \<langle>\<phi>\<rangle>))"
using not_prfOf_PPf prfOf_PPf by blast
text \<open>... hence that predicate is complete w.r.t. truth:\<close>
lemma isTrue_PPf_Implies_bprv_PPf:
"prf \<in> proof \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
isTrue (PPf (encPf prf) \<langle>\<phi>\<rangle>) \<Longrightarrow> bprv (PPf (encPf prf) \<langle>\<phi>\<rangle>)"
by (metis FvarsT_num Fvars_PPf Fvars_fls PPf
Un_empty empty_iff enc encPf fls in_num isTrue_prv_PPf_prf_or_neg
neg_def B.not_isTrue_fls B.prv_imp_implies_isTrue)
text \<open>... and thanks cleanness we can replace encoded proofs
with arbitrary numerals in the completeness property:\<close>
lemma isTrue_implies_bprv_PPf:
assumes [simp]: "n \<in> num" "\<phi> \<in> fmla" "Fvars \<phi> = {}"
and iT: "isTrue (PPf n \<langle>\<phi>\<rangle>)"
shows "bprv (PPf n \<langle>\<phi>\<rangle>)"
proof(cases "n \<in> encPf ` proof")
case True
thus ?thesis
using iT isTrue_PPf_Implies_bprv_PPf by auto
next
case False
hence "bprv (neg (PPf n \<langle>\<phi>\<rangle>))" by (simp add: Clean_PPf_encPf)
hence "isTrue (neg (PPf n \<langle>\<phi>\<rangle>))" by (intro B.sound_isTrue) auto
hence False using iT by (intro B.isTrue_neg_excl) auto
thus ?thesis by auto
qed
text \<open>... in fact, by @{locale Minimal_Truth_Soundness} we even have an iff:\<close>
lemma isTrue_iff_bprv_PPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> isTrue (PPf n \<langle>\<phi>\<rangle>) \<longleftrightarrow> bprv (PPf n \<langle>\<phi>\<rangle>)"
using isTrue_implies_bprv_PPf
using exists_no_Fvars B.not_isTrue_fls B.sound_isTrue by auto
text \<open>Truth of the provability representation implies provability (TIP):\<close>
lemma TIP:
assumes \<phi>[simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}"
and iPP: "isTrue (wrepr.PP \<langle>\<phi>\<rangle>)"
shows "prv \<phi>"
proof-
have "isTrue (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))" using iPP unfolding PP_PPf[OF \<phi>(1)] .
from B.isTrue_exi[OF _ _ _ this]
obtain n where n[simp]: "n \<in> num" and "isTrue (PPf n \<langle>\<phi>\<rangle>)" by auto
hence pP: "bprv (PPf n \<langle>\<phi>\<rangle>)" using isTrue_implies_bprv_PPf by auto
hence "\<not> bprv (neg (PPf n \<langle>\<phi>\<rangle>))"
using B.prv_neg_excl[of "PPf n \<langle>\<phi>\<rangle>"] by auto
then obtain "prf" where "prf"[simp]: "prf \<in> proof" and nn: "n = encPf prf"
using assms n Clean_PPf_encPf \<phi>(1) by blast
have "prfOf prf \<phi>" using pP unfolding nn using prfOf_iff_PPf by auto
thus ?thesis using prv_prfOf by auto
qed
text \<open>The reverse HBL1 -- now without the $\omega$-consistency assumption which holds here
thanks to our truth-in-standard-model assumption:\<close>
lemmas HBL1_rev = \<omega>consistentStd1_HBL1_rev[OF B.\<omega>consistentStd1]
text \<open>Note that the above would also follow by @{locale Minimal_Truth_Soundness} from TIP:\<close>
lemma TIP_implies_HBL1_rev:
assumes TIP: "\<forall>\<phi>. \<phi> \<in> fmla \<and> Fvars \<phi> = {} \<and> isTrue (wrepr.PP \<langle>\<phi>\<rangle>) \<longrightarrow> prv \<phi>"
shows "\<forall>\<phi>. \<phi> \<in> fmla \<and> Fvars \<phi> = {} \<and> bprv (wrepr.PP \<langle>\<phi>\<rangle>) \<longrightarrow> prv \<phi>"
using B.sound_isTrue[of "wrepr.PP \<langle>_\<rangle>"] TIP by auto
end \<comment> \<open>context @{locale Minimal_Truth_Soundness_Proof_Repr}\<close>
section \<open>Proof recovery from $\mathit{HBL1\_iff}$\<close>
locale Minimal_Truth_Soundness_HBL1iff_Compl_Pf =
HBL1
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv bprv
enc
P
+
B : Minimal_Truth_Soundness
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
bprv
isTrue
+
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and enc ("\<langle>_\<rangle>")
and prv bprv
and P
and isTrue
+
fixes Pf :: 'fmla
assumes
\<comment> \<open>@{term Pf} is a formula with free variables @{term xx} @{term yy}:\<close>
Pf[simp,intro!]: "Pf \<in> fmla"
and
Fvars_Pf[simp]: "Fvars Pf = {yy,xx}"
and
\<comment> \<open>@{term P} relates to @{term Pf} internally (inside basic provability)
just like a @{term prv} and a @{term prfOf} would relate---via an existential:\<close>
P_Pf:
"\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
let PPf = (\<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]) in
bprv (eqv (subst P \<langle>\<phi>\<rangle> xx) (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)))"
assumes
\<comment> \<open>We assume both $\mathit{HBL1}$ and $\mathit{HBL1\_rev}$, i.e., an iff version:\<close>
HBL1_iff: "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv (PP \<langle>\<phi>\<rangle>) \<longleftrightarrow> prv \<phi>"
and
Compl_Pf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
let PPf = (\<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]) in
isTrue (PPf n \<langle>\<phi>\<rangle>) \<longrightarrow> bprv (PPf n \<langle>\<phi>\<rangle>)"
begin
definition PPf where "PPf \<equiv> \<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]"
lemma PP_deff: "PP t = subst P t xx" using PP_def by auto
lemma PP_PPf_eqv:
"\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv (eqv (PP \<langle>\<phi>\<rangle>) (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)))"
using PP_deff PPf_def P_Pf by auto
(* *)
lemma PPf[simp,intro!]: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> xx \<notin> FvarsT t1 \<Longrightarrow> PPf t1 t2 \<in> fmla"
unfolding PPf_def by auto
lemma PPf_def2: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> xx \<notin> FvarsT t1 \<Longrightarrow>
PPf t1 t2 = subst (subst Pf t1 yy) t2 xx"
unfolding PPf_def by (rule psubst_eq_rawpsubst2[simplified]) auto
lemma Fvars_PPf[simp]:
"t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> xx \<notin> FvarsT t1 \<Longrightarrow> Fvars (PPf t1 t2) = FvarsT t1 \<union> FvarsT t2"
by (auto simp add: PPf_def2 subst2_fresh_switch)
lemma [simp]:
"n \<in> num \<Longrightarrow> subst (PPf (Var yy) (Var xx)) n xx = PPf (Var yy) n"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (PPf (Var yy) m) n yy = PPf n m"
"n \<in> num \<Longrightarrow> subst (PPf (Var yy) (Var xx)) n yy = PPf n (Var xx)"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (PPf m (Var xx)) n xx = PPf m n"
"m \<in> num \<Longrightarrow> subst (PPf (Var zz) (Var xx')) m zz = PPf m (Var xx')"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (PPf m (Var xx')) n xx' = PPf m n"
"n \<in> num \<Longrightarrow> subst (PPf (Var zz) (Var xx')) n xx' = PPf (Var zz) n"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (PPf (Var zz) n) m zz = PPf m n"
by (auto simp: PPf_def2 subst2_fresh_switch)
(* *)
lemma PP_PPf:
assumes "\<phi> \<in> fmla" "Fvars \<phi> = {}" shows "bprv (PP \<langle>\<phi>\<rangle>) \<longleftrightarrow> bprv (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))"
using assms PP_PPf_eqv[OF assms] B.prv_eqv_sym[OF _ _ PP_PPf_eqv[OF assms]]
by (auto intro!: B.prv_eqv_prv[of "PP \<langle>\<phi>\<rangle>" "exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)"]
B.prv_eqv_prv[of "exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)" "PP \<langle>\<phi>\<rangle>"])
lemma isTrue_implies_bprv_PPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
isTrue (PPf n \<langle>\<phi>\<rangle>) \<Longrightarrow> bprv (PPf n \<langle>\<phi>\<rangle>)"
using Compl_Pf by(simp add: PPf_def)
lemma isTrue_iff_bprv_PPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> isTrue (PPf n \<langle>\<phi>\<rangle>) \<longleftrightarrow> bprv (PPf n \<langle>\<phi>\<rangle>)"
using isTrue_implies_bprv_PPf
using exists_no_Fvars B.not_isTrue_fls B.sound_isTrue by auto
text \<open>Preparing to instantiate this "proof recovery" alternative into our
mainstream locale hierarchy, which assumes proofs.
We define the "missing" proofs to be numerals, we encode them as the identity,
and we "copy" @{term prfOf} from the corresponding predicate one-level-up, @{term PPf}:\<close>
definition "proof" :: "'trm set" where [simp]: "proof = num"
definition prfOf :: "'trm \<Rightarrow> 'fmla \<Rightarrow> bool" where
"prfOf n \<phi> \<equiv> bprv (PPf n \<langle>\<phi>\<rangle>)"
definition encPf :: "'trm \<Rightarrow> 'trm" where [simp]: "encPf \<equiv> id"
(* *)
lemma prv_exi_PPf_iff_isTrue:
assumes [simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}"
shows "bprv (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)) \<longleftrightarrow> isTrue (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L thus ?R by (intro B.sound_isTrue) auto
next
assume ?R
- obtain n where n[simp]: "n \<in> num" and i: "isTrue (PPf n \<langle>\<phi>\<rangle>)" using B.isTrue_exi[OF _ _ _ `?R`] by auto
+ obtain n where n[simp]: "n \<in> num" and i: "isTrue (PPf n \<langle>\<phi>\<rangle>)" using B.isTrue_exi[OF _ _ _ \<open>?R\<close>] by auto
hence "bprv (PPf n \<langle>\<phi>\<rangle>)" by (auto simp: isTrue_iff_bprv_PPf)
thus ?L by (intro B.prv_exiI[of _ _ n]) auto
qed
lemma isTrue_exi_iff:
assumes [simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}"
shows "isTrue (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)) \<longleftrightarrow> (\<exists>n\<in>num. isTrue (PPf n \<langle>\<phi>\<rangle>))" (is "?L \<longleftrightarrow> ?R")
proof
- assume ?L thus ?R using B.isTrue_exi[OF _ _ _ `?L`] by auto
+ assume ?L thus ?R using B.isTrue_exi[OF _ _ _ \<open>?L\<close>] by auto
next
assume ?R
then obtain n where n[simp]: "n \<in> num" and i: "isTrue (PPf n \<langle>\<phi>\<rangle>)" by auto
hence "bprv (PPf n \<langle>\<phi>\<rangle>)" by (auto simp: isTrue_iff_bprv_PPf)
hence "bprv (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))" by (intro B.prv_exiI[of _ _ n]) auto
thus ?L by (intro B.sound_isTrue) auto
qed
lemma prv_prfOf:
assumes "\<phi> \<in> fmla" "Fvars \<phi> = {}"
shows "prv \<phi> \<longleftrightarrow> (\<exists>n\<in>num. prfOf n \<phi>)"
proof-
have "prv \<phi> \<longleftrightarrow> bprv (PP \<langle>\<phi>\<rangle>)" using HBL1_iff[OF assms] by simp
also have "\<dots> \<longleftrightarrow> bprv (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))" unfolding PP_PPf[OF assms] ..
also have "\<dots> \<longleftrightarrow> isTrue (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))" using prv_exi_PPf_iff_isTrue[OF assms] .
also have "\<dots> \<longleftrightarrow> (\<exists>n\<in>num. isTrue (PPf n \<langle>\<phi>\<rangle>))" using isTrue_exi_iff[OF assms] .
also have "\<dots> \<longleftrightarrow> (\<exists>n\<in>num. bprv (PPf n \<langle>\<phi>\<rangle>))" using isTrue_iff_bprv_PPf[OF _ assms] by auto
also have "\<dots> \<longleftrightarrow> (\<exists>n\<in>num. prfOf n \<phi>)" unfolding prfOf_def ..
finally show ?thesis .
qed
lemma prfOf_prv_Pf:
assumes "n \<in> num" and "\<phi> \<in> fmla" "Fvars \<phi> = {}" and "prfOf n \<phi>"
shows "bprv (psubst Pf [(n, yy), (\<langle>\<phi>\<rangle>, xx)])"
using assms unfolding prfOf_def by (auto simp: PPf_def2 psubst_eq_rawpsubst2)
lemma isTrue_exi_iff_PP:
assumes [simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {}"
shows "isTrue (PP \<langle>\<phi>\<rangle>) \<longleftrightarrow> (\<exists>n\<in>num. isTrue (PPf n \<langle>\<phi>\<rangle>))"
proof-
have "bprv (eqv (PP \<langle>\<phi>\<rangle>) (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)))"
using PP_PPf_eqv by auto
hence "bprv (imp (PP \<langle>\<phi>\<rangle>) (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)))"
and "bprv (imp (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)) (PP \<langle>\<phi>\<rangle>))"
by (simp_all add: B.prv_imp_eqvEL B.prv_imp_eqvER)
thus ?thesis unfolding isTrue_exi_iff[OF assms, symmetric]
by (intro iffI) (rule B.prv_imp_implies_isTrue; simp)+
qed
lemma bprv_compl_isTrue_PP_enc:
assumes 1: "\<phi> \<in> fmla" "Fvars \<phi> = {}" and 2: "isTrue (PP \<langle>\<phi>\<rangle>)"
shows "bprv (PP \<langle>\<phi>\<rangle>)"
proof-
obtain n where nn: "n \<in> num" and i: "isTrue (PPf n \<langle>\<phi>\<rangle>)"
using 2 unfolding isTrue_exi_iff_PP[OF 1] ..
hence "bprv (PPf n \<langle>\<phi>\<rangle>)"
using i using nn assms isTrue_iff_bprv_PPf by blast
hence "bprv (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>))"
using 2 assms isTrue_exi_iff isTrue_exi_iff_PP prv_exi_PPf_iff_isTrue by auto
thus ?thesis using PP_PPf 1 by blast
qed
lemma TIP:
assumes 1: "\<phi> \<in> fmla" "Fvars \<phi> = {}" and 2: "isTrue (PP \<langle>\<phi>\<rangle>)"
shows "prv \<phi>"
using bprv_compl_isTrue_PP_enc[OF assms] HBL1_iff assms by blast
end \<comment> \<open>context @{locale Minimal_Truth_Soundness_HBL1iff_Compl_Pf}\<close>
locale Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf =
Minimal_Truth_Soundness_HBL1iff_Compl_Pf
+
assumes
Compl_NegPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
let PPf = (\<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]) in
isTrue (B.neg (PPf n \<langle>\<phi>\<rangle>)) \<longrightarrow> bprv (B.neg (PPf n \<langle>\<phi>\<rangle>))"
begin
lemma isTrue_implies_prv_neg_PPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
isTrue (B.neg (PPf n \<langle>\<phi>\<rangle>)) \<Longrightarrow> bprv (B.neg (PPf n \<langle>\<phi>\<rangle>))"
using Compl_NegPf by(simp add: PPf_def)
lemma isTrue_iff_prv_neg_PPf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> isTrue (B.neg (PPf n \<langle>\<phi>\<rangle>)) \<longleftrightarrow> bprv (B.neg (PPf n \<langle>\<phi>\<rangle>))"
using isTrue_implies_prv_neg_PPf
using exists_no_Fvars B.not_isTrue_fls B.sound_isTrue by auto
lemma prv_PPf_decide:
assumes [simp]: "n \<in> num" "\<phi> \<in> fmla" "Fvars \<phi> = {}"
and np: "\<not> bprv (PPf n \<langle>\<phi>\<rangle>)"
shows "bprv (B.neg (PPf n \<langle>\<phi>\<rangle>))"
proof-
have "\<not> isTrue (PPf n \<langle>\<phi>\<rangle>)" using assms by (auto simp: isTrue_iff_bprv_PPf)
hence "isTrue (B.neg (PPf n \<langle>\<phi>\<rangle>))" using B.isTrue_neg[of "PPf n \<langle>\<phi>\<rangle>"] by auto
thus ?thesis by (auto simp: isTrue_iff_prv_neg_PPf)
qed
lemma not_prfOf_prv_neg_Pf:
assumes n\<phi>: "n \<in> num" "\<phi> \<in> fmla" "Fvars \<phi> = {}" and "\<not> prfOf n \<phi>"
shows "bprv (B.neg (psubst Pf [(n, yy), (\<langle>\<phi>\<rangle>, xx)]))"
using assms prv_PPf_decide[OF n\<phi>] by (auto simp: prfOf_def PPf_def2 psubst_eq_rawpsubst2)
end \<comment> \<open>context @{locale Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf}\<close>
sublocale Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf <
repr: CleanRepr_Proofs
(* added label to avoid duplicated constant P, which is assumed
in Minimal_Truth_Soundness_HBL1iff_Compl_Pf but defined in CleanRepr_Proofs
(they are of course extensionally equal).
*)
where "proof" = "proof" and prfOf = prfOf and encPf = encPf
by standard (auto simp: bprv_prv prv_prfOf prfOf_prv_Pf not_prfOf_prv_neg_Pf)
(* Lemma 6 ("proof recovery") from the JAR paper: *)
sublocale Minimal_Truth_Soundness_HBL1iff_Compl_Pf_Compl_NegPf <
min_truth: Minimal_Truth_Soundness_Proof_Repr
where "proof" = "proof" and prfOf = prfOf and encPf = encPf
by standard
(* FOR THE CLASSICAL REASONING VERSION *)
locale Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf =
HBL1
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv bprv
enc
P
+
B: Minimal_Truth_Soundness
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
bprv
isTrue
+
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and enc ("\<langle>_\<rangle>")
and prv bprv
and P
and isTrue
+
fixes Pf :: 'fmla
assumes
(* Pf is a formula with free variables xx yy *)
Pf[simp,intro!]: "Pf \<in> fmla"
and
Fvars_Pf[simp]: "Fvars Pf = {yy,xx}"
and
(* P relates to Pf internally just like a prv and a proofOf would
relate: via an existential *)
P_Pf:
"\<phi> \<in> fmla \<Longrightarrow>
let PPf = (\<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]) in
bprv (eqv (subst P \<langle>\<phi>\<rangle> xx) (exi yy (PPf (Var yy) \<langle>\<phi>\<rangle>)))"
assumes
(* We assume, in addition to HBL1, the strong form of HBL1_rev: *)
HBL1_rev_prv: "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> prv (PP \<langle>\<phi>\<rangle>) \<Longrightarrow> prv \<phi>"
and
Compl_Pf:
"\<And> n \<phi>. n \<in> num \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow>
let PPf = (\<lambda> t1 t2. psubst Pf [(t1,yy), (t2,xx)]) in
isTrue (PPf n \<langle>\<phi>\<rangle>) \<longrightarrow> bprv (PPf n \<langle>\<phi>\<rangle>)"
begin
lemma HBL1_rev:
assumes f: "\<phi> \<in> fmla" and fv: "Fvars \<phi> = {}" and bp: "bprv (PP \<langle>\<phi>\<rangle>)"
shows "prv \<phi>"
using assms by (auto intro!: HBL1_rev_prv bprv_prv[OF _ _ bp])
lemma HBL1_iff: "\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> bprv (PP \<langle>\<phi>\<rangle>) \<longleftrightarrow> prv \<phi>"
using HBL1 HBL1_rev unfolding PP_def by auto
lemma HBL1_iff_prv: "\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> prv (PP \<langle>\<phi>\<rangle>) \<longleftrightarrow> prv \<phi>"
by (intro iffI bprv_prv[OF _ _ HBL1_PP], elim HBL1_rev_prv[rotated -1]) auto
end \<comment> \<open>context @{locale Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf}\<close>
sublocale Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf <
mts_prv_mts: Minimal_Truth_Soundness_HBL1iff_Compl_Pf where Pf = Pf
using P_Pf HBL1_rev HBL1_PP Compl_Pf
by unfold_locales auto
locale Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf_Classical =
Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf
+
assumes
\<comment> \<open>NB: we don't really need to assume classical reasoning (double negation) all throughout,
but only for the provability predicate:\<close>
classical_P: "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> let PP = (\<lambda>t. subst P t xx) in
prv (B.neg (B.neg (PP \<langle>\<phi>\<rangle>))) \<Longrightarrow> prv (PP \<langle>\<phi>\<rangle>)"
begin
lemma classical_PP: "\<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> = {} \<Longrightarrow> prv (B.neg (B.neg (PP \<langle>\<phi>\<rangle>))) \<Longrightarrow> prv (PP \<langle>\<phi>\<rangle>)"
using classical_P unfolding PP_def by auto
end \<comment> \<open>context @{locale Minimal_Truth_Soundness_HBL1iff_prv_Compl_Pf_Classical}\<close>
(*<*)
end
(*>*)
diff --git a/thys/HOL-CSP/Mndet.thy b/thys/HOL-CSP/Mndet.thy
--- a/thys/HOL-CSP/Mndet.thy
+++ b/thys/HOL-CSP/Mndet.thy
@@ -1,242 +1,242 @@
(*<*)
\<comment>\<open> ********************************************************************
* Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL
* Version : 2.0
*
* Author : Burkhart Wolff, Safouan Taha, Lina Ye.
* (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff)
*
* This file : A Combined CSP Theory
*
* Copyright (c) 2009 Université Paris-Sud, France
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are
* met:
*
* * Redistributions of source code must retain the above copyright
* notice, this list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above
* copyright notice, this list of conditions and the following
* disclaimer in the documentation and/or other materials provided
* with the distribution.
*
* * Neither the name of the copyright holders nor the names of its
* contributors may be used to endorse or promote products derived
* from this software without specific prior written permission.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
* "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
* LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
* A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
* OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
* SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
* LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
* DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
* THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
* (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************\<close>
(*>*)
theory Mndet
imports Process Stop Mprefix Ndet
begin
section\<open>Multiple non deterministic operator\<close>
definition
mndet :: "['\<alpha> set, '\<alpha> \<Rightarrow> '\<alpha> process] \<Rightarrow> '\<alpha> process"
where "mndet A P \<equiv> if A = {}
then STOP
else Abs_process(\<Union> x\<in>A. F(x \<rightarrow> P x),
\<Union> x\<in>A. D(x \<rightarrow> P x))"
syntax
"_Mndet" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a process \<Rightarrow> 'a process" ("(3\<sqinter>_\<in>_ \<rightarrow> / _)" [0, 0, 10] 10)
translations
"\<sqinter>x\<in>A\<rightarrow> P" \<rightleftharpoons> "CONST mndet A (\<lambda>x. P)"
lemma mt_mndet[simp] : "mndet {} P = STOP"
unfolding mndet_def by simp
lemma mndet_is_process : "A \<noteq> {} \<Longrightarrow> is_process (\<Union> x\<in>A. F(x \<rightarrow> P x), \<Union> x\<in>A. D(x \<rightarrow> P x))"
unfolding is_process_def FAILURES_def DIVERGENCES_def
apply auto
using is_processT1 apply auto[1]
using is_processT2 apply blast
using is_processT3_SR apply blast
using is_processT4 apply blast
using is_processT5_S1 apply blast
using is_processT6 apply blast
using is_processT7 apply blast
using NF_ND apply auto[1]
using is_processT9 by blast
lemma T_mdet1 : "T(mndet {} P) = {[]}"
unfolding mndet_def by(simp add: T_STOP)
lemma rep_abs_mndet[simp]: "A \<noteq> {} \<Longrightarrow>
(Rep_process (Abs_process(\<Union>x\<in>A. F (x \<rightarrow> P x),\<Union>x\<in>A. D (x \<rightarrow> P x)))) =
(\<Union>x\<in>A. F (x \<rightarrow> P x), \<Union>x\<in>A. D (x \<rightarrow> P x))"
apply(subst Process.process.Abs_process_inverse)
by(auto intro: mndet_is_process[simplified])
lemma T_mndet: "A\<noteq>{} \<Longrightarrow> T(mndet A P) = (\<Union> x\<in>A. T (x \<rightarrow> P x))"
unfolding mndet_def
apply(simp, subst T_def, simp add: TRACES_def FAILURES_def)
apply(auto intro: mndet_is_process[simplified])
unfolding TRACES_def FAILURES_def apply(cases "A = {}")
apply(auto intro: F_T D_T simp: Nil_elem_T)
using NF_NT by blast
lemma F_mndet1 : "F(mndet {} P) = {(s, X). s = []}"
unfolding mndet_def by(simp add: F_STOP)
lemma F_mndet: "A\<noteq>{} \<Longrightarrow> F(mndet A P) = (\<Union> x\<in>A. F (x \<rightarrow> P x))"
unfolding mndet_def
by(simp, subst F_def, auto simp : FAILURES_def F_mndet1)
lemma D_mndet1 : "D(mndet {} P) = {}"
unfolding mndet_def by(simp add: D_STOP)
lemma D_mndet : "A\<noteq>{} \<Longrightarrow> D(mndet A P) = (\<Union> x\<in>A. D (x \<rightarrow> P x))"
unfolding mndet_def
apply(simp, subst D_def, subst Process.process.Abs_process_inverse)
by(auto intro: mndet_is_process[simplified] simp: DIVERGENCES_def)
text\<open> Thus we know now, that mndet yields processes. Direct consequences are the following
distributivities: \<close>
lemma mndet_unit : "(\<sqinter> x \<in> {a} \<rightarrow> P x) = (a \<rightarrow> P a)"
by(auto simp : Process.Process_eq_spec F_mndet D_mndet)
lemma mndet_distrib : "A \<noteq>{} \<Longrightarrow> B \<noteq>{} \<Longrightarrow> (\<sqinter> x \<in> A\<union>B \<rightarrow> P x) = ((\<sqinter> x \<in> A \<rightarrow> P x) \<sqinter> (\<sqinter> x \<in> B \<rightarrow> P x))"
by(auto simp : Process.Process_eq_spec F_ndet D_ndet F_mndet D_mndet)
text\<open> The two lemmas @{thm mndet_unit} and @{thm mndet_distrib} together give us that @{const mndet}
can be represented by a fold in the finite case. \<close>
lemma mndet_distrib_unit : "A-{a} \<noteq> {} \<Longrightarrow> (\<sqinter> x \<in> insert a A \<rightarrow> P x) = ((a \<rightarrow> P a) \<sqinter> (\<sqinter> x \<in> A-{a} \<rightarrow> P x))"
by (metis Un_Diff_cancel insert_is_Un insert_not_empty mndet_distrib mndet_unit)
subsection\<open>Finite case Continuity\<close>
text\<open> This also implies that mndet is continuous for the
finite @{term A} and an arbitrary body @{term f}: \<close>
lemma Mndet_cont_finite[simp]:
assumes "finite A"
and "\<And>x. cont (f x)"
shows "cont (\<lambda>y. \<sqinter> z \<in> A \<rightarrow> f z y)"
-proof(rule Finite_Set.finite_induct[OF `finite A`])
+proof(rule Finite_Set.finite_induct[OF \<open>finite A\<close>])
show "cont (\<lambda>y. \<sqinter>z\<in>{} \<rightarrow> f z y)" by auto
next
fix A fix a
assume "cont (\<lambda>y. \<sqinter>z\<in>A \<rightarrow> f z y)" and "a \<notin> A"
show "cont (\<lambda>y. \<sqinter>z\<in>insert a A \<rightarrow> f z y)"
proof(cases "A={}")
case True
- then show ?thesis by(simp add: mndet_unit True `\<And>x. cont (f x)`)
+ then show ?thesis by(simp add: mndet_unit True \<open>\<And>x. cont (f x)\<close>)
next
case False
have * : "A-{a} \<noteq> {}" by (simp add: False \<open>a \<notin> A\<close>)
have ** : "A-{a} = A" by (simp add: \<open>a \<notin> A\<close>)
show ?thesis
apply(simp only: mndet_distrib_unit[OF *], simp only: **)
by (simp add: \<open>cont (\<lambda>y. \<sqinter>z\<in>A \<rightarrow> f z y)\<close> assms(2))
qed
qed
subsection\<open>General case Continuity\<close>
lemma mono_Mndet[simp] : "monofun (mndet (A::'a set))"
proof(cases "A={}")
case True
then show ?thesis by(auto simp: monofun_def)
next
case False
then show ?thesis apply(simp add: monofun_def, intro allI impI)
unfolding le_approx_def
proof(simp add:T_mndet F_mndet D_mndet, intro conjI)
fix x::"'a \<Rightarrow> 'a process" and y::"'a \<Rightarrow> 'a process"
assume "A \<noteq> {}" and "x \<sqsubseteq> y"
show "(\<Union>x\<in>A. D (x\<rightarrow> y x)) \<subseteq> (\<Union>xa\<in>A. D (xa \<rightarrow> x xa))"
by (metis (mono_tags, lifting) SUP_mono \<open>x \<sqsubseteq> y\<close> fun_below_iff le_approx1 mono_Mprefix0 write0_def)
next
fix x::"'a \<Rightarrow> 'a process" and y::"'a \<Rightarrow> 'a process"
assume *:"A \<noteq> {}" and **:"x \<sqsubseteq> y"
have *** : "\<And>z. z \<in> A \<Longrightarrow> x z \<sqsubseteq> y z " by (simp add: \<open>x \<sqsubseteq> y\<close> fun_belowD)
with * show "\<forall>s. (\<forall>xa\<in>A. s \<notin> D (xa \<rightarrow> x xa)) \<longrightarrow> Ra (mndet A x) s = Ra (mndet A y) s"
unfolding Ra_def
by (auto simp:proc_ord2a mono_Mprefix0 write0_def F_mndet) (meson le_approx2 mono_Mprefix0 write0_def)+
next
fix x::"'a \<Rightarrow> 'a process" and y::"'a \<Rightarrow> 'a process"
assume *:"A \<noteq> {}" and "x \<sqsubseteq> y"
have *** : "\<And>z. z \<in> A \<Longrightarrow> (z \<rightarrow> x z) \<sqsubseteq> (z \<rightarrow> y z) "
by (metis \<open>x \<sqsubseteq> y\<close> fun_below_iff mono_Mprefix0 write0_def)
with * show "min_elems (\<Union>xa\<in>A. D (xa \<rightarrow> x xa)) \<subseteq> (\<Union>x\<in>A. T (x \<rightarrow> y x))"
unfolding min_elems_def apply auto
by (metis Set.set_insert elem_min_elems insert_subset le_approx3 le_less min_elems5)
qed
qed
lemma mndet_chainpreserving: "chain Y \<Longrightarrow> chain (\<lambda>i. (\<sqinter> z \<in> A \<rightarrow> Y i z))"
apply(rule chainI, rename_tac i)
apply(frule_tac i="i" in chainE)
by (simp add: below_fun_def mono_Mprefix0 write0_def monofunE)
lemma contlub_Mndet : "contlub (mndet A)"
proof(cases "A={}")
case True
then show ?thesis by(auto simp: contlub_def)
next
case False
show ?thesis
proof(rule contlubI, rule proc_ord_proc_eq_spec)
fix Y :: "nat \<Rightarrow> 'a \<Rightarrow> 'a process"
assume a:"chain Y"
show "(\<sqinter>x\<in>A \<rightarrow> (\<Squnion>i. Y i) x) \<sqsubseteq> (\<Squnion>i. \<sqinter>x\<in>A \<rightarrow> Y i x)"
proof(simp add:le_approx_def, intro conjI allI impI)
show "D (\<Squnion>i. \<sqinter>x\<in>A \<rightarrow> Y i x) \<subseteq> D (\<sqinter>x\<in>A \<rightarrow> Lub Y x)"
using a False D_LUB[OF mndet_chainpreserving[OF a], of A]
limproc_is_thelub[OF mndet_chainpreserving[OF a], of A]
apply (auto simp add:write0_def D_Mprefix D_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a] D_mndet)
by (metis (mono_tags, hide_lams) event.inject)
next
fix s :: "'a event list"
assume "s \<notin> D (\<sqinter>x\<in>A \<rightarrow> Lub Y x)"
show "Ra (\<sqinter>x\<in>A \<rightarrow> Lub Y x) s = Ra (\<Squnion>i. \<sqinter>x\<in>A \<rightarrow> Y i x) s"
unfolding Ra_def
using a False F_LUB[OF mndet_chainpreserving[OF a], of A]
limproc_is_thelub[OF mndet_chainpreserving[OF a], of A]
apply (auto simp add:write0_def F_Mprefix F_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a] F_mndet)
by (metis (mono_tags, hide_lams) event.inject)
next
show "min_elems (D (\<sqinter>x\<in>A \<rightarrow> Lub Y x)) \<subseteq> T (\<Squnion>i. \<sqinter>x\<in>A \<rightarrow> Y i x)"
unfolding min_elems_def
using a False limproc_is_thelub[OF mndet_chainpreserving[OF a], of A]
D_LUB[OF mndet_chainpreserving[OF a], of A]
F_LUB[OF mndet_chainpreserving[OF a], of A]
by (auto simp add:D_T write0_def D_Mprefix F_Mprefix D_mndet F_mndet D_LUB[OF ch2ch_fun[OF a]] F_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a])
qed
next
fix Y :: "nat \<Rightarrow> 'a \<Rightarrow> 'a process"
assume a:"chain Y"
show "D (\<sqinter>x\<in>A \<rightarrow> (\<Squnion>i. Y i) x) \<subseteq> D (\<Squnion>i. \<sqinter>x\<in>A \<rightarrow> Y i x)"
using a False D_LUB[OF mndet_chainpreserving[OF a], of A]
limproc_is_thelub[OF mndet_chainpreserving[OF a], of A]
by (auto simp add:write0_def D_Mprefix D_mndet D_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a])
qed
qed
lemma Mndet_cont[simp]: "(\<And>x. cont (f x)) \<Longrightarrow> cont (\<lambda>y. (\<sqinter> z \<in> A \<rightarrow> (f z y)))"
apply(rule_tac f = "\<lambda>z y. (f y z)" in Cont.cont_compose, rule monocontlub2cont)
by (auto intro: mono_Mndet contlub_Mndet Fun_Cpo.cont2cont_lambda)
end
\ No newline at end of file
diff --git a/thys/HOL-CSP/Mprefix.thy b/thys/HOL-CSP/Mprefix.thy
--- a/thys/HOL-CSP/Mprefix.thy
+++ b/thys/HOL-CSP/Mprefix.thy
@@ -1,429 +1,429 @@
(*<*)
\<comment>\<open> ********************************************************************
* Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL
* Version : 2.0
*
* Author : Burkhart Wolff, Safouan Taha, Lina Ye.
* (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff)
*
* This file : A Combined CSP Theory
*
* Copyright (c) 2009 Université Paris-Sud, France
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are
* met:
*
* * Redistributions of source code must retain the above copyright
* notice, this list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above
* copyright notice, this list of conditions and the following
* disclaimer in the documentation and/or other materials provided
* with the distribution.
*
* * Neither the name of the copyright holders nor the names of its
* contributors may be used to endorse or promote products derived
* from this software without specific prior written permission.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
* "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
* LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
* A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
* OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
* SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
* LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
* DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
* THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
* (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************\<close>
(*>*)
section\<open> The Multi-Prefix Operator Definition \<close>
theory Mprefix
imports Process
begin
subsection\<open> The Definition and some Consequences \<close>
definition Mprefix :: "['a set,'a => 'a process] => 'a process" where
"Mprefix A P \<equiv> Abs_process(
{(tr,ref). tr = [] \<and> ref Int (ev ` A) = {}} \<union>
{(tr,ref). tr \<noteq> [] \<and> hd tr \<in> (ev ` A) \<and>
(\<exists> a. ev a = (hd tr)\<and>(tl tr,ref)\<in>F(P a))},
{d. d \<noteq> [] \<and> hd d \<in> (ev ` A) \<and>
(\<exists> a. ev a = hd d \<and> tl d \<in> D(P a))})"
syntax(xsymbols)
"@mprefix" :: "[pttrn,'a set,'a process] \<Rightarrow> 'a process" ("(3\<box>(_) \<in> (_) \<rightarrow> (_))"[0,0,64]64)
translations
"\<box>x\<in>A \<rightarrow> P" => "CONST Mprefix A (\<lambda>x. P)"
subsection\<open> Well-foundedness of Mprefix \<close>
lemma is_process_REP_Mp :
"is_process ({(tr,ref). tr=[] \<and> ref \<inter> (ev ` A) = {}} \<union>
{(tr,ref). tr \<noteq> [] \<and> hd tr \<in> (ev ` A) \<and>
(\<exists> a. ev a = (hd tr) \<and> (tl tr,ref) \<in> F(P a))},
{d. d \<noteq> [] \<and> hd d \<in> (ev ` A) \<and>
(\<exists> a. ev a = hd d \<and> tl d \<in> D(P a))})"
(is "is_process(?f, ?d)")
proof (simp only:is_process_def FAILURES_def DIVERGENCES_def
Product_Type.fst_conv Product_Type.snd_conv,
intro conjI allI impI)
show "([],{}) \<in> ?f" by(simp)
next
fix s:: "'a event list" fix X::"'a event set"
assume H : "(s, X) \<in> ?f"
show "front_tickFree s"
apply(insert H, auto simp: front_tickFree_def tickFree_def
dest!:list_nonMt_append)
apply(rename_tac aa ta x, case_tac "ta", auto simp: front_tickFree_charn
dest! : is_processT2[rule_format])
apply(simp add: tickFree_def)
done
next
fix s t :: "'a event list"
assume "(s @ t, {}) \<in> ?f"
then show "(s, {}) \<in> ?f"
by(auto elim: is_processT3[rule_format])
next
fix s:: "'a event list" fix X Y::"'a event set"
assume "(s, Y) \<in> ?f \<and> X \<subseteq> Y"
then show "(s, X) \<in> ?f"
by(auto intro: is_processT4[rule_format])
next
fix s:: "'a event list" fix X Y::"'a event set"
assume "(s, X) \<in> ?f \<and> (\<forall> c. c\<in>Y \<longrightarrow> (s @ [c], {}) \<notin> ?f)"
then show "(s, X \<union> Y) \<in> ?f "
by(auto intro!: is_processT1 is_processT5[rule_format])
next
fix s:: "'a event list" fix X::"'a event set"
assume "(s @ [tick], {}) \<in> ?f"
then show "(s, X - {tick}) \<in> ?f"
by(cases s, auto dest!: is_processT6[rule_format])
next
fix s t:: "'a event list" fix X::"'a event set"
assume "s \<in> ?d \<and> tickFree s \<and> front_tickFree t"
then show "s @ t \<in> ?d"
by(auto intro!: is_processT7_S, cases s, simp_all)
next
fix s:: "'a event list" fix X::"'a event set"
assume "s \<in> ?d"
then show "(s, X) \<in> ?f"
by(auto simp: is_processT8_S)
next
fix s:: "'a event list"
assume "s @ [tick] \<in> ?d"
then show "s \<in> ?d"
apply(auto)
apply(cases s, simp_all)
apply(cases s, auto intro: is_processT9[rule_format])
done
qed
(* bizarre exercise in style proposed by Makarius... *)
lemma is_process_REP_Mp' :
"is_process ({(tr,ref). tr=[] \<and> ref \<inter> (ev ` A) = {}} \<union>
{(tr,ref). tr \<noteq> [] \<and> hd tr \<in> (ev ` A) \<and>
(\<exists> a. ev a = (hd tr) \<and> (tl tr,ref) \<in> F(P a))},
{d. d \<noteq> [] \<and> hd d \<in> (ev ` A) \<and>
(\<exists> a. ev a = hd d \<and> tl d \<in> D(P a))})"
(is "is_process(?f, ?d)")
proof (simp only:is_process_def FAILURES_def DIVERGENCES_def
Product_Type.fst_conv Product_Type.snd_conv,
intro conjI allI impI,goal_cases)
case 1
show "([],{}) \<in> ?f" by(simp )
next
case (2 s X)
assume H : "(s, X) \<in> ?f"
have "front_tickFree s"
apply(insert H, auto simp: front_tickFree_def tickFree_def
dest!:list_nonMt_append)
apply(rename_tac aa ta x, case_tac "ta", auto simp: front_tickFree_charn
dest! : is_processT2[rule_format])
apply(simp add: tickFree_def)
done
then show "front_tickFree s" by(auto)
next
case (3 s t)
assume H : "(s @ t, {}) \<in> ?f"
show "(s, {}) \<in> ?f" using H by(auto elim: is_processT3[rule_format])
next
case (4 s X Y)
assume H1: "(s, Y) \<in> ?f \<and> X \<subseteq> Y"
then show "(s, X) \<in> ?f" by(auto intro: is_processT4[rule_format])
next
case (5 s X Y)
assume "(s, X) \<in> ?f \<and> (\<forall> c. c\<in>Y \<longrightarrow> (s @ [c], {}) \<notin> ?f)"
then show "(s, X \<union> Y) \<in> ?f" by(auto intro!: is_processT1 is_processT5[rule_format])
next
case (6 s X)
assume "(s @ [tick], {}) \<in> ?f"
then show "(s, X - {tick}) \<in> ?f" by(cases s, auto dest!: is_processT6[rule_format])
next
case (7 s t)
assume H1 : "s \<in> ?d \<and> tickFree s \<and> front_tickFree t"
have 7: "s @ t \<in> ?d"
using H1 by(auto intro!: is_processT7_S, cases s, simp_all)
then show "s @ t \<in> ?d" using H1 by(auto)
next
case (8 s X)
assume H : "s \<in> ?d"
then show "(s, X) \<in> ?f" using H by(auto simp: is_processT8_S)
next
case (9 s)
assume H: "s @ [tick] \<in> ?d"
then have 9: "s \<in> ?d"
apply(auto)
apply(cases s, simp_all)
apply(cases s, auto intro: is_processT9[rule_format])
done
then show "s \<in> ?d" by(auto)
qed
lemmas Rep_Abs_Mp[simp] = Abs_process_inverse[simplified, OF is_process_REP_Mp]
thm Rep_Abs_Mp
lemma Rep_Abs_Mp'' :
assumes H1 : "f = {(tr, ref). tr = [] \<and> ref \<inter> ev ` A = {}} \<union>
{(tr, ref). tr \<noteq> [] \<and> hd tr \<in> ev ` A
\<and> (\<exists>a. ev a = hd tr \<and> (tl tr, ref) \<in> F (P a))}"
and H2 : "d = {d. d \<noteq> [] \<and> hd d \<in> (ev ` A) \<and>
(\<exists> a. ev a = hd d \<and> tl d \<in> D(P a))}"
shows "Rep_process (Abs_process (f,d)) = (f,d)"
by(subst Abs_process_inverse,
simp_all only: H1 H2 CollectI Rep_process is_process_REP_Mp)
subsection \<open> Projections in Prefix \<close>
lemma F_Mprefix :
"F(\<box> x \<in> A \<rightarrow> P x) = {(tr,ref). tr=[] \<and> ref \<inter> (ev ` A) = {}} \<union>
{(tr,ref). tr \<noteq> [] \<and> hd tr \<in> (ev ` A) \<and>
(\<exists> a. ev a = (hd tr) \<and> (tl tr,ref) \<in> F(P a))}"
by(simp add:Mprefix_def F_def Rep_Abs_Mp'' FAILURES_def)
lemma D_Mprefix:
"D(\<box> x \<in> A \<rightarrow> P x) = {d. d \<noteq> [] \<and> hd d \<in> (ev ` A) \<and>
(\<exists> a. ev a = hd d \<and> tl d \<in> D(P a))}"
by(simp add:Mprefix_def D_def Rep_Abs_Mp'' DIVERGENCES_def)
lemma T_Mprefix:
"T(\<box> x \<in> A \<rightarrow> P x)={s. s=[] \<or> (\<exists> a. a\<in>A \<and> s\<noteq>[] \<and> hd s = ev a \<and> tl s \<in> T(P a))}"
by(auto simp: T_F_spec[symmetric] F_Mprefix)
subsection \<open> Basic Properties \<close>
lemma tick_T_Mprefix [simp]: "[tick] \<notin> T(\<box> x \<in> A \<rightarrow> P x)"
by(simp add:T_Mprefix)
lemma Nil_Nin_D_Mprefix [simp]: "[] \<notin> D(\<box> x \<in> A \<rightarrow> P x)"
by(simp add: D_Mprefix)
subsection\<open> Proof of Continuity Rule \<close>
subsubsection\<open> Backpatch Isabelle 2009-1\<close>
\<comment>\<open>re-introduced from Isabelle/HOLCF 2009-1; clearly
a derived concept, but still a useful shortcut\<close>
definition
contlub :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool"
where
"contlub f = (\<forall>Y. chain Y \<longrightarrow> f (\<Squnion> i. Y i) = (\<Squnion> i. f (Y i)))"
lemma contlubE:
"\<lbrakk>contlub f; chain Y\<rbrakk> \<Longrightarrow> f (\<Squnion> i. Y i) = (\<Squnion> i. f (Y i))"
by (simp add: contlub_def)
lemma monocontlub2cont: "\<lbrakk>monofun f; contlub f\<rbrakk> \<Longrightarrow> cont f"
apply (rule contI)
apply (rule thelubE)
apply (erule (1) ch2ch_monofun)
apply (erule (1) contlubE [symmetric])
done
lemma contlubI:
"(\<And>Y. chain Y \<Longrightarrow> f (\<Squnion> i. Y i) = (\<Squnion> i. f (Y i))) \<Longrightarrow> contlub f"
by (simp add: contlub_def)
lemma cont2contlub: "cont f \<Longrightarrow> contlub f"
apply (rule contlubI)
apply (rule Porder.po_class.lub_eqI [symmetric])
apply (erule (1) contE)
done
subsubsection\<open> Core of Proof \<close>
lemma mono_Mprefix1:
"\<forall>a\<in>A. P a \<sqsubseteq> Q a \<Longrightarrow> D (Mprefix A Q) \<subseteq> D (Mprefix A P)"
apply(auto simp: D_Mprefix) using le_approx1 by blast
lemma mono_Mprefix2:
"\<forall>x\<in>A. P x \<sqsubseteq> Q x \<Longrightarrow>
\<forall>s. s \<notin> D (Mprefix A P) \<longrightarrow> Ra (Mprefix A P) s = Ra (Mprefix A Q) s"
apply (auto simp: Ra_def D_Mprefix F_Mprefix) using proc_ord2a by blast+
lemma mono_Mprefix3 :
assumes H:"\<forall>x\<in>A. P x \<sqsubseteq> Q x"
shows " min_elems (D (Mprefix A P)) \<subseteq> T (Mprefix A Q)"
proof(auto simp: min_elems_def D_Mprefix T_Mprefix image_def, goal_cases)
case (1 x a)
with H[rule_format, of a, OF 1(2)] show ?case
apply(auto dest!: le_approx3 simp: min_elems_def)
apply(subgoal_tac "\<forall>t. t \<in> D (P a) \<longrightarrow> \<not> t < tl x", auto)
apply(rename_tac t, erule_tac x="(ev a)#t" in allE, auto)
using less_cons hd_Cons_tl by metis
qed
lemma mono_Mprefix0:
"\<forall>x\<in>A. P x \<sqsubseteq> Q x \<Longrightarrow> Mprefix A P \<sqsubseteq> Mprefix A Q"
apply(simp add: le_approx_def mono_Mprefix1 mono_Mprefix3)
apply(rule mono_Mprefix2)
apply(auto simp: le_approx_def)
done
lemma mono_Mprefix[simp]: "monofun(Mprefix A)"
by(auto simp: Fun_Cpo.below_fun_def monofun_def mono_Mprefix0)
lemma proc_ord2_set:
"P \<sqsubseteq> Q \<Longrightarrow> {(s, X). s \<notin> D P \<and> (s, X) \<in> F P} = {(s, X). s \<notin> D P \<and> (s, X) \<in> F Q}"
by(auto simp: le_approx2)
lemma proc_ord_proc_eq_spec: "P \<sqsubseteq> Q \<Longrightarrow> D P \<subseteq> D Q \<Longrightarrow> P = Q"
by (metis (mono_tags, lifting) below_antisym below_refl le_approx_def subset_antisym)
lemma mprefix_chainpreserving: "chain Y \<Longrightarrow> chain (\<lambda>i. Mprefix A (Y i))"
apply(rule chainI, rename_tac i)
apply(frule_tac i="i" in chainE)
by(simp add: mono_Mprefix0 fun_belowD)
lemma limproc_is_thelub_fun:
assumes "chain S"
shows "(Lub S c) = lim_proc (range (\<lambda>x. (S x c)))"
proof -
have "\<And>xa. chain (\<lambda>x. S x xa)"
- using `chain S` by(auto intro!: chainI simp: chain_def fun_belowD )
+ using \<open>chain S\<close> by(auto intro!: chainI simp: chain_def fun_belowD )
then show ?thesis by (metis contlub_lambda limproc_is_thelub)
qed
lemma contlub_Mprefix : "contlub(%P. Mprefix A P)"
proof(rule contlubI, rule proc_ord_proc_eq_spec)
fix Y :: "nat \<Rightarrow> 'a \<Rightarrow> 'a process"
assume C : "chain Y"
have C': "\<And>xa. chain (\<lambda>x. Y x xa)"
apply(insert C,rule chainI)
by(auto simp: chain_def fun_belowD)
show "Mprefix A (\<Squnion> i. Y i) \<sqsubseteq> (\<Squnion> i. Mprefix A (Y i))"
by(auto simp: Process.le_approx_def F_Mprefix D_Mprefix T_Mprefix C C'
mprefix_chainpreserving limproc_is_thelub limproc_is_thelub_fun
D_T D_LUB D_LUB_2 F_LUB T_LUB_2 Ra_def min_elems_def)
next
fix Y :: "nat \<Rightarrow> 'a \<Rightarrow> 'a process"
assume C : "chain Y"
show "D (Mprefix A (\<Squnion> i. Y i)) \<subseteq> D (\<Squnion> i. Mprefix A (Y i))"
apply(auto simp: Process.le_approx_def F_Mprefix D_Mprefix T_Mprefix
C mprefix_chainpreserving limproc_is_thelub D_LUB_2)
by (meson C fun_below_iff in_mono is_ub_thelub le_approx1)
qed
lemma Mprefix_cont [simp]:
"(\<And>x. cont (f x)) \<Longrightarrow> cont (\<lambda>y. \<box> z \<in> A \<rightarrow> f z y)"
apply(rule_tac f = "\<lambda>z y. (f y z)" in Cont.cont_compose)
apply(rule monocontlub2cont)
apply(auto intro: mono_Mprefix contlub_Mprefix Fun_Cpo.cont2cont_lambda)
done
subsection\<open> High-level Syntax for Read and Write \<close>
text\<open> The following syntax introduces the usual channel notation for CSP.
Slightly deviating from conventional wisdom, we view a channel not as a tag in
a pair, rather than as a function of type @{typ "'a\<Rightarrow>'b"}. This paves the way
for \emph{typed} channels over a common universe of events. \<close>
definition read :: "['a \<Rightarrow> 'b,'a set, 'a \<Rightarrow> 'b process] \<Rightarrow> 'b process"
where "read c A P \<equiv> Mprefix(c ` A) (P o (inv c))"
definition "write" :: "['a \<Rightarrow> 'b, 'a, 'b process] \<Rightarrow> 'b process"
where "write c a P \<equiv> Mprefix {c a} (\<lambda> x. P)"
definition write0 :: "['a, 'a process] \<Rightarrow> 'a process"
where "write0 a P \<equiv> Mprefix {a} (\<lambda> x. P)"
syntax
"_read" :: "[id, pttrn, 'a process] => 'a process"
("(3_`?`_ /\<rightarrow> _)" [0,0,28] 28)
"_readX" :: "[id, pttrn, bool,'a process] => 'a process"
("(3_`?`_`|`_ /\<rightarrow> _)" [0,0,28] 28)
"_readS" :: "[id, pttrn, 'b set,'a process] => 'a process"
("(3_`?`_`:`_ /\<rightarrow> _)" [0,0,28] 28)
"_write" :: "[id, 'b, 'a process] => 'a process"
("(3_`!`_ /\<rightarrow> _)" [0,0,28] 28)
"_writeS" :: "['a, 'a process] => 'a process"
("(3_ /\<rightarrow> _)" [0,28] 28)
subsection\<open>CSP$_M$-Style Syntax for Communication Primitives\<close>
translations
"_read c p P" == "CONST read c CONST UNIV (\<lambda>p. P)"
"_write c p P" == "CONST write c p P"
"_readX c p b P" => "CONST read c {p. b} (\<lambda>p. P)"
"_writeS a P" == "CONST write0 a P"
"_readS c p A P" => "CONST read c A (\<lambda>p. P)"
lemma read_cont[simp]:
"(\<And>x. cont (f x)) \<Longrightarrow> cont (\<lambda>y. c`?`x \<rightarrow> f x y)"
by(simp add:read_def o_def inv_def)
lemma write_cont[simp]:
"(\<And>x. cont (P::('b::cpo \<Rightarrow> 'a process))) \<Longrightarrow> cont(\<lambda>x. (c `!` a \<rightarrow> P x))"
by(simp add:write_def)
corollary write0_cont_lub : "contlub(Mprefix {a})"
using contlub_Mprefix by blast
lemma write0_contlub : "contlub(write0 a)"
unfolding write0_def contlub_def
proof (auto)
fix Y :: "nat \<Rightarrow> 'a process"
assume "chain Y"
have * : "chain (\<lambda>i. (\<lambda>_. Y i)::'b \<Rightarrow> 'a process)"
by (meson \<open>chain Y\<close> fun_below_iff po_class.chain_def)
have **: "(\<lambda>x. Lub Y) = Lub (\<lambda>i. (\<lambda>_. Y i))"
by(rule ext,metis * ch2ch_fun limproc_is_thelub limproc_is_thelub_fun lub_eq)
show "Mprefix {a} (\<lambda>x. Lub Y) = (\<Squnion>i. Mprefix {a} (\<lambda>x. Y i))"
apply(subst **, subst contlubE[OF contlub_Mprefix])
by (simp_all add: *)
qed
lemma write0_cont[simp]:
"cont (P::('b::cpo \<Rightarrow> 'a process)) \<Longrightarrow> cont(\<lambda>x. (a \<rightarrow> P x))"
by(simp add:write0_def)
end
diff --git a/thys/Inductive_Inference/CONS_LIM.thy b/thys/Inductive_Inference/CONS_LIM.thy
--- a/thys/Inductive_Inference/CONS_LIM.thy
+++ b/thys/Inductive_Inference/CONS_LIM.thy
@@ -1,544 +1,544 @@
section \<open>CONS is a proper subset of LIM\label{s:cons_lim}\<close>
theory CONS_LIM
imports Inductive_Inference_Basics
begin
text \<open>That there are classes in @{term "LIM - CONS"} was noted by
Barzdin~\cite{b-iiafp-74,b-iiafp-77} and Blum and Blum~\cite{bb-tmtii-75}. It
was proven by Wiehagen~\cite{w-lerfss-76} (see also Wiehagen and
Zeugmann~\cite{wz-idmowle-94}). The proof uses this class:\<close>
definition U_LIMCONS :: "partial1 set" ("U\<^bsub>LIM-CONS\<^esub>") where
"U\<^bsub>LIM-CONS\<^esub> \<equiv> {vs @ [j] \<odot> p| vs j p. j \<ge> 2 \<and> p \<in> \<R>\<^sub>0\<^sub>1 \<and> \<phi> j = vs @ [j] \<odot> p}"
text \<open>Every function in @{term "U\<^bsub>LIM-CONS\<^esub>"} carries a Gödel number
greater or equal two of itself, after which only zeros and ones occur.
Thus, a strategy that always outputs the rightmost value greater or equal two
in the given prefix will converge to this Gödel number.
The next function searches an encoded list for the rightmost element
greater or equal two.\<close>
definition rmge2 :: partial1 where
"rmge2 e \<equiv>
if \<forall>i<e_length e. e_nth e i < 2 then Some 0
else Some (e_nth e (GREATEST i. i < e_length e \<and> e_nth e i \<ge> 2))"
lemma rmge2:
assumes "xs = list_decode e"
shows "rmge2 e =
(if \<forall>i<length xs. xs ! i < 2 then Some 0
else Some (xs ! (GREATEST i. i < length xs \<and> xs ! i \<ge> 2)))"
proof -
have "(i < e_length e \<and> e_nth e i \<ge> 2) = (i < length xs \<and> xs ! i \<ge> 2)" for i
using assms by simp
then have "(GREATEST i. i < e_length e \<and> e_nth e i \<ge> 2) =
(GREATEST i. i < length xs \<and> xs ! i \<ge> 2)"
by simp
moreover have "(\<forall>i<length xs. xs ! i < 2) = (\<forall>i<e_length e. e_nth e i < 2)"
using assms by simp
moreover have "(GREATEST i. i < length xs \<and> xs ! i \<ge> 2) < length xs" (is "Greatest ?P < _")
if "\<not> (\<forall>i<length xs. xs ! i < 2)"
using that GreatestI_ex_nat[of ?P] le_less_linear order.asym by blast
ultimately show ?thesis using rmge2_def assms by auto
qed
lemma rmge2_init:
"rmge2 (f \<triangleright> n) =
(if \<forall>i<Suc n. the (f i) < 2 then Some 0
else Some (the (f (GREATEST i. i < Suc n \<and> the (f i) \<ge> 2))))"
proof -
let ?xs = "prefix f n"
have "f \<triangleright> n = list_encode ?xs" by (simp add: init_def)
moreover have "(\<forall>i<Suc n. the (f i) < 2) = (\<forall>i<length ?xs. ?xs ! i < 2)"
by simp
moreover have "(GREATEST i. i < Suc n \<and> the (f i) \<ge> 2) =
(GREATEST i. i < length ?xs \<and> ?xs ! i \<ge> 2)"
using length_prefix[of f n] prefix_nth[of _ n f] by metis
moreover have "(GREATEST i. i < Suc n \<and> the (f i) \<ge> 2) < Suc n"
if "\<not> (\<forall>i<Suc n. the (f i) < 2)"
using that GreatestI_ex_nat[of "\<lambda>i. i<Suc n \<and> the (f i) \<ge> 2" n] by fastforce
ultimately show ?thesis using rmge2 by auto
qed
corollary rmge2_init_total:
assumes "total1 f"
shows "rmge2 (f \<triangleright> n) =
(if \<forall>i<Suc n. the (f i) < 2 then Some 0
else f (GREATEST i. i < Suc n \<and> the (f i) \<ge> 2))"
using assms total1_def rmge2_init by auto
lemma rmge2_in_R1: "rmge2 \<in> \<R>"
proof -
define g where
"g = Cn 3 r_ifle [r_constn 2 2, Cn 3 r_nth [Id 3 2, Id 3 0], Cn 3 r_nth [Id 3 2, Id 3 0], Id 3 1]"
then have "recfn 3 g" by simp
then have g: "eval g [j, r, e] \<down>= (if 2 \<le> e_nth e j then e_nth e j else r)" for j r e
using g_def by simp
let ?h = "Pr 1 Z g"
have "recfn 2 ?h"
by (simp add: \<open>recfn 3 g\<close>)
have h: "eval ?h [j, e] =
(if \<forall>i<j. e_nth e i < 2 then Some 0
else Some (e_nth e (GREATEST i. i < j \<and> e_nth e i \<ge> 2)))" for j e
proof (induction j)
case 0
then show ?case using \<open>recfn 2 ?h\<close> by auto
next
case (Suc j)
then have "eval ?h [Suc j, e] = eval g [j, the (eval ?h [j, e]), e]"
using \<open>recfn 2 ?h\<close> by auto
then have *: "eval ?h [Suc j, e] \<down>=
(if 2 \<le> e_nth e j then e_nth e j
else if \<forall>i<j. e_nth e i < 2 then 0
else (e_nth e (GREATEST i. i < j \<and> e_nth e i \<ge> 2)))"
using g Suc by auto
show ?case
proof (cases "\<forall>i<Suc j. e_nth e i < 2")
case True
then show ?thesis using * by auto
next
case ex: False
show ?thesis
proof (cases "2 \<le> e_nth e j")
case True
then have "eval ?h [Suc j, e] \<down>= e_nth e j"
using * by simp
moreover have "(GREATEST i. i < Suc j \<and> e_nth e i \<ge> 2) = j"
using ex True Greatest_equality[of "\<lambda>i. i < Suc j \<and> e_nth e i \<ge> 2"]
by simp
ultimately show ?thesis using ex by auto
next
case False
then have "\<exists>i<j. e_nth e i \<ge> 2"
using ex leI less_Suc_eq by blast
with * have "eval ?h [Suc j, e] \<down>= e_nth e (GREATEST i. i < j \<and> e_nth e i \<ge> 2)"
using False by (smt leD)
moreover have "(GREATEST i. i < Suc j \<and> e_nth e i \<ge> 2) =
(GREATEST i. i < j \<and> e_nth e i \<ge> 2)"
using False ex by (metis less_SucI less_Suc_eq less_antisym numeral_2_eq_2)
ultimately show ?thesis using ex by metis
qed
qed
qed
let ?hh = "Cn 1 ?h [Cn 1 r_length [Id 1 0], Id 1 0]"
have "recfn 1 ?hh"
- using `recfn 2 ?h` by simp
+ using \<open>recfn 2 ?h\<close> by simp
with h have hh: "eval ?hh [e] \<down>=
(if \<forall>i<e_length e. e_nth e i < 2 then 0
else e_nth e (GREATEST i. i < e_length e \<and> e_nth e i \<ge> 2))" for e
by auto
then have "eval ?hh [e] = rmge2 e" for e
unfolding rmge2_def by auto
moreover have "total ?hh"
- using hh totalI1 `recfn 1 ?hh` by simp
- ultimately show ?thesis using `recfn 1 ?hh` by blast
+ using hh totalI1 \<open>recfn 1 ?hh\<close> by simp
+ ultimately show ?thesis using \<open>recfn 1 ?hh\<close> by blast
qed
text \<open>The first part of the main result is that @{term "U\<^bsub>LIM-CONS\<^esub> \<in> LIM"}.\<close>
lemma U_LIMCONS_in_Lim: "U\<^bsub>LIM-CONS\<^esub> \<in> LIM"
proof -
have "U\<^bsub>LIM-CONS\<^esub> \<subseteq> \<R>"
unfolding U_LIMCONS_def using prepend_in_R1 RPred1_subseteq_R1 by blast
have "learn_lim \<phi> U\<^bsub>LIM-CONS\<^esub> rmge2"
proof (rule learn_limI)
show "environment \<phi> U\<^bsub>LIM-CONS\<^esub> rmge2"
using \<open>U_LIMCONS \<subseteq> \<R>\<close> phi_in_P2 rmge2_def rmge2_in_R1 by simp
show "\<exists>i. \<phi> i = f \<and> (\<forall>\<^sup>\<infinity>n. rmge2 (f \<triangleright> n) \<down>= i)" if "f \<in> U\<^bsub>LIM-CONS\<^esub>" for f
proof -
from that obtain vs j p where
j: "j \<ge> 2"
and p: "p \<in> \<R>\<^sub>0\<^sub>1"
and s: "\<phi> j = vs @ [j] \<odot> p"
and f: "f = vs @ [j] \<odot> p"
unfolding U_LIMCONS_def by auto
then have "\<phi> j = f" by simp
from that have "total1 f"
- using `U\<^bsub>LIM-CONS\<^esub> \<subseteq> \<R>` R1_imp_total1 total1_def by auto
+ using \<open>U\<^bsub>LIM-CONS\<^esub> \<subseteq> \<R>\<close> R1_imp_total1 total1_def by auto
define n\<^sub>0 where "n\<^sub>0 = length vs"
have f_gr_n0: "f n \<down>= 0 \<or> f n \<down>= 1" if "n > n\<^sub>0" for n
proof -
have "f n = p (n - n\<^sub>0 - 1)"
using that n\<^sub>0_def f by simp
with RPred1_def p show ?thesis by auto
qed
have "rmge2 (f \<triangleright> n) \<down>= j" if "n \<ge> n\<^sub>0" for n
proof -
have n0_greatest: "(GREATEST i. i < Suc n \<and> the (f i) \<ge> 2) = n\<^sub>0"
proof (rule Greatest_equality)
show "n\<^sub>0 < Suc n \<and> the (f n\<^sub>0) \<ge> 2"
using n\<^sub>0_def f that j by simp
show "\<And>y. y < Suc n \<and> the (f y) \<ge> 2 \<Longrightarrow> y \<le> n\<^sub>0"
proof -
fix y assume "y < Suc n \<and> 2 \<le> the (f y)"
moreover have "p \<in> \<R> \<and> (\<forall>n. p n \<down>= 0 \<or> p n \<down>= 1)"
using RPred1_def p by blast
ultimately show "y \<le> n\<^sub>0"
using f_gr_n0
by (metis Suc_1 Suc_n_not_le_n Zero_neq_Suc le_less_linear le_zero_eq option.sel)
qed
qed
have "f n\<^sub>0 \<down>= j"
using n\<^sub>0_def f by simp
then have "\<not> (\<forall>i<Suc n. the (f i) < 2)"
using j that less_Suc_eq_le by auto
then have "rmge2 (f \<triangleright> n) = f (GREATEST i. i < Suc n \<and> the (f i) \<ge> 2)"
- using rmge2_init_total `total1 f` by auto
- with n0_greatest `f n\<^sub>0 \<down>= j` show ?thesis by simp
+ using rmge2_init_total \<open>total1 f\<close> by auto
+ with n0_greatest \<open>f n\<^sub>0 \<down>= j\<close> show ?thesis by simp
qed
- with `\<phi> j = f` show ?thesis by auto
+ with \<open>\<phi> j = f\<close> show ?thesis by auto
qed
qed
then show ?thesis using Lim_def by auto
qed
text \<open>The class @{term "U_LIMCONS"} is \emph{prefix-complete}, which
means that every non-empty list is the prefix of some function in @{term
"U_LIMCONS"}. To show this we use an auxiliary lemma: For every $f \in
\mathcal{R}$ and $k \in \mathbb{N}$ the value of $f$ at $k$ can be replaced
by a Gödel number of the function resulting from the replacement.\<close>
lemma goedel_at:
fixes m :: nat and k :: nat
assumes "f \<in> \<R>"
shows "\<exists>n\<ge>m. \<phi> n = (\<lambda>x. if x = k then Some n else f x)"
proof -
define psi :: "partial1 \<Rightarrow> nat \<Rightarrow> partial2" where
"psi = (\<lambda>f k i x. (if x = k then Some i else f x))"
have "psi f k \<in> \<R>\<^sup>2"
proof -
obtain r where r: "recfn 1 r" "total r" "eval r [x] = f x" for x
using assms by auto
define r_psi where
"r_psi = Cn 2 r_ifeq [Id 2 1, r_dummy 1 (r_const k), Id 2 0, Cn 2 r [Id 2 1]]"
show ?thesis
proof (rule R2I[of r_psi])
from r_psi_def show "recfn 2 r_psi"
using r(1) by simp
have "eval r_psi [i, x] = (if x = k then Some i else f x)" for i x
proof -
have "eval (Cn 2 r [Id 2 1]) [i, x] = f x"
using r by simp
then have "eval r_psi [i, x] = eval r_ifeq [x, k, i, the (f x)]"
unfolding r_psi_def using \<open>recfn 2 r_psi\<close> r R1_imp_total1[OF assms]
by simp
then show ?thesis using assms by simp
qed
then show "\<And>x y. eval r_psi [x, y] = psi f k x y"
unfolding psi_def by simp
then show "total r_psi"
using totalI2[of r_psi] \<open>recfn 2 r_psi\<close> assms psi_def by fastforce
qed
qed
then obtain n where "n \<ge> m" "\<phi> n = psi f k n"
using assms kleene_fixed_point[of "psi f k" m] by auto
then show ?thesis unfolding psi_def by auto
qed
lemma U_LIMCONS_prefix_complete:
assumes "length vs > 0"
shows "\<exists>f\<in>U\<^bsub>LIM-CONS\<^esub>. prefix f (length vs - 1) = vs"
proof -
let ?p = "\<lambda>_. Some 0"
let ?f = "vs @ [0] \<odot> ?p"
have "?f \<in> \<R>"
using prepend_in_R1 RPred1_subseteq_R1 const0_in_RPred1 by blast
with goedel_at[of ?f 2 "length vs"] obtain j where
j: "j \<ge> 2" "\<phi> j = (\<lambda>x. if x = length vs then Some j else ?f x)" (is "_ = ?g")
by auto
moreover have g: "?g x = (vs @ [j] \<odot> ?p) x" for x
by (simp add: nth_append)
ultimately have "?g \<in> U\<^bsub>LIM-CONS\<^esub>"
unfolding U_LIMCONS_def using const0_in_RPred1 by fastforce
moreover have "prefix ?g (length vs - 1) = vs"
using g assms prefixI prepend_associative by auto
ultimately show ?thesis by auto
qed
text \<open>Roughly speaking, a strategy learning a prefix-complete class
must be total because it must be defined for every prefix in
the class. Technically, however, the empty list is not a prefix, and thus a
strategy may diverge on input 0. We can work around this by
showing that if there is a strategy learning a prefix-complete class then
there is also a total strategy learning this class. We need the result only
for consistent learning.\<close>
lemma U_prefix_complete_imp_total_strategy:
assumes "\<And>vs. length vs > 0 \<Longrightarrow> \<exists>f\<in>U. prefix f (length vs - 1) = vs"
and "learn_cons \<psi> U s"
shows "\<exists>t. total1 t \<and> learn_cons \<psi> U t"
proof -
define t where "t = (\<lambda>e. if e = 0 then Some 0 else s e)"
have "s e \<down>" if "e > 0" for e
proof -
from that have "list_decode e \<noteq> []" (is "?vs \<noteq> _")
using list_encode_0 list_encode_decode by (metis less_imp_neq)
then have "length ?vs > 0" by simp
with assms(1) obtain f where f: "f \<in> U" "prefix f (length ?vs - 1) = ?vs"
by auto
with learn_cons_def learn_limE have "s (f \<triangleright> (length ?vs - 1)) \<down>"
using assms(2) by auto
then show "s e \<down>"
using f(2) init_def by auto
qed
then have "total1 t"
using t_def by auto
have "t \<in> \<P>"
proof -
from assms(2) have "s \<in> \<P>"
using learn_consE by simp
then obtain rs where rs: "recfn 1 rs" "eval rs [x] = s x" for x
by auto
define rt where "rt = Cn 1 (r_lifz Z rs) [Id 1 0, Id 1 0]"
then have "recfn 1 rt"
using rs by auto
moreover have "eval rt [x] = t x" for x
using rs rt_def t_def by simp
ultimately show ?thesis by blast
qed
have "s (f \<triangleright> n) = t (f \<triangleright> n)" if "f \<in> U" for f n
unfolding t_def by (simp add: init_neq_zero)
then have "learn_cons \<psi> U t"
- using `t \<in> \<P>` assms(2) learn_consE[of \<psi> U s] learn_consI[of \<psi> U t] by simp
- with `total1 t` show ?thesis by auto
+ using \<open>t \<in> \<P>\<close> assms(2) learn_consE[of \<psi> U s] learn_consI[of \<psi> U t] by simp
+ with \<open>total1 t\<close> show ?thesis by auto
qed
text \<open>The proof of @{prop "U\<^bsub>LIM-CONS\<^esub> \<notin> CONS"} is by contradiction.
Assume there is a consistent learning strategy $S$. By the previous
lemma $S$ can be assumed to be total. Moreover it outputs a consistent
hypothesis for every prefix. Thus for every $e \in \mathbb{N}^+$, $S(e) \neq
S(e0)$ or $S(e) \neq S(e1)$ because $S(e)$ cannot be consistent with both
$e0$ and $e1$. We use this property of $S$ to construct a function in @{term
"U\<^bsub>LIM-CONS\<^esub>"} for which $S$ fails as a learning strategy. To
this end we define a numbering $\psi \in \mathcal{R}^2$ with $\psi_i(0) = i$
and
\[
\psi_i(x + 1) = \left\{\begin{array}{ll}
0 & \mbox{if } S(\psi_i^x0) \neq S(\psi_i^x),\\
1 & \mbox{otherwise}.
\end{array}\right.
\]
This numbering is recursive because $S$ is total. The ``otherwise'' case is
equivalent to $S(\psi_i^x1) \neq S(\psi_i^x)$ because $S(\psi_i^x)$ cannot be
consistent with both $\psi_i^x0$ and $\psi_i^x1$. Therefore every prefix
$\psi_i^x$ is extended in such a way that $S$ changes its hypothesis. Hence
$S$ does not learn $\psi_i$ in the limit. Kleene's fixed-point theorem
ensures that for some $j \geq 2$, $\varphi_j = \psi_j$. This $\psi_j$ is the
sought function in @{term "U\<^bsub>LIM-CONS\<^esub>"}.
The following locale formalizes the construction of $\psi$ for a total
strategy $S$.\<close>
locale cons_lim =
fixes s :: partial1
assumes s_in_R1: "s \<in> \<R>"
begin
text \<open>A @{typ recf} computing the strategy:\<close>
definition r_s :: recf where
"r_s \<equiv> SOME r_s. recfn 1 r_s \<and> total r_s \<and> s = (\<lambda>x. eval r_s [x])"
lemma r_s_recfn [simp]: "recfn 1 r_s"
and r_s_total [simp]: "\<And>x. eval r_s [x] \<down>"
and eval_r_s: "s = (\<lambda>x. eval r_s [x])"
using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all
text \<open>The next function represents the prefixes of $\psi_i$.\<close>
fun prefixes :: "nat \<Rightarrow> nat \<Rightarrow> nat list" where
"prefixes i 0 = [i]"
| "prefixes i (Suc x) = (prefixes i x) @
[if s (e_snoc (list_encode (prefixes i x)) 0) = s (list_encode (prefixes i x))
then 1 else 0]"
definition "r_prefixes_aux \<equiv>
Cn 3 r_ifeq
[Cn 3 r_s [Cn 3 r_snoc [Id 3 1, r_constn 2 0]],
Cn 3 r_s [Id 3 1],
Cn 3 r_snoc [Id 3 1, r_constn 2 1],
Cn 3 r_snoc [Id 3 1, r_constn 2 0]]"
lemma r_prefixes_aux_recfn: "recfn 3 r_prefixes_aux"
unfolding r_prefixes_aux_def by simp
lemma r_prefixes_aux:
"eval r_prefixes_aux [j, v, i] \<down>=
e_snoc v (if eval r_s [e_snoc v 0] = eval r_s [v] then 1 else 0)"
unfolding r_prefixes_aux_def by auto
definition "r_prefixes \<equiv> r_swap (Pr 1 r_singleton_encode r_prefixes_aux)"
lemma r_prefixes_recfn: "recfn 2 r_prefixes"
unfolding r_prefixes_def r_prefixes_aux_def by simp
lemma r_prefixes: "eval r_prefixes [i, n] \<down>= list_encode (prefixes i n)"
proof -
let ?h = "Pr 1 r_singleton_encode r_prefixes_aux"
have "eval ?h [n, i] \<down>= list_encode (prefixes i n)"
proof (induction n)
case 0
then show ?case
using r_prefixes_def r_prefixes_aux_recfn r_singleton_encode by simp
next
case (Suc n)
then show ?case
using r_prefixes_aux_recfn r_prefixes_aux eval_r_s
by auto metis+
qed
moreover have "eval ?h [n, i] = eval r_prefixes [i, n]" for i n
unfolding r_prefixes_def by (simp add: r_prefixes_aux_recfn)
ultimately show ?thesis by simp
qed
lemma prefixes_neq_nil: "length (prefixes i x) > 0"
by (induction x) auto
text \<open>The actual numbering can then be defined via @{term prefixes}.\<close>
definition psi :: "partial2" ("\<psi>") where
"\<psi> i x \<equiv> Some (last (prefixes i x))"
lemma psi_in_R2: "\<psi> \<in> \<R>\<^sup>2"
proof
define r_psi where "r_psi \<equiv> Cn 2 r_last [r_prefixes]"
have "recfn 2 r_psi"
unfolding r_psi_def by (simp add: r_prefixes_recfn)
then have "eval r_psi [i, n] \<down>= last (prefixes i n)" for n i
unfolding r_psi_def using r_prefixes r_prefixes_recfn prefixes_neq_nil by simp
then have "(\<lambda>i x. Some (last (prefixes i x))) \<in> \<P>\<^sup>2"
- using `recfn 2 r_psi` P2I[of "r_psi"] by simp
+ using \<open>recfn 2 r_psi\<close> P2I[of "r_psi"] by simp
with psi_def show "\<psi> \<in> \<P>\<^sup>2" by presburger
moreover show "total2 psi"
unfolding psi_def by auto
qed
lemma psi_0_or_1:
assumes "n > 0"
shows "\<psi> i n \<down>= 0 \<or> \<psi> i n \<down>= 1"
proof -
from assms obtain m where "n = Suc m"
using gr0_implies_Suc by blast
then have "last (prefixes i (Suc m)) = 0 \<or> last (prefixes i (Suc m)) = 1"
by simp
then show ?thesis using \<open>n = Suc m\<close> psi_def by simp
qed
text \<open>The function @{term "prefixes"} does indeed provide the prefixes
for @{term "\<psi>"}.\<close>
lemma psi_init: "(\<psi> i) \<triangleright> x = list_encode (prefixes i x)"
proof -
have "prefix (\<psi> i) x = prefixes i x"
unfolding psi_def
by (induction x) (simp_all add: prefix_0 prefix_Suc)
with init_def show ?thesis by simp
qed
text \<open>One of the functions $\psi_i$ is in @{term "U\<^bsub>LIM-CONS\<^esub>"}.\<close>
lemma ex_psi_in_U: "\<exists>j. \<psi> j \<in> U\<^bsub>LIM-CONS\<^esub>"
proof -
obtain j where j: "j \<ge> 2" "\<psi> j = \<phi> j"
using kleene_fixed_point[of \<psi>] psi_in_R2 R2_imp_P2 by metis
then have "\<psi> j \<in> \<P>" by (simp add: phi_in_P2)
define p where "p = (\<lambda>x. \<psi> j (x + 1))"
have "p \<in> \<R>\<^sub>0\<^sub>1"
proof -
- from p_def `\<psi> j \<in> \<P>` skip_P1 have "p \<in> \<P>" by blast
+ from p_def \<open>\<psi> j \<in> \<P>\<close> skip_P1 have "p \<in> \<P>" by blast
from psi_in_R2 have "total1 (\<psi> j)" by simp
with p_def have "total1 p"
by (simp add: total1_def)
with psi_0_or_1 have "p n \<down>= 0 \<or> p n \<down>= 1" for n
using psi_def p_def by simp
then show ?thesis
by (simp add: RPred1_def P1_total_imp_R1 \<open>p \<in> \<P>\<close> \<open>total1 p\<close>)
qed
moreover have "\<psi> j = [j] \<odot> p"
proof
fix x
show "\<psi> j x = ([j] \<odot> p) x"
proof (cases "x = 0")
case True
then show ?thesis using psi_def psi_def prepend_at_less by simp
next
case False
then show ?thesis using p_def by simp
qed
qed
ultimately have "\<psi> j \<in> U\<^bsub>LIM-CONS\<^esub>"
using j U_LIMCONS_def by (metis (mono_tags, lifting) append_Nil mem_Collect_eq)
then show ?thesis by auto
qed
text \<open>The strategy fails to learn @{term U_LIMCONS} because it changes
its hypothesis all the time on functions $\psi_j \in V_0$.\<close>
lemma U_LIMCONS_not_learn_cons: "\<not> learn_cons \<phi> U\<^bsub>LIM-CONS\<^esub> s"
proof
assume learn: "learn_cons \<phi> U\<^bsub>LIM-CONS\<^esub> s"
have "s (list_encode (vs @ [0])) \<noteq> s (list_encode (vs @ [1]))" for vs
proof -
obtain f\<^sub>0 where f0: "f\<^sub>0 \<in> U\<^bsub>LIM-CONS\<^esub>" "prefix f\<^sub>0 (length vs) = vs @ [0]"
using U_LIMCONS_prefix_complete[of "vs @ [0]"] by auto
obtain f\<^sub>1 where f1: "f\<^sub>1 \<in> U\<^bsub>LIM-CONS\<^esub>" "prefix f\<^sub>1 (length vs) = vs @ [1]"
using U_LIMCONS_prefix_complete[of "vs @ [1]"] by auto
have "f\<^sub>0 (length vs) \<noteq> f\<^sub>1 (length vs)"
using f0 f1 by (metis lessI nth_append_length prefix_nth zero_neq_one)
moreover have "\<phi> (the (s (f\<^sub>0 \<triangleright> length vs))) (length vs) = f\<^sub>0 (length vs)"
using learn_consE(3)[of \<phi> U_LIMCONS s, OF learn, of f\<^sub>0 "length vs", OF f0(1)]
by simp
moreover have "\<phi> (the (s (f\<^sub>1 \<triangleright> length vs))) (length vs) = f\<^sub>1 (length vs)"
using learn_consE(3)[of \<phi> U_LIMCONS s, OF learn, of f\<^sub>1 "length vs", OF f1(1)]
by simp
ultimately have "the (s (f\<^sub>0 \<triangleright> length vs)) \<noteq> the (s (f\<^sub>1 \<triangleright> length vs))"
by auto
then have "s (f\<^sub>0 \<triangleright> length vs) \<noteq> s (f\<^sub>1 \<triangleright> length vs)"
by auto
with f0(2) f1(2) show ?thesis by (simp add: init_def)
qed
then have "s (list_encode (vs @ [0])) \<noteq> s (list_encode vs) \<or>
s (list_encode (vs @ [1])) \<noteq> s (list_encode vs)"
for vs
by metis
then have "s (list_encode (prefixes i (Suc x))) \<noteq> s (list_encode (prefixes i x))" for i x
by simp
then have "\<not> learn_lim \<phi> {\<psi> i} s" for i
using psi_def psi_init always_hyp_change_not_Lim by simp
then have "\<not> learn_lim \<phi> U_LIMCONS s"
using ex_psi_in_U learn_lim_closed_subseteq by blast
then show False
using learn learn_cons_def by simp
qed
end
text \<open>With the locale we can now show the second part of the main
result:\<close>
lemma U_LIMCONS_not_in_CONS: "U\<^bsub>LIM-CONS\<^esub> \<notin> CONS"
proof
assume "U\<^bsub>LIM-CONS\<^esub> \<in> CONS"
then have "U\<^bsub>LIM-CONS\<^esub> \<in> CONS_wrt \<phi>"
by (simp add: CONS_wrt_phi_eq_CONS)
then obtain almost_s where "learn_cons \<phi> U\<^bsub>LIM-CONS\<^esub> almost_s"
using CONS_wrt_def by auto
then obtain s where s: "total1 s" "learn_cons \<phi> U\<^bsub>LIM-CONS\<^esub> s"
using U_LIMCONS_prefix_complete U_prefix_complete_imp_total_strategy by blast
then have "s \<in> \<R>"
using learn_consE(1) P1_total_imp_R1 by blast
with cons_lim_def interpret cons_lim s by simp
show False
using s(2) U_LIMCONS_not_learn_cons by simp
qed
text \<open>The main result of this section:\<close>
theorem CONS_subset_Lim: "CONS \<subset> LIM"
using U_LIMCONS_in_Lim U_LIMCONS_not_in_CONS CONS_subseteq_Lim by auto
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/CP_FIN_NUM.thy b/thys/Inductive_Inference/CP_FIN_NUM.thy
--- a/thys/Inductive_Inference/CP_FIN_NUM.thy
+++ b/thys/Inductive_Inference/CP_FIN_NUM.thy
@@ -1,1032 +1,1032 @@
section \<open>FIN is a proper subset of CP\label{s:fin_cp}\<close>
theory CP_FIN_NUM
imports Inductive_Inference_Basics
begin
text \<open>Let $S$ be a FIN strategy for a non-empty class $U$. Let $T$ be a
strategy that hypothesizes an arbitrary function from $U$ while $S$ outputs
``don't know'' and the hypothesis of $S$ otherwise. Then $T$ is a CP strategy
for $U$.\<close>
lemma nonempty_FIN_wrt_impl_CP:
assumes "U \<noteq> {}" and "U \<in> FIN_wrt \<psi>"
shows "U \<in> CP_wrt \<psi>"
proof -
obtain s where "learn_fin \<psi> U s"
using assms(2) FIN_wrt_def by auto
then have env: "environment \<psi> U s" and
fin: "\<And>f. f \<in> U \<Longrightarrow>
\<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
using learn_finE by auto
from assms(1) obtain f\<^sub>0 where "f\<^sub>0 \<in> U"
by auto
with fin obtain i\<^sub>0 where "\<psi> i\<^sub>0 = f\<^sub>0"
by blast
define t where "t x \<equiv>
(if s x \<up> then None else if s x \<down>= 0 then Some i\<^sub>0 else Some (the (s x) - 1))"
for x
have "t \<in> \<P>"
proof -
from env obtain rs where rs: "recfn 1 rs" "\<And>x. eval rs [x] = s x"
by auto
define rt where "rt = Cn 1 r_ifz [rs, r_const i\<^sub>0, Cn 1 r_dec [rs]]"
then have "recfn 1 rt"
using rs(1) by simp
then have "eval rt [x] \<down>= (if s x \<down>= 0 then i\<^sub>0 else (the (s x)) - 1)" if "s x \<down>" for x
using rs rt_def that by auto
moreover have "eval rt [x] \<up>" if "eval rs [x] \<up>" for x
using rs rt_def that by simp
ultimately have "eval rt [x] = t x" for x
using rs(2) t_def by simp
with \<open>recfn 1 rt\<close> show ?thesis by auto
qed
have "learn_cp \<psi> U t"
proof (rule learn_cpI)
show "environment \<psi> U t"
using env t_def \<open>t \<in> \<P>\<close> by simp
show "\<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. t (f \<triangleright> n) \<down>= i)" if "f \<in> U" for f
proof -
from that fin obtain i n\<^sub>0 where
i: "\<psi> i = f" "\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0" "\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i"
by blast
moreover have "\<forall>n\<ge>n\<^sub>0. t (f \<triangleright> n) \<down>= i"
using that t_def i(3) by simp
ultimately show ?thesis by auto
qed
show "\<psi> (the (t (f \<triangleright> n))) \<in> U" if "f \<in> U" for f n
using \<open>\<psi> i\<^sub>0 = f\<^sub>0\<close> \<open>f\<^sub>0 \<in> U\<close> t_def fin env that
by (metis (no_types, lifting) diff_Suc_1 not_less option.sel)
qed
then show ?thesis using CP_wrt_def env by auto
qed
lemma FIN_wrt_impl_CP:
assumes "U \<in> FIN_wrt \<psi>"
shows "U \<in> CP_wrt \<psi>"
proof (cases "U = {}")
case True
then have "\<psi> \<in> \<P>\<^sup>2 \<Longrightarrow> U \<in> CP_wrt \<psi>"
using CP_wrt_def learn_cpI[of \<psi> "{}" "\<lambda>x. Some 0"] const_in_Prim1 by auto
moreover have "\<psi> \<in> \<P>\<^sup>2"
using assms FIN_wrt_def learn_finE by auto
ultimately show "U \<in> CP_wrt \<psi>" by simp
next
case False
with nonempty_FIN_wrt_impl_CP assms show ?thesis
by simp
qed
corollary FIN_subseteq_CP: "FIN \<subseteq> CP"
proof
fix U
assume "U \<in> FIN"
then have "\<exists>\<psi>. U \<in> FIN_wrt \<psi>"
using FIN_def FIN_wrt_def by auto
then have "\<exists>\<psi>. U \<in> CP_wrt \<psi>"
using FIN_wrt_impl_CP by auto
then show "U \<in> CP"
by (simp add: CP_def CP_wrt_def)
qed
text \<open>In order to show the \emph{proper} inclusion, we show @{term
"U\<^sub>0 \<in> CP - FIN"}. A CP strategy for @{term "U\<^sub>0"} simply
hypothesizes the function in @{term U0} with the longest prefix of $f^n$ not
ending in zero. For that we define a function computing the index of the
rightmost non-zero value in a list, returning the length of the list if there
is no such value.\<close>
definition findr :: partial1 where
"findr e \<equiv>
if \<exists>i<e_length e. e_nth e i \<noteq> 0
then Some (GREATEST i. i < e_length e \<and> e_nth e i \<noteq> 0)
else Some (e_length e)"
lemma findr_total: "findr e \<down>"
unfolding findr_def by simp
lemma findr_ex:
assumes "\<exists>i<e_length e. e_nth e i \<noteq> 0"
shows "the (findr e) < e_length e"
and "e_nth e (the (findr e)) \<noteq> 0"
and "\<forall>i. the (findr e) < i \<and> i < e_length e \<longrightarrow> e_nth e i = 0"
proof -
let ?P = "\<lambda>i. i < e_length e \<and> e_nth e i \<noteq> 0"
from assms have "\<exists>i. ?P i" by simp
then have "?P (Greatest ?P)"
using GreatestI_ex_nat[of ?P "e_length e"] by fastforce
moreover have *: "findr e = Some (Greatest ?P)"
using assms findr_def by simp
ultimately show "the (findr e) < e_length e" and "e_nth e (the (findr e)) \<noteq> 0"
by fastforce+
show "\<forall>i. the (findr e) < i \<and> i < e_length e \<longrightarrow> e_nth e i = 0"
using * Greatest_le_nat[of ?P _ "e_length e"] by fastforce
qed
definition "r_findr \<equiv>
let g =
Cn 3 r_ifz
[Cn 3 r_nth [Id 3 2, Id 3 0],
Cn 3 r_ifeq [Id 3 0, Id 3 1, Cn 3 S [Id 3 0], Id 3 1],
Id 3 0]
in Cn 1 (Pr 1 Z g) [Cn 1 r_length [Id 1 0], Id 1 0]"
lemma r_findr_prim [simp]: "prim_recfn 1 r_findr"
unfolding r_findr_def by simp
lemma r_findr [simp]: "eval r_findr [e] = findr e"
proof -
define g where "g =
Cn 3 r_ifz
[Cn 3 r_nth [Id 3 2, Id 3 0],
Cn 3 r_ifeq [Id 3 0, Id 3 1, Cn 3 S [Id 3 0], Id 3 1],
Id 3 0]"
then have "recfn 3 g"
by simp
with g_def have g: "eval g [j, r, e] \<down>=
(if e_nth e j \<noteq> 0 then j else if j = r then Suc j else r)" for j r e
by simp
let ?h = "Pr 1 Z g"
have "recfn 2 ?h"
by (simp add: \<open>recfn 3 g\<close>)
let ?P = "\<lambda>e j i. i < j \<and> e_nth e i \<noteq> 0"
let ?G = "\<lambda>e j. Greatest (?P e j)"
have h: "eval ?h [j, e] =
(if \<forall>i<j. e_nth e i = 0 then Some j else Some (?G e j))" for j e
proof (induction j)
case 0
then show ?case using \<open>recfn 2 ?h\<close> by auto
next
case (Suc j)
then have "eval ?h [Suc j, e] = eval g [j, the (eval ?h [j, e]), e]"
using \<open>recfn 2 ?h\<close> by auto
then have "eval ?h [Suc j, e] =
eval g [j, if \<forall>i<j. e_nth e i = 0 then j else ?G e j, e]"
using Suc by auto
then have *: "eval ?h [Suc j, e] \<down>=
(if e_nth e j \<noteq> 0 then j
else if j = (if \<forall>i<j. e_nth e i = 0 then j else ?G e j)
then Suc j
else (if \<forall>i<j. e_nth e i = 0 then j else ?G e j))"
using g by simp
show ?case
proof (cases "\<forall>i<Suc j. e_nth e i = 0")
case True
then show ?thesis using * by simp
next
case False
then have ex: "\<exists>i<Suc j. e_nth e i \<noteq> 0"
by auto
show ?thesis
proof (cases "e_nth e j = 0")
case True
then have ex': "\<exists>i<j. e_nth e i \<noteq> 0"
using ex less_Suc_eq by fastforce
then have "(if \<forall>i<j. e_nth e i = 0 then j else ?G e j) = ?G e j"
by metis
moreover have "?G e j < j"
using ex' GreatestI_nat[of "?P e j"] less_imp_le_nat by blast
ultimately have "eval ?h [Suc j, e] \<down>= ?G e j"
using * True by simp
moreover have "?G e j = ?G e (Suc j)"
using True by (metis less_SucI less_Suc_eq)
ultimately show ?thesis using ex by metis
next
case False
then have "eval ?h [Suc j, e] \<down>= j"
using * by simp
moreover have "?G e (Suc j) = j"
using ex False Greatest_equality[of "?P e (Suc j)"] by simp
ultimately show ?thesis using ex by simp
qed
qed
qed
let ?hh = "Cn 1 ?h [Cn 1 r_length [Id 1 0], Id 1 0]"
have "recfn 1 ?hh"
- using `recfn 2 ?h` by simp
+ using \<open>recfn 2 ?h\<close> by simp
with h have hh: "eval ?hh [e] \<down>=
(if \<forall>i<e_length e. e_nth e i = 0 then e_length e else ?G e (e_length e))" for e
by auto
then have "eval ?hh [e] = findr e" for e
unfolding findr_def by auto
moreover have "total ?hh"
- using hh totalI1 `recfn 1 ?hh` by simp
+ using hh totalI1 \<open>recfn 1 ?hh\<close> by simp
ultimately show ?thesis
- using `recfn 1 ?hh` g_def r_findr_def findr_def by metis
+ using \<open>recfn 1 ?hh\<close> g_def r_findr_def findr_def by metis
qed
lemma U0_in_CP: "U\<^sub>0 \<in> CP"
proof -
define s where
"s \<equiv> \<lambda>x. if findr x \<down>= e_length x then Some 0 else Some (e_take (Suc (the (findr x))) x)"
have "s \<in> \<P>"
proof -
define r where
"r \<equiv> Cn 1 r_ifeq [r_findr, r_length, Z, Cn 1 r_take [Cn 1 S [r_findr], Id 1 0]]"
then have "\<And>x. eval r [x] = s x"
using s_def findr_total by fastforce
moreover have "recfn 1 r"
using r_def by simp
ultimately show ?thesis by auto
qed
moreover have "learn_cp prenum U\<^sub>0 s"
proof (rule learn_cpI)
show "environment prenum U\<^sub>0 s"
using \<open>s \<in> \<P>\<close> s_def prenum_in_R2 U0_in_NUM by auto
show "\<exists>i. prenum i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)" if "f \<in> U\<^sub>0" for f
proof (cases "f = (\<lambda>_. Some 0)")
case True
then have "s (f \<triangleright> n) \<down>= 0" for n
using findr_def s_def by simp
then have "\<forall>n\<ge>0. s (f \<triangleright> n) \<down>= 0" by simp
moreover have "prenum 0 = f"
using True by auto
ultimately show ?thesis by auto
next
case False
then obtain ws where ws: "length ws > 0" "last ws \<noteq> 0" "f = ws \<odot> 0\<^sup>\<infinity>"
using U0_def \<open>f \<in> U\<^sub>0\<close> almost0_canonical by blast
let ?m = "length ws - 1"
let ?i = "list_encode ws"
have "prenum ?i = f"
using ws by auto
moreover have "s (f \<triangleright> n) \<down>= ?i" if "n \<ge> ?m" for n
proof -
have "e_nth (f \<triangleright> n) ?m \<noteq> 0"
using ws that by (simp add: last_conv_nth)
then have "\<exists>k<Suc n. e_nth (f \<triangleright> n) k \<noteq> 0"
using le_imp_less_Suc that by blast
moreover have
"(GREATEST k. k < e_length (f \<triangleright> n) \<and> e_nth (f \<triangleright> n) k \<noteq> 0) = ?m"
proof (rule Greatest_equality)
show "?m < e_length (f \<triangleright> n) \<and> e_nth (f \<triangleright> n) ?m \<noteq> 0"
using \<open>e_nth (f \<triangleright> n) ?m \<noteq> 0\<close> that by auto
show "\<And>y. y < e_length (f \<triangleright> n) \<and> e_nth (f \<triangleright> n) y \<noteq> 0 \<Longrightarrow> y \<le> ?m"
using ws less_Suc_eq_le by fastforce
qed
ultimately have "findr (f \<triangleright> n) \<down>= ?m"
using that findr_def by simp
moreover have "?m < e_length (f \<triangleright> n)"
using that by simp
ultimately have "s (f \<triangleright> n) \<down>= e_take (Suc ?m) (f \<triangleright> n)"
using s_def by simp
moreover have "e_take (Suc ?m) (f \<triangleright> n) = list_encode ws"
proof -
have "take (Suc ?m) (prefix f n) = prefix f ?m"
using take_prefix[of f ?m n] ws that by (simp add: almost0_in_R1)
then have "take (Suc ?m) (prefix f n) = ws"
using ws prefixI by auto
then show ?thesis by simp
qed
ultimately show ?thesis by simp
qed
ultimately show ?thesis by auto
qed
show "\<And>f n. f \<in> U\<^sub>0 \<Longrightarrow> prenum (the (s (f \<triangleright> n))) \<in> U\<^sub>0"
using U0_def by fastforce
qed
ultimately show ?thesis using CP_def by blast
qed
text \<open>As a bit of an interlude, we can now show that CP is not
closed under the subset relation. This works by removing functions from
@{term "U\<^sub>0"} in a ``noncomputable'' way such that a strategy cannot ensure
that every intermediate hypothesis is in that new class.\<close>
lemma CP_not_closed_subseteq: "\<exists>V U. V \<subseteq> U \<and> U \<in> CP \<and> V \<notin> CP"
proof -
\<comment> \<open>The numbering $g\in\mathcal{R}^2$ enumerates all
functions $i0^\infty \in U_0$.\<close>
define g where "g \<equiv> \<lambda>i. [i] \<odot> 0\<^sup>\<infinity>"
have g_inj: "i = j" if "g i = g j" for i j
proof -
have "g i 0 \<down>= i" and "g j 0 \<down>= j"
by (simp_all add: g_def)
with that show "i = j"
by (metis option.inject)
qed
\<comment> \<open>Define a class $V$. If the strategy $\varphi_i$ learns
$g_i$, it outputs a hypothesis for $g_i$ on some shortest prefix $g_i^m$.
Then the function $g_i^m10^\infty$ is included in the class $V$; otherwise
$g_i$ is included.\<close>
define V where "V \<equiv>
{if learn_lim \<phi> {g i} (\<phi> i)
then (prefix (g i) (LEAST n. \<phi> (the (\<phi> i ((g i) \<triangleright> n))) = g i)) @ [1] \<odot> 0\<^sup>\<infinity>
else g i |
i. i \<in> UNIV}"
have "V \<notin> CP_wrt \<phi>"
proof
\<comment> \<open>Assuming $V \in CP_\varphi$, there is a CP strategy
$\varphi_i$ for $V$.\<close>
assume "V \<in> CP_wrt \<phi>"
then obtain s where s: "s \<in> \<P>" "learn_cp \<phi> V s"
using CP_wrt_def learn_cpE(1) by auto
then obtain i where i: "\<phi> i = s"
using phi_universal by auto
show False
proof (cases "learn_lim \<phi> {g i} (\<phi> i)")
case learn: True
\<comment> \<open>If $\varphi_i$ learns $g_i$, it hypothesizes $g_i$ on
some shortest prefix $g_i^m$. Thus it hypothesizes $g_i$ on some prefix
of $g_i^m10^\infty \in V$, too. But $g_i$ is not a class-preserving
hypothesis because $g_i \notin V$.\<close>
let ?P = "\<lambda>n. \<phi> (the (\<phi> i ((g i) \<triangleright> n))) = g i"
let ?m = "Least ?P"
have "\<exists>n. ?P n"
using i s by (meson learn infinite_hyp_wrong_not_Lim insertI1 lessI)
then have "?P ?m"
using LeastI_ex[of ?P] by simp
define h where "h = (prefix (g i) ?m) @ [1] \<odot> 0\<^sup>\<infinity>"
then have "h \<in> V"
using V_def learn by auto
have "(g i) \<triangleright> ?m = h \<triangleright> ?m"
proof -
have "prefix (g i) ?m = prefix h ?m"
unfolding h_def by (simp add: prefix_prepend_less)
then show ?thesis by auto
qed
then have "\<phi> (the (\<phi> i (h \<triangleright> ?m))) = g i"
- using `?P ?m` by simp
+ using \<open>?P ?m\<close> by simp
moreover have "g i \<notin> V"
proof
assume "g i \<in> V"
then obtain j where j: "g i =
(if learn_lim \<phi> {g j} (\<phi> j)
then (prefix (g j) (LEAST n. \<phi> (the (\<phi> j ((g j) \<triangleright> n))) = g j)) @ [1] \<odot> 0\<^sup>\<infinity>
else g j)"
using V_def by auto
show False
proof (cases "learn_lim \<phi> {g j} (\<phi> j)")
case True
then have "g i =
(prefix (g j) (LEAST n. \<phi> (the (\<phi> j ((g j) \<triangleright> n))) = g j)) @ [1] \<odot> 0\<^sup>\<infinity>"
(is "g i = ?vs @ [1] \<odot> 0\<^sup>\<infinity>")
using j by simp
moreover have len: "length ?vs > 0" by simp
ultimately have "g i (length ?vs) \<down>= 1"
by (simp add: prepend_associative)
moreover have "g i (length ?vs) \<down>= 0"
using g_def len by simp
ultimately show ?thesis by simp
next
case False
then show ?thesis
using j g_inj learn by auto
qed
qed
ultimately have "\<phi> (the (\<phi> i (h \<triangleright> ?m))) \<notin> V" by simp
then have "\<not> learn_cp \<phi> V (\<phi> i)"
- using `h \<in> V` learn_cpE(3) by auto
+ using \<open>h \<in> V\<close> learn_cpE(3) by auto
then show ?thesis by (simp add: i s(2))
next
\<comment> \<open>If $\varphi_i$ does not learn $g_i$, then $g_i\in V$.
Hence $\varphi_i$ does not learn $V$.\<close>
case False
then have "g i \<in> V"
using V_def by auto
with False have "\<not> learn_lim \<phi> V (\<phi> i)"
using learn_lim_closed_subseteq by auto
then show ?thesis
using s(2) i by (simp add: learn_cp_def)
qed
qed
then have "V \<notin> CP"
using CP_wrt_phi by simp
moreover have "V \<subseteq> U\<^sub>0"
using V_def g_def U0_def by auto
ultimately show ?thesis using U0_in_CP by auto
qed
text \<open>Continuing with the main result of this section, we show that
@{term "U\<^sub>0"} cannot be learned finitely. Any FIN strategy would have
to output a hypothesis for the constant zero function on some prefix. But
@{term "U\<^sub>0"} contains infinitely many other functions starting with
the same prefix, which the strategy then would not learn finitely.\<close>
lemma U0_not_in_FIN: "U\<^sub>0 \<notin> FIN"
proof
assume "U\<^sub>0 \<in> FIN"
then obtain \<psi> s where "learn_fin \<psi> U\<^sub>0 s"
using FIN_def by blast
with learn_finE have cp: "\<And>f. f \<in> U\<^sub>0 \<Longrightarrow>
\<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
by simp_all
define z where "z = [] \<odot> 0\<^sup>\<infinity>"
then have "z \<in> U\<^sub>0"
using U0_def by auto
with cp obtain i n\<^sub>0 where i: "\<psi> i = z" and n0: "\<forall>n\<ge>n\<^sub>0. s (z \<triangleright> n) \<down>= Suc i"
by blast
define w where "w = replicate (Suc n\<^sub>0) 0 @ [1] \<odot> 0\<^sup>\<infinity>"
then have "prefix w n\<^sub>0 = replicate (Suc n\<^sub>0) 0"
by (simp add: prefix_prepend_less)
moreover have "prefix z n\<^sub>0 = replicate (Suc n\<^sub>0) 0"
using prefixI[of "replicate (Suc n\<^sub>0) 0" z] less_Suc_eq_0_disj unfolding z_def
by fastforce
ultimately have "z \<triangleright> n\<^sub>0 = w \<triangleright> n\<^sub>0"
by (simp add: init_prefixE)
with n0 have *: "s (w \<triangleright> n\<^sub>0) \<down>= Suc i" by auto
have "w \<in> U\<^sub>0" using w_def U0_def by auto
with cp obtain i' n\<^sub>0' where i': "\<psi> i' = w"
and n0': "\<forall>n<n\<^sub>0'. s (w \<triangleright> n) \<down>= 0" "\<forall>n\<ge>n\<^sub>0'. s (w \<triangleright> n) \<down>= Suc i'"
by blast
have "i \<noteq> i'"
proof
assume "i = i'"
then have "w = z"
using i i' by simp
have "w (Suc n\<^sub>0) \<down>= 1"
using w_def prepend[of "replicate (Suc n\<^sub>0) 0 @ [1]" "0\<^sup>\<infinity>" "Suc n\<^sub>0"]
by (metis length_append_singleton length_replicate lessI nth_append_length)
moreover have "z (Suc n\<^sub>0) \<down>= 0"
using z_def by simp
ultimately show False
using \<open>w = z\<close> by simp
qed
then have "s (w \<triangleright> n\<^sub>0) \<down>\<noteq> Suc i"
using n0' by (cases "n\<^sub>0 < n\<^sub>0'") simp_all
with * show False by simp
qed
theorem FIN_subset_CP: "FIN \<subset> CP"
using U0_in_CP U0_not_in_FIN FIN_subseteq_CP by auto
section \<open>NUM and FIN are incomparable\label{s:num_fin}\<close>
text \<open>The class $V_0$ of all total recursive functions $f$ where $f(0)$
is a Gödel number of $f$ can be learned finitely by always hypothesizing
$f(0)$. The class is not in NUM and therefore serves to separate NUM and
FIN.\<close>
definition V0 :: "partial1 set" ("V\<^sub>0") where
"V\<^sub>0 = {f. f \<in> \<R> \<and> \<phi> (the (f 0)) = f}"
lemma V0_altdef: "V\<^sub>0 = {[i] \<odot> f| i f. f \<in> \<R> \<and> \<phi> i = [i] \<odot> f}"
(is "V\<^sub>0 = ?W")
proof
show "V\<^sub>0 \<subseteq> ?W"
proof
fix f
assume "f \<in> V\<^sub>0"
then have "f \<in> \<R>"
unfolding V0_def by simp
then obtain i where i: "f 0 \<down>= i" by fastforce
define g where "g = (\<lambda>x. f (x + 1))"
then have "g \<in> \<R>"
- using skip_R1[OF `f \<in> \<R>`] by blast
+ using skip_R1[OF \<open>f \<in> \<R>\<close>] by blast
moreover have "[i] \<odot> g = f"
using g_def i by auto
moreover have "\<phi> i = f"
- using `f \<in> V\<^sub>0` V0_def i by force
+ using \<open>f \<in> V\<^sub>0\<close> V0_def i by force
ultimately show "f \<in> ?W" by auto
qed
show "?W \<subseteq> V\<^sub>0"
proof
fix g
assume "g \<in> ?W"
then have "\<phi> (the (g 0)) = g" by auto
moreover have "g \<in> \<R>"
- using prepend_in_R1 `g \<in> ?W` by auto
+ using prepend_in_R1 \<open>g \<in> ?W\<close> by auto
ultimately show "g \<in> V\<^sub>0"
by (simp add: V0_def)
qed
qed
lemma V0_in_FIN: "V\<^sub>0 \<in> FIN"
proof -
define s where "s = (\<lambda>x. Some (Suc (e_hd x)))"
have "s \<in> \<P>"
proof -
define r where "r = Cn 1 S [r_hd]"
then have "recfn 1 r" by simp
moreover have "eval r [x] \<down>= Suc (e_hd x)" for x
unfolding r_def by simp
ultimately show ?thesis
using s_def by blast
qed
have s: "s (f \<triangleright> n) \<down>= Suc (the (f 0))" for f n
unfolding s_def by simp
have "learn_fin \<phi> V\<^sub>0 s"
proof (rule learn_finI)
show "environment \<phi> V\<^sub>0 s"
using s_def \<open>s \<in> \<P>\<close> phi_in_P2 V0_def by auto
show "\<exists>i n\<^sub>0. \<phi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
if "f \<in> V\<^sub>0" for f
using that V0_def s by auto
qed
then show ?thesis using FIN_def by auto
qed
text \<open>To every @{term "f \<in> \<R>"} a number can be prepended that is
a Gödel number of the resulting function. Such a function is then in $V_0$.
If $V_0$ was in NUM, it would be embedded in a total numbering. Shifting this
numbering to the left, essentially discarding the values at point $0$, would
yield a total numbering for @{term "\<R>"}, which contradicts @{thm[source]
R1_not_in_NUM}. This proves @{prop "V\<^sub>0 \<notin> NUM"}.\<close>
lemma prepend_goedel:
assumes "f \<in> \<R>"
shows "\<exists>i. \<phi> i = [i] \<odot> f"
proof -
obtain r where r: "recfn 1 r" "total r" "\<And>x. eval r [x] = f x"
using assms by auto
define r_psi where "r_psi = Cn 2 r_ifz [Id 2 1, Id 2 0, Cn 2 r [Cn 2 r_dec [Id 2 1]]]"
then have "recfn 2 r_psi"
using r(1) by simp
have "eval r_psi [i, x] = (if x = 0 then Some i else f (x - 1))" for i x
proof -
have "eval (Cn 2 r [Cn 2 r_dec [Id 2 1]]) [i, x] = f (x - 1)"
using r by simp
then have "eval r_psi [i, x] = eval r_ifz [x, i, the (f (x - 1))]"
unfolding r_psi_def using \<open>recfn 2 r_psi\<close> r R1_imp_total1[OF assms] by auto
then show ?thesis
using assms by simp
qed
with \<open>recfn 2 r_psi\<close> have "(\<lambda>i x. if x = 0 then Some i else f (x - 1)) \<in> \<P>\<^sup>2"
by auto
with kleene_fixed_point obtain i where
"\<phi> i = (\<lambda>x. if x = 0 then Some i else f (x - 1))"
by blast
then have "\<phi> i = [i] \<odot> f" by auto
then show ?thesis by auto
qed
lemma V0_in_FIN_minus_NUM: "V\<^sub>0 \<in> FIN - NUM"
proof -
have "V\<^sub>0 \<notin> NUM"
proof
assume "V\<^sub>0 \<in> NUM"
then obtain \<psi> where \<psi>: "\<psi> \<in> \<R>\<^sup>2" "\<And>f. f \<in> V\<^sub>0 \<Longrightarrow> \<exists>i. \<psi> i = f"
by auto
define \<psi>' where "\<psi>' i x = \<psi> i (Suc x)" for i x
have "\<psi>' \<in> \<R>\<^sup>2"
proof
from \<psi>(1) obtain r_psi where
r_psi: "recfn 2 r_psi" "total r_psi" "\<And>i x. eval r_psi [i, x] = \<psi> i x"
by blast
define r_psi' where "r_psi' = Cn 2 r_psi [Id 2 0, Cn 2 S [Id 2 1]]"
then have "recfn 2 r_psi'" and "\<And>i x. eval r_psi' [i, x] = \<psi>' i x"
unfolding r_psi'_def \<psi>'_def using r_psi by simp_all
then show "\<psi>' \<in> \<P>\<^sup>2" by blast
show "total2 \<psi>'"
using \<psi>'_def \<psi>(1) by (simp add: total2I)
qed
have "\<exists>i. \<psi>' i = f" if "f \<in> \<R>" for f
proof -
from that obtain j where j: "\<phi> j = [j] \<odot> f"
using prepend_goedel by auto
then have "\<phi> j \<in> V\<^sub>0"
using that V0_altdef by auto
with \<psi> obtain i where "\<psi> i = \<phi> j" by auto
then have "\<psi>' i = f"
using \<psi>'_def j by (auto simp add: prepend_at_ge)
then show ?thesis by auto
qed
with \<open>\<psi>' \<in> \<R>\<^sup>2\<close> have "\<R> \<in> NUM" by auto
with R1_not_in_NUM show False by simp
qed
then show ?thesis
using V0_in_FIN by auto
qed
corollary FIN_not_subseteq_NUM: "\<not> FIN \<subseteq> NUM"
using V0_in_FIN_minus_NUM by auto
section \<open>NUM and CP are incomparable\label{s:num_cp}\<close>
text \<open>There are FIN classes outside of NUM, and CP encompasses FIN.
Hence there are CP classes outside of NUM, too.\<close>
theorem CP_not_subseteq_NUM: "\<not> CP \<subseteq> NUM"
using FIN_subseteq_CP FIN_not_subseteq_NUM by blast
text \<open>Conversely there is a subclass of @{term "U\<^sub>0"} that
is in NUM but cannot be learned in a class-preserving way. The following
proof is due to Jantke and Beick~\cite{jb-cpnii-81}. The idea is to
diagonalize against all strategies, that is, all partial recursive
functions.\<close>
theorem NUM_not_subseteq_CP: "\<not> NUM \<subseteq> CP"
proof-
\<comment> \<open>Define a family of functions $f_k$.\<close>
define f where "f \<equiv> \<lambda>k. [k] \<odot> 0\<^sup>\<infinity>"
then have "f k \<in> \<R>" for k
using almost0_in_R1 by auto
\<comment> \<open>If the strategy $\varphi_k$ learns $f_k$ it hypothesizes
$f_k$ for some shortest prefix $f_k^{a_k}$. Define functions $f'_k =
k0^{a_k}10^\infty$.\<close>
define a where
"a \<equiv> \<lambda>k. LEAST x. (\<phi> (the ((\<phi> k) ((f k) \<triangleright> x)))) = f k"
define f' where "f' \<equiv> \<lambda>k. (k # (replicate (a k) 0) @ [1]) \<odot> 0\<^sup>\<infinity>"
then have "f' k \<in> \<R>" for k
using almost0_in_R1 by auto
\<comment> \<open>Although $f_k$ and $f'_k$ differ, they share the prefix of length $a_k + 1$.\<close>
have init_eq: "(f' k) \<triangleright> (a k) = (f k) \<triangleright> (a k)" for k
proof (rule init_eqI)
fix x assume "x \<le> a k"
then show "f' k x = f k x"
by (cases "x = 0") (simp_all add: nth_append f'_def f_def)
qed
have "f k \<noteq> f' k" for k
proof -
have "f k (Suc (a k)) \<down>= 0" using f_def by auto
moreover have "f' k (Suc (a k)) \<down>= 1"
using f'_def prepend[of "(k # (replicate (a k) 0) @ [1])" "0\<^sup>\<infinity>" "Suc (a k)"]
by (metis length_Cons length_append_singleton length_replicate lessI nth_Cons_Suc
nth_append_length)
ultimately show ?thesis by auto
qed
\<comment> \<open>The separating class $U$ contains $f'_k$ if $\varphi_k$
learns $f_k$; otherwise it contains $f_k$.\<close>
define U where
"U \<equiv> {if learn_lim \<phi> {f k} (\<phi> k) then f' k else f k |k. k \<in> UNIV}"
have "U \<notin> CP"
proof
assume "U \<in> CP"
have "\<exists>k. learn_cp \<phi> U (\<phi> k)"
proof -
have "\<exists>\<psi> s. learn_cp \<psi> U s"
- using CP_def `U \<in> CP` by auto
+ using CP_def \<open>U \<in> CP\<close> by auto
then obtain s where s: "learn_cp \<phi> U s"
using learn_cp_wrt_goedel[OF goedel_numbering_phi] by blast
then obtain k where "\<phi> k = s"
using phi_universal learn_cp_def learn_lim_def by auto
then show ?thesis using s by auto
qed
then obtain k where k: "learn_cp \<phi> U (\<phi> k)" by auto
then have learn: "learn_lim \<phi> U (\<phi> k)"
using learn_cp_def by simp
\<comment> \<open>If $f_k$ was in $U$, $\varphi_k$ would learn it. But then,
by definition of $U$, $f_k$ would not be in $U$. Hence $f_k \notin U$.\<close>
have "f k \<notin> U"
proof
assume "f k \<in> U"
then obtain m where m: "f k = (if learn_lim \<phi> {f m} (\<phi> m) then f' m else f m)"
using U_def by auto
have "f k 0 \<down>= m"
using f_def f'_def m by simp
moreover have "f k 0 \<down>= k" by (simp add: f_def)
ultimately have "m = k" by simp
with m have "f k = (if learn_lim \<phi> {f k} (\<phi> k) then f' k else f k)"
by auto
moreover have "learn_lim \<phi> {f k} (\<phi> k)"
using \<open>f k \<in> U\<close> learn_lim_closed_subseteq[OF learn] by simp
ultimately have "f k = f' k"
by simp
then show False
using \<open>f k \<noteq> f' k\<close> by simp
qed
then have "f' k \<in> U" using U_def by fastforce
then have in_U: "\<forall>n. \<phi> (the ((\<phi> k) ((f' k) \<triangleright> n))) \<in> U"
using learn_cpE(3)[OF k] by simp
\<comment> \<open>Since $f'_k \in U$, the strategy $\varphi_k$ learns $f_k$.
Then $a_k$ is well-defined, $f'^{a_k} = f^{a_k}$, and $\varphi_k$
hypothesizes $f_k$ on $f'^{a_k}$, which is not a class-preserving
hypothesis.\<close>
have "learn_lim \<phi> {f k} (\<phi> k)" using U_def \<open>f k \<notin> U\<close> by fastforce
then have "\<exists>i n\<^sub>0. \<phi> i = f k \<and> (\<forall>n\<ge>n\<^sub>0. \<phi> k ((f k) \<triangleright> n) \<down>= i)"
using learn_limE(2) by simp
then obtain i n\<^sub>0 where "\<phi> i = f k \<and> (\<forall>n\<ge>n\<^sub>0. \<phi> k ((f k) \<triangleright> n) \<down>= i)"
by auto
then have "\<phi> (the (\<phi> k ((f k) \<triangleright> (a k)))) = f k"
using a_def LeastI[of "\<lambda>x. (\<phi> (the ((\<phi> k) ((f k) \<triangleright> x)))) = f k" n\<^sub>0]
by simp
then have "\<phi> (the ((\<phi> k) ((f' k) \<triangleright> (a k)))) = f k"
using init_eq by simp
then show False
using \<open>f k \<notin> U\<close> in_U by metis
qed
moreover have "U \<in> NUM"
using NUM_closed_subseteq[OF U0_in_NUM, of U] f_def f'_def U0_def U_def
by fastforce
ultimately show ?thesis by auto
qed
section \<open>NUM is a proper subset of TOTAL\label{s:num_total}\<close>
text \<open>A NUM class $U$ is embedded in a total numbering @{term \<psi>}.
The strategy $S$ with $S(f^n) = \min \{i \mid \forall k \le n: \psi_i(k) =
f(k)\}$ for $f \in U$ converges to the least index of $f$ in @{term \<psi>},
and thus learns $f$ in the limit. Moreover it will be a TOTAL strategy
because @{term \<psi>} contains only total functions. This shows @{prop "NUM
\<subseteq> TOTAL"}.\<close>
text \<open>First we define, for every hypothesis space $\psi$, a
function that tries to determine for a given list $e$ and index $i$ whether
$e$ is a prefix of $\psi_i$. In other words it tries to decide whether $i$ is
a consistent hypothesis for $e$. ``Tries'' refers to the fact that the
function will diverge if $\psi_i(x)\uparrow$ for any $x \le |e|$. We start
with a version that checks the list only up to a given length.\<close>
definition r_consist_upto :: "recf \<Rightarrow> recf" where
"r_consist_upto r_psi \<equiv>
let g = Cn 4 r_ifeq
[Cn 4 r_psi [Id 4 2, Id 4 0], Cn 4 r_nth [Id 4 3, Id 4 0], Id 4 1, r_constn 3 1]
in Pr 2 (r_constn 1 0) g"
lemma r_consist_upto_recfn: "recfn 2 r_psi \<Longrightarrow> recfn 3 (r_consist_upto r_psi)"
using r_consist_upto_def by simp
lemma r_consist_upto:
assumes "recfn 2 r_psi"
shows "\<forall>k<j. eval r_psi [i, k] \<down> \<Longrightarrow>
eval (r_consist_upto r_psi) [j, i, e] =
(if \<forall>k<j. eval r_psi [i, k] \<down>= e_nth e k then Some 0 else Some 1)"
and "\<not> (\<forall>k<j. eval r_psi [i, k] \<down>) \<Longrightarrow> eval (r_consist_upto r_psi) [j, i, e] \<up>"
proof -
define g where "g =
Cn 4 r_ifeq
[Cn 4 r_psi [Id 4 2, Id 4 0], Cn 4 r_nth [Id 4 3, Id 4 0], Id 4 1, r_constn 3 1]"
then have "recfn 4 g"
using assms by simp
moreover have "eval (Cn 4 r_nth [Id 4 3, Id 4 0]) [j, r, i, e] \<down>= e_nth e j" for j r i e
by simp
moreover have "eval (r_constn 3 1) [j, r, i, e] \<down>= 1" for j r i e
by simp
moreover have "eval (Cn 4 r_psi [Id 4 2, Id 4 0]) [j, r, i, e] = eval r_psi [i, j]" for j r i e
using assms(1) by simp
ultimately have g: "eval g [j, r, i, e] =
(if eval r_psi [i, j] \<up> then None
else if eval r_psi [i, j] \<down>= e_nth e j then Some r else Some 1)"
for j r i e
- using `recfn 4 g` g_def assms by auto
+ using \<open>recfn 4 g\<close> g_def assms by auto
have goal1: "\<forall>k<j. eval r_psi [i, k] \<down> \<Longrightarrow>
eval (r_consist_upto r_psi) [j, i, e] =
(if \<forall>k<j. eval r_psi [i, k] \<down>= e_nth e k then Some 0 else Some 1)"
for j i e
proof (induction j)
case 0
then show ?case
using r_consist_upto_def r_consist_upto_recfn assms eval_Pr_0 by simp
next
case (Suc j)
then have "eval (r_consist_upto r_psi) [Suc j, i, e] =
eval g [j, the (eval (r_consist_upto r_psi) [j, i, e]), i, e]"
using assms eval_Pr_converg_Suc g_def r_consist_upto_def r_consist_upto_recfn
by simp
also have "... = eval g [j, if \<forall>k<j. eval r_psi [i, k] \<down>= e_nth e k then 0 else 1, i, e]"
using Suc by auto
also have "... \<down>= (if eval r_psi [i, j] \<down>= e_nth e j
then if \<forall>k<j. eval r_psi [i, k] \<down>= e_nth e k then 0 else 1 else 1)"
using g by (simp add: Suc.prems)
also have "... \<down>= (if \<forall>k<Suc j. eval r_psi [i, k] \<down>= e_nth e k then 0 else 1)"
by (simp add: less_Suc_eq)
finally show ?case by simp
qed
then show "\<forall>k<j. eval r_psi [i, k] \<down> \<Longrightarrow>
eval (r_consist_upto r_psi) [j, i, e] =
(if \<forall>k<j. eval r_psi [i, k] \<down>= e_nth e k then Some 0 else Some 1)"
by simp
show "\<not> (\<forall>k<j. eval r_psi [i, k] \<down>) \<Longrightarrow> eval (r_consist_upto r_psi) [j, i, e] \<up>"
proof -
assume "\<not> (\<forall>k<j. eval r_psi [i, k] \<down>)"
then have "\<exists>k<j. eval r_psi [i, k] \<up>" by simp
let ?P = "\<lambda>k. k < j \<and> eval r_psi [i, k] \<up>"
define kmin where "kmin = Least ?P"
then have "?P kmin"
using LeastI_ex[of ?P] \<open>\<exists>k<j. eval r_psi [i, k] \<up>\<close> by auto
from kmin_def have "\<And>k. k < kmin \<Longrightarrow> \<not> ?P k"
using kmin_def not_less_Least[of _ ?P] by blast
then have "\<forall>k < kmin. eval r_psi [i, k] \<down>"
- using `?P kmin` by simp
+ using \<open>?P kmin\<close> by simp
then have "eval (r_consist_upto r_psi) [kmin, i, e] =
(if \<forall>k<kmin. eval r_psi [i, k] \<down>= e_nth e k then Some 0 else Some 1)"
using goal1 by simp
moreover have "eval r_psi [i, kmin] \<up>"
- using `?P kmin` by simp
+ using \<open>?P kmin\<close> by simp
ultimately have "eval (r_consist_upto r_psi) [Suc kmin, i, e] \<up>"
using r_consist_upto_def g assms by simp
moreover have "j \<ge> kmin"
- using `?P kmin` by simp
+ using \<open>?P kmin\<close> by simp
ultimately show "eval (r_consist_upto r_psi) [j, i, e] \<up>"
- using r_consist_upto_def r_consist_upto_recfn `?P kmin` eval_Pr_converg_le assms
+ using r_consist_upto_def r_consist_upto_recfn \<open>?P kmin\<close> eval_Pr_converg_le assms
by (metis (full_types) Suc_leI length_Cons list.size(3) numeral_2_eq_2 numeral_3_eq_3)
qed
qed
text \<open>The next function provides the consistency decision functions we
need.\<close>
definition consistent :: "partial2 \<Rightarrow> partial2" where
"consistent \<psi> i e \<equiv>
if \<forall>k<e_length e. \<psi> i k \<down>
then if \<forall>k<e_length e. \<psi> i k \<down>= e_nth e k
then Some 0 else Some 1
else None"
text \<open>Given $i$ and $e$, @{term "consistent \<psi>"} decides whether $e$
is a prefix of $\psi_i$, provided $\psi_i$ is defined for the length of
$e$.\<close>
definition r_consistent :: "recf \<Rightarrow> recf" where
"r_consistent r_psi \<equiv>
Cn 2 (r_consist_upto r_psi) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]"
lemma r_consistent_recfn [simp]: "recfn 2 r_psi \<Longrightarrow> recfn 2 (r_consistent r_psi)"
using r_consistent_def r_consist_upto_recfn by simp
lemma r_consistent_converg:
assumes "recfn 2 r_psi" and "\<forall>k<e_length e. eval r_psi [i, k] \<down>"
shows "eval (r_consistent r_psi) [i, e] \<down>=
(if \<forall>k<e_length e. eval r_psi [i, k] \<down>= e_nth e k then 0 else 1)"
proof -
have "eval (r_consistent r_psi) [i, e] = eval (r_consist_upto r_psi) [e_length e, i, e]"
using r_consistent_def r_consist_upto_recfn assms(1) by simp
then show ?thesis using assms r_consist_upto(1) by simp
qed
lemma r_consistent_diverg:
assumes "recfn 2 r_psi" and "\<exists>k<e_length e. eval r_psi [i, k] \<up>"
shows "eval (r_consistent r_psi) [i, e] \<up>"
unfolding r_consistent_def
using r_consist_upto_recfn[OF assms(1)] r_consist_upto[OF assms(1)] assms(2)
by simp
lemma r_consistent:
assumes "recfn 2 r_psi" and "\<forall>x y. eval r_psi [x, y] = \<psi> x y"
shows "eval (r_consistent r_psi) [i, e] = consistent \<psi> i e"
proof (cases "\<forall>k<e_length e. \<psi> i k \<down>")
case True
then have "\<forall>k<e_length e. eval r_psi [i, k] \<down>"
using assms by simp
then show ?thesis
unfolding consistent_def using True by (simp add: assms r_consistent_converg)
next
case False
then have "consistent \<psi> i e \<up>"
unfolding consistent_def by auto
moreover have "eval (r_consistent r_psi) [i, e] \<up>"
using r_consistent_diverg[OF assms(1)] assms False by simp
ultimately show ?thesis by simp
qed
lemma consistent_in_P2:
assumes "\<psi> \<in> \<P>\<^sup>2"
shows "consistent \<psi> \<in> \<P>\<^sup>2"
using assms r_consistent P2E[OF assms(1)] P2I r_consistent_recfn by metis
lemma consistent_for_R2:
assumes "\<psi> \<in> \<R>\<^sup>2"
shows "consistent \<psi> i e =
(if \<forall>j<e_length e. \<psi> i j \<down>= e_nth e j then Some 0 else Some 1)"
using assms by (simp add: consistent_def)
lemma consistent_init:
assumes "\<psi> \<in> \<R>\<^sup>2" and "f \<in> \<R>"
shows "consistent \<psi> i (f \<triangleright> n) = (if \<psi> i \<triangleright> n = f \<triangleright> n then Some 0 else Some 1)"
using consistent_def[of _ _ "init f n"] assms init_eq_iff_eq_upto by simp
lemma consistent_in_R2:
assumes "\<psi> \<in> \<R>\<^sup>2"
shows "consistent \<psi> \<in> \<R>\<^sup>2"
using total2I consistent_in_P2 consistent_for_R2[OF assms] P2_total_imp_R2 R2_imp_P2 assms
by (metis option.simps(3))
text \<open>For total hypothesis spaces the next function computes the
minimum hypothesis consistent with a given prefix. It diverges if no such
hypothesis exists.\<close>
definition min_cons_hyp :: "partial2 \<Rightarrow> partial1" where
"min_cons_hyp \<psi> e \<equiv>
if \<exists>i. consistent \<psi> i e \<down>= 0 then Some (LEAST i. consistent \<psi> i e \<down>= 0) else None"
lemma min_cons_hyp_in_P1:
assumes "\<psi> \<in> \<R>\<^sup>2"
shows "min_cons_hyp \<psi> \<in> \<P>"
proof -
from assms consistent_in_R2 obtain rc where
rc: "recfn 2 rc" "total rc" "\<And>i e. eval rc [i, e] = consistent \<psi> i e"
using R2E[of "consistent \<psi>"] by metis
define r where "r = Mn 1 rc"
then have "recfn 1 r"
using rc(1) by simp
moreover from this have "eval r [e] = min_cons_hyp \<psi> e" for e
using r_def eval_Mn'[of 1 rc "[e]"] rc min_cons_hyp_def assms
by (auto simp add: consistent_in_R2)
ultimately show ?thesis by auto
qed
text \<open>The function @{term "min_cons_hyp \<psi>"} is a strategy for
learning all NUM classes embedded in @{term \<psi>}. It is an example of an
``identification-by-enumeration'' strategy.\<close>
lemma NUM_imp_learn_total:
assumes "\<psi> \<in> \<R>\<^sup>2" and "U \<in> NUM_wrt \<psi>"
shows "learn_total \<psi> U (min_cons_hyp \<psi>)"
proof (rule learn_totalI)
have ex_psi_i_f: "\<exists>i. \<psi> i = f" if "f \<in> U" for f
using assms that NUM_wrt_def by simp
moreover have consistent_eq_0: "consistent \<psi> i ((\<psi> i) \<triangleright> n) \<down>= 0" for i n
using assms by (simp add: consistent_init)
ultimately have "\<And>f n. f \<in> U \<Longrightarrow> min_cons_hyp \<psi> (f \<triangleright> n) \<down>"
using min_cons_hyp_def assms(1) by fastforce
then show env: "environment \<psi> U (min_cons_hyp \<psi>)"
using assms NUM_wrt_def min_cons_hyp_in_P1 NUM_E(1) NUM_I by auto
show "\<And>f n. f \<in> U \<Longrightarrow> \<psi> (the (min_cons_hyp \<psi> (f \<triangleright> n))) \<in> \<R>"
using assms by (simp)
show "\<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. min_cons_hyp \<psi> (f \<triangleright> n) \<down>= i)" if "f \<in> U" for f
proof -
from that env have "f \<in> \<R>" by auto
let ?P = "\<lambda>i. \<psi> i = f"
define imin where "imin \<equiv> Least ?P"
with ex_psi_i_f that have imin: "?P imin" "\<And>j. ?P j \<Longrightarrow> j \<ge> imin"
using LeastI_ex[of ?P] Least_le[of ?P] by simp_all
then have f_neq: "\<psi> i \<noteq> f" if "i < imin" for i
using leD that by auto
let ?Q = "\<lambda>i n. \<psi> i \<triangleright> n \<noteq> f \<triangleright> n"
define nu :: "nat \<Rightarrow> nat" where "nu = (\<lambda>i. SOME n. ?Q i n)"
have nu_neq: "\<psi> i \<triangleright> (nu i) \<noteq> f \<triangleright> (nu i)" if "i < imin" for i
proof -
from assms have "\<psi> i \<in> \<R>" by simp
moreover from assms imin(1) have "f \<in> \<R>" by auto
moreover have "f \<noteq> \<psi> i"
using that f_neq by auto
ultimately have "\<exists>n. f \<triangleright> n \<noteq> (\<psi> i) \<triangleright> n"
using neq_fun_neq_init by simp
then show "?Q i (nu i)"
unfolding nu_def using someI_ex[of "\<lambda>n. ?Q i n"] by metis
qed
have "\<exists>n\<^sub>0. \<forall>n\<ge>n\<^sub>0. min_cons_hyp \<psi> (f \<triangleright> n) \<down>= imin"
proof (cases "imin = 0")
case True
then have "\<forall>n. min_cons_hyp \<psi> (f \<triangleright> n) \<down>= imin"
using consistent_eq_0 assms(1) imin(1) min_cons_hyp_def by auto
then show ?thesis by simp
next
case False
define n\<^sub>0 where "n\<^sub>0 = Max (set (map nu [0..<imin]))" (is "_ = Max ?N")
have "nu i \<le> n\<^sub>0" if "i < imin" for i
proof -
have "finite ?N"
using n\<^sub>0_def by simp
moreover have "?N \<noteq> {}"
using False n\<^sub>0_def by simp
moreover have "nu i \<in> ?N"
using that by simp
ultimately show ?thesis
using that Max_ge n\<^sub>0_def by blast
qed
then have "\<psi> i \<triangleright> n\<^sub>0 \<noteq> f \<triangleright> n\<^sub>0" if "i < imin" for i
using nu_neq neq_init_forall_ge that by blast
then have *: "\<psi> i \<triangleright> n \<noteq> f \<triangleright> n" if "i < imin" and "n \<ge> n\<^sub>0" for i n
using nu_neq neq_init_forall_ge that by blast
have "\<psi> imin \<triangleright> n = f \<triangleright> n" for n
using imin(1) by simp
moreover have "(consistent \<psi> i (f \<triangleright> n) \<down>= 0) = (\<psi> i \<triangleright> n = f \<triangleright> n)" for i n
by (simp add: \<open>f \<in> \<R>\<close> assms(1) consistent_init)
ultimately have "min_cons_hyp \<psi> (f \<triangleright> n) \<down>= (LEAST i. \<psi> i \<triangleright> n = f \<triangleright> n)" for n
using min_cons_hyp_def[of \<psi> "f \<triangleright> n"] by auto
moreover have "(LEAST i. \<psi> i \<triangleright> n = f \<triangleright> n) = imin" if "n \<ge> n\<^sub>0" for n
proof (rule Least_equality)
show "\<psi> imin \<triangleright> n = f \<triangleright> n"
using imin(1) by simp
show "\<And>y. \<psi> y \<triangleright> n = f \<triangleright> n \<Longrightarrow> imin \<le> y"
using imin * leI that by blast
qed
ultimately have "min_cons_hyp \<psi> (f \<triangleright> n) \<down>= imin" if "n \<ge> n\<^sub>0" for n
using that by blast
then show ?thesis by auto
qed
with imin(1) show ?thesis by auto
qed
qed
corollary NUM_subseteq_TOTAL: "NUM \<subseteq> TOTAL"
proof
fix U
assume "U \<in> NUM"
then have "\<exists>\<psi>\<in>\<R>\<^sup>2. \<forall>f\<in>U. \<exists>i. \<psi> i = f" by auto
then have "\<exists>\<psi>\<in>\<R>\<^sup>2. U \<in> NUM_wrt \<psi>"
using NUM_wrt_def by simp
then have "\<exists>\<psi> s. learn_total \<psi> U s"
using NUM_imp_learn_total by auto
then show "U \<in> TOTAL"
using TOTAL_def by auto
qed
text \<open>The class @{term V0} is in @{term "TOTAL - NUM"}. \<close>
theorem NUM_subset_TOTAL: "NUM \<subset> TOTAL"
using CP_subseteq_TOTAL FIN_not_subseteq_NUM FIN_subseteq_CP NUM_subseteq_TOTAL
by auto
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/Inductive_Inference_Basics.thy b/thys/Inductive_Inference/Inductive_Inference_Basics.thy
--- a/thys/Inductive_Inference/Inductive_Inference_Basics.thy
+++ b/thys/Inductive_Inference/Inductive_Inference_Basics.thy
@@ -1,1222 +1,1222 @@
chapter \<open>Inductive inference of recursive functions\label{c:iirf}\<close>
theory Inductive_Inference_Basics
imports Standard_Results
begin
text \<open>Inductive inference originates from work by
Solomonoff~\cite{s-ftiip1-64,s-ftiip2-64} and Gold~\cite{g-lil-67,g-lr-65}
and comes in many variations. The common theme is to infer additional
information about objects, such as formal languages or functions, from incomplete
data, such as finitely many words contained in the language or argument-value
pairs of the function. Oftentimes ``additional information'' means complete
information, such that the task becomes identification of the object.
The basic setting in inductive inference of recursive functions is as follows.
Let us denote, for a total function $f$, by $f^n$ the code of the list
$[f(0), ..., f(n)]$. Let $U$ be a set (called \emph{class}) of total
recursive functions, and $\psi$ a binary partial recursive function
(called \emph{hypothesis space}).
A partial recursive function $S$ (called \emph{strategy})
is said to \emph{learn $U$ in the limit with respect to $\psi$} if
for all $f \in U$,
\begin{itemize}
\item the value $S(f^n)$ is defined for all $n\in\mathbb{N}$,
\item the sequence $S(f^0), S(f^1), \ldots$ converges to an
$i\in\mathbb{N}$ with $\psi_i = f$.
\end{itemize}
Both the output $S(f^n)$ of the strategy and its interpretation
as a function $\psi_{S(f^n)}$ are called \emph{hypothesis}. The set
of all classes learnable in the limit by $S$ with respect to $\psi$ is
denoted by $\mathrm{LIM}_\psi(S)$. Moreover we set $\mathrm{LIM}_\psi =
\bigcup_{S\in\mathcal{P}} \mathrm{LIM}_\psi(S)$ and $\mathrm{LIM} =
\bigcup_{\psi\in\mathcal{P}^2} \mathrm{LIM}_\psi$. We call the latter set the
\emph{inference type} $\mathrm{LIM}$.
Many aspects of this setting can be varied. We shall consider:
\begin{itemize}
\item Intermediate hypotheses: $\psi_{S(f^n)}$ can be required to be total or
to be in the class $U$, or to coincide with $f$ on arguments up to $n$, or
a myriad of other conditions or combinations thereof.
\item Convergence of hypotheses:
\begin{itemize}
\item The strategy can be required to output not a sequence but a single
hypothesis, which must be correct.
\item The strategy can be required to converge to a \emph{function} rather
than an index.
\end{itemize}
\end{itemize}
We formalize five kinds of results (\<open>\<I>\<close> and \<open>\<I>'\<close> stand for
inference types):
\begin{itemize}
\item Comparison of learning power: results of the form @{prop "\<I>
\<subset> \<I>'"}, in particular showing that the inclusion is proper
(Sections~\ref{s:fin_cp}, \ref{s:num_fin}, \ref{s:num_cp},
\ref{s:num_total}, \ref{s:cons_lim}, \ref{s:lim_bc}, \ref{s:total_cons},
\ref{s:r1_bc}).
\item Whether \<open>\<I>\<close> is closed under the subset relation: @{prop "U
\<in> \<I> \<and> V \<subseteq> U \<Longrightarrow> V \<in> \<I>"}.
\item Whether \<open>\<I>\<close> is closed under union: @{prop "U \<in> \<I> \<and>
V \<in> \<I> \<Longrightarrow> U \<union> V \<in> \<I>"} (Section~\ref{s:union}).
\item Whether every class in \<open>\<I>\<close> can be learned with respect to a
Gödel numbering as hypothesis space (Section~\ref{s:inference_types}).
\item Whether every class in \<open>\<I>\<close> can be learned by a \emph{total}
recursive strategy (Section~\ref{s:lemma_r}).
\end{itemize}
The bulk of this chapter is devoted to the first category of results. Most
results that we are going to formalize have been called ``classical'' by
Jantke and Beick~\cite{jb-cpnii-81}, who compare a large number of inference
types. Another comparison is by Case and Smith~\cite{cs-cicmii-83}. Angluin
and Smith~\cite{as-ii-87} give an overview of various forms of inductive
inference.
All (interesting) proofs herein are based on my lecture notes of the
\emph{Induktive Inferenz} lectures by Rolf Wiehagen from 1999/2000 and
2000/2001 at the University of Kaiserslautern. I have given references to the
original proofs whenever I was able to find them. For the other proofs, as
well as for those that I had to contort beyond recognition, I provide proof
sketches.\<close>
section \<open>Preliminaries\<close>
text \<open>Throughout the chapter, in particular in proof sketches, we use
the following notation.
Let $b\in\mathbb{N}^*$ be a list of numbers. We write $|b|$ for its length
and $b_i$ for the $i$-th element ($i=0,\dots, |b| - 1$). Concatenation of
numbers and lists works in the obvious way; for instance, $jbk$ with
$j,k\in\mathbb{N}$, $b\in\mathbb{N}^*$ refers to the list $jb_0\dots
b_{|b|-1}k$. For $0 \leq i < |b|$, the term $b_{i:=v}$ denotes the list
$b_0\dots b_{i-1}vb_{i+1}\dots b_{|b|-1}$. The notation $b_{<i}$ refers to
$b_0\dots b_{i-1}$ for $0 < i \leq |b|$. Moreover, $v^n$ is short for the
list consisting of $n$ times the value $v \in \mathbb{N}$.
Unary partial functions can be regarded as infinite sequences consisting of
numbers and the symbol~$\uparrow$ denoting undefinedness. We abbreviate the
empty function by $\uparrow^\infty$ and the constant zero function by
$0^\infty$. A function can be written as a list concatenated with a partial
function. For example, $jb\uparrow^\infty$ is the function
\[
x \mapsto \left\{\begin{array}{ll}
j & \mbox{if } x = 0,\\
b_{x-1} & \mbox{if } 0 < x \leq |b|,\\
\uparrow & \mbox{otherwise,}
\end{array}\right.
\]
and $jp$, where $p$ is a function, means
\[
x \mapsto \left\{\begin{array}{ll}
j & \mbox{if } x = 0,\\
p(x-1) & \mbox{otherwise.}
\end{array}\right.
\]
A \emph{numbering} is a function $\psi \in \mathcal{P}^2$.\<close>
subsection \<open>The prefixes of a function\<close>
text \<open>A \emph{prefix}, also called \emph{initial segment}, is a list of
initial values of a function.\<close>
definition prefix :: "partial1 \<Rightarrow> nat \<Rightarrow> nat list" where
"prefix f n \<equiv> map (\<lambda>x. the (f x)) [0..<Suc n]"
lemma length_prefix [simp]: "length (prefix f n) = Suc n"
unfolding prefix_def by simp
lemma prefix_nth [simp]:
assumes "k < Suc n"
shows "prefix f n ! k = the (f k)"
unfolding prefix_def using assms nth_map_upt[of k "Suc n" 0 "\<lambda>x. the (f x)"] by simp
lemma prefixI:
assumes "length vs > 0" and "\<And>x. x < length vs \<Longrightarrow> f x \<down>= vs ! x"
shows "prefix f (length vs - 1) = vs"
using assms nth_equalityI[of "prefix f (length vs - 1)" vs] by simp
lemma prefixI':
assumes "length vs = Suc n" and "\<And>x. x < Suc n \<Longrightarrow> f x \<down>= vs ! x"
shows "prefix f n = vs"
using assms nth_equalityI[of "prefix f (length vs - 1)" vs] by simp
lemma prefixE:
assumes "prefix f (length vs - 1) = vs"
and "f \<in> \<R>"
and "length vs > 0"
and "x < length vs"
shows "f x \<down>= vs ! x"
using assms length_prefix prefix_nth[of x "length vs - 1" f] by simp
lemma prefix_eqI:
assumes "\<And>x. x \<le> n \<Longrightarrow> f x = g x"
shows "prefix f n = prefix g n"
using assms prefix_def by simp
lemma prefix_0: "prefix f 0 = [the (f 0)]"
using prefix_def by simp
lemma prefix_Suc: "prefix f (Suc n) = prefix f n @ [the (f (Suc n))]"
unfolding prefix_def by simp
lemma take_prefix:
assumes "f \<in> \<R>" and "k \<le> n"
shows "prefix f k = take (Suc k) (prefix f n)"
proof -
let ?vs = "take (Suc k) (prefix f n)"
have "length ?vs = Suc k"
using assms(2) by simp
then have "\<And>x. x < length ?vs \<Longrightarrow> f x \<down>= ?vs ! x"
using assms by auto
then show ?thesis
- using prefixI[where ?vs="?vs"] `length ?vs = Suc k` by simp
+ using prefixI[where ?vs="?vs"] \<open>length ?vs = Suc k\<close> by simp
qed
text \<open>Strategies receive prefixes in the form of encoded lists. The
term ``prefix'' refers to both encoded and unencoded lists. We use the
notation @{text "f \<triangleright> n"} for the prefix $f^n$.\<close>
definition init :: "partial1 \<Rightarrow> nat \<Rightarrow> nat" (infix "\<triangleright>" 110) where
"f \<triangleright> n \<equiv> list_encode (prefix f n)"
lemma init_neq_zero: "f \<triangleright> n \<noteq> 0"
unfolding init_def prefix_def using list_encode_0 by fastforce
lemma init_prefixE [elim]: "prefix f n = prefix g n \<Longrightarrow> f \<triangleright> n = g \<triangleright> n"
unfolding init_def by simp
lemma init_eqI:
assumes "\<And>x. x \<le> n \<Longrightarrow> f x = g x"
shows "f \<triangleright> n = g \<triangleright> n"
unfolding init_def using prefix_eqI[OF assms] by simp
lemma initI:
assumes "e_length e > 0" and "\<And>x. x < e_length e \<Longrightarrow> f x \<down>= e_nth e x"
shows "f \<triangleright> (e_length e - 1) = e"
unfolding init_def using assms prefixI by simp
lemma initI':
assumes "e_length e = Suc n" and "\<And>x. x < Suc n \<Longrightarrow> f x \<down>= e_nth e x"
shows "f \<triangleright> n = e"
unfolding init_def using assms prefixI' by simp
lemma init_iff_list_eq_upto:
assumes "f \<in> \<R>" and "e_length vs > 0"
shows "(\<forall>x<e_length vs. f x \<down>= e_nth vs x) \<longleftrightarrow> prefix f (e_length vs - 1) = list_decode vs"
using prefixI[OF assms(2)] prefixE[OF _ assms] by auto
lemma length_init [simp]: "e_length (f \<triangleright> n) = Suc n"
unfolding init_def by simp
lemma init_Suc_snoc: "f \<triangleright> (Suc n) = e_snoc (f \<triangleright> n) (the (f (Suc n)))"
unfolding init_def by (simp add: prefix_Suc)
lemma nth_init: "i < Suc n \<Longrightarrow> e_nth (f \<triangleright> n) i = the (f i)"
unfolding init_def using prefix_nth by auto
lemma hd_init [simp]: "e_hd (f \<triangleright> n) = the (f 0)"
unfolding init_def using init_neq_zero by (simp add: e_hd_nth0)
lemma list_decode_init [simp]: "list_decode (f \<triangleright> n) = prefix f n"
unfolding init_def by simp
lemma init_eq_iff_eq_upto:
assumes "g \<in> \<R>" and "f \<in> \<R>"
shows "(\<forall>j<Suc n. g j = f j) \<longleftrightarrow> g \<triangleright> n = f \<triangleright> n"
using assms initI' init_iff_list_eq_upto length_init list_decode_init
by (metis diff_Suc_1 zero_less_Suc)
definition is_init_of :: "nat \<Rightarrow> partial1 \<Rightarrow> bool" where
"is_init_of t f \<equiv> \<forall>i<e_length t. f i \<down>= e_nth t i"
lemma not_initial_imp_not_eq:
assumes "\<And>x. x < Suc n \<Longrightarrow> f x \<down>" and "\<not> (is_init_of (f \<triangleright> n) g)"
shows "f \<noteq> g"
using is_init_of_def assms by auto
lemma all_init_eq_imp_fun_eq:
assumes "f \<in> \<R>" and "g \<in> \<R>" and "\<And>n. f \<triangleright> n = g \<triangleright> n"
shows "f = g"
proof
fix n
from assms have "prefix f n = prefix g n"
by (metis init_def list_decode_encode)
then have "the (f n) = the (g n)"
unfolding init_def prefix_def by simp
then show "f n = g n"
using assms(1,2) by (meson R1_imp_total1 option.expand total1E)
qed
corollary neq_fun_neq_init:
assumes "f \<in> \<R>" and "g \<in> \<R>" and "f \<noteq> g"
shows "\<exists>n. f \<triangleright> n \<noteq> g \<triangleright> n"
using assms all_init_eq_imp_fun_eq by auto
lemma eq_init_forall_le:
assumes "f \<triangleright> n = g \<triangleright> n" and "m \<le> n"
shows "f \<triangleright> m = g \<triangleright> m"
proof -
from assms(1) have "prefix f n = prefix g n"
by (metis init_def list_decode_encode)
then have "the (f k) = the (g k)" if "k \<le> n" for k
using prefix_def that by auto
then have "the (f k) = the (g k)" if "k \<le> m" for k
using assms(2) that by simp
then have "prefix f m = prefix g m"
using prefix_def by simp
then show ?thesis by (simp add: init_def)
qed
corollary neq_init_forall_ge:
assumes "f \<triangleright> n \<noteq> g \<triangleright> n" and "m \<ge> n"
shows "f \<triangleright> m \<noteq> g \<triangleright> m"
using eq_init_forall_le assms by blast
lemma e_take_init:
assumes "f \<in> \<R>" and "k < Suc n"
shows "e_take (Suc k) (f \<triangleright> n) = f \<triangleright> k"
using assms take_prefix by (simp add: init_def less_Suc_eq_le)
lemma init_butlast_init:
assumes "total1 f" and "f \<triangleright> n = e" and "n > 0"
shows "f \<triangleright> (n - 1) = e_butlast e"
proof -
let ?e = "e_butlast e"
have "e_length e = Suc n"
using assms(2) by auto
then have len: "e_length ?e = n"
by simp
have "f \<triangleright> (e_length ?e - 1) = ?e"
proof (rule initI)
show "0 < e_length ?e"
using assms(3) len by simp
have "\<And>x. x < e_length e \<Longrightarrow> f x \<down>= e_nth e x"
using assms(1,2) total1_def \<open>e_length e = Suc n\<close> by auto
then show "\<And>x. x < e_length ?e \<Longrightarrow> f x \<down>= e_nth ?e x"
by (simp add: butlast_conv_take)
qed
with len show ?thesis by simp
qed
text \<open>Some definitions make use of recursive predicates, that is,
$01$-valued functions.\<close>
definition RPred1 :: "partial1 set" ("\<R>\<^sub>0\<^sub>1") where
"\<R>\<^sub>0\<^sub>1 \<equiv> {f. f \<in> \<R> \<and> (\<forall>x. f x \<down>= 0 \<or> f x \<down>= 1)}"
lemma RPred1_subseteq_R1: "\<R>\<^sub>0\<^sub>1 \<subseteq> \<R>"
unfolding RPred1_def by auto
lemma const0_in_RPred1: "(\<lambda>_. Some 0) \<in> \<R>\<^sub>0\<^sub>1"
using RPred1_def const_in_Prim1 by fast
lemma RPred1_altdef: "\<R>\<^sub>0\<^sub>1 = {f. f \<in> \<R> \<and> (\<forall>x. the (f x) \<le> 1)}"
(is "\<R>\<^sub>0\<^sub>1 = ?S")
proof
show "\<R>\<^sub>0\<^sub>1 \<subseteq> ?S"
proof
fix f
assume f: "f \<in> \<R>\<^sub>0\<^sub>1"
with RPred1_def have "f \<in> \<R>" by auto
from f have "\<forall>x. f x \<down>= 0 \<or> f x \<down>= 1"
by (simp add: RPred1_def)
then have "\<forall>x. the (f x) \<le> 1"
by (metis eq_refl less_Suc_eq_le zero_less_Suc option.sel)
- with `f \<in> \<R>` show "f \<in> ?S" by simp
+ with \<open>f \<in> \<R>\<close> show "f \<in> ?S" by simp
qed
show "?S \<subseteq> \<R>\<^sub>0\<^sub>1"
proof
fix f
assume f: "f \<in> ?S"
then have "f \<in> \<R>" by simp
then have total: "\<And>x. f x \<down>" by auto
from f have "\<forall>x. the (f x) = 0 \<or> the (f x) = 1"
by (simp add: le_eq_less_or_eq)
with total have "\<forall>x. f x \<down>= 0 \<or> f x \<down>= 1"
by (metis option.collapse)
then show "f \<in> \<R>\<^sub>0\<^sub>1"
- using `f \<in> \<R>` RPred1_def by auto
+ using \<open>f \<in> \<R>\<close> RPred1_def by auto
qed
qed
subsection \<open>NUM\<close>
text \<open>A class of recursive functions is in NUM if it can be
embedded in a total numbering. Thus, for learning such classes there is
always a total hypothesis space available.\<close>
definition NUM :: "partial1 set set" where
"NUM \<equiv> {U. \<exists>\<psi>\<in>\<R>\<^sup>2. \<forall>f\<in>U. \<exists>i. \<psi> i = f}"
definition NUM_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"\<psi> \<in> \<R>\<^sup>2 \<Longrightarrow> NUM_wrt \<psi> \<equiv> {U. \<forall>f\<in>U. \<exists>i. \<psi> i = f}"
lemma NUM_I [intro]:
assumes "\<psi> \<in> \<R>\<^sup>2" and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f"
shows "U \<in> NUM"
using assms NUM_def by blast
lemma NUM_E [dest]:
assumes "U \<in> NUM"
shows "U \<subseteq> \<R>"
and "\<exists>\<psi>\<in>\<R>\<^sup>2. \<forall>f\<in>U. \<exists>i. \<psi> i = f"
using NUM_def assms by (force, auto)
lemma NUM_closed_subseteq:
assumes "U \<in> NUM" and "V \<subseteq> U"
shows "V \<in> NUM"
using assms subset_eq[of V U] NUM_I by auto
text \<open>This is the classical diagonalization proof showing that there is
no total numbering containing all total recursive functions.\<close>
lemma R1_not_in_NUM: "\<R> \<notin> NUM"
proof
assume "\<R> \<in> NUM"
then obtain \<psi> where num: "\<psi> \<in> \<R>\<^sup>2" "\<forall>f\<in>\<R>. \<exists>i. \<psi> i = f"
by auto
then obtain psi where psi: "recfn 2 psi" "total psi" "eval psi [i, x] = \<psi> i x" for i x
by auto
define d where "d = Cn 1 S [Cn 1 psi [Id 1 0, Id 1 0]]"
then have "recfn 1 d"
using psi(1) by simp
moreover have d: "eval d [x] \<down>= Suc (the (\<psi> x x))" for x
unfolding d_def using num psi by simp
ultimately have "(\<lambda>x. eval d [x]) \<in> \<R>"
using R1I by blast
then obtain i where "\<psi> i = (\<lambda>x. eval d [x])"
using num(2) by auto
then have "\<psi> i i = eval d [i]" by simp
with d have "\<psi> i i \<down>= Suc (the (\<psi> i i))" by simp
then show False
using option.sel[of "Suc (the (\<psi> i i))"] by simp
qed
text \<open>A hypothesis space that contains a function for every prefix will
come in handy. The following is a total numbering with this property.\<close>
definition "r_prenum \<equiv>
Cn 2 r_ifless [Id 2 1, Cn 2 r_length [Id 2 0], Cn 2 r_nth [Id 2 0, Id 2 1], r_constn 1 0]"
lemma r_prenum_prim [simp]: "prim_recfn 2 r_prenum"
unfolding r_prenum_def by simp_all
lemma r_prenum [simp]:
"eval r_prenum [e, x] \<down>= (if x < e_length e then e_nth e x else 0)"
by (simp add: r_prenum_def)
definition prenum :: partial2 where
"prenum e x \<equiv> Some (if x < e_length e then e_nth e x else 0)"
lemma prenum_in_R2: "prenum \<in> \<R>\<^sup>2"
using prenum_def Prim2I[OF r_prenum_prim, of prenum] by simp
lemma prenum [simp]: "prenum e x \<down>= (if x < e_length e then e_nth e x else 0)"
unfolding prenum_def ..
lemma prenum_encode:
"prenum (list_encode vs) x \<down>= (if x < length vs then vs ! x else 0)"
using prenum_def by (cases "x < length vs") simp_all
text \<open>Prepending a list of numbers to a function:\<close>
definition prepend :: "nat list \<Rightarrow> partial1 \<Rightarrow> partial1" (infixr "\<odot>" 64) where
"vs \<odot> f \<equiv> \<lambda>x. if x < length vs then Some (vs ! x) else f (x - length vs)"
lemma prepend [simp]:
"(vs \<odot> f) x = (if x < length vs then Some (vs ! x) else f (x - length vs))"
unfolding prepend_def ..
lemma prepend_total: "total1 f \<Longrightarrow> total1 (vs \<odot> f)"
unfolding total1_def by simp
lemma prepend_at_less:
assumes "n < length vs"
shows "(vs \<odot> f) n \<down>= vs ! n"
using assms by simp
lemma prepend_at_ge:
assumes "n \<ge> length vs"
shows "(vs \<odot> f) n = f (n - length vs)"
using assms by simp
lemma prefix_prepend_less:
assumes "n < length vs"
shows "prefix (vs \<odot> f) n = take (Suc n) vs"
using assms length_prefix by (intro nth_equalityI) simp_all
lemma prepend_eqI:
assumes "\<And>x. x < length vs \<Longrightarrow> g x \<down>= vs ! x"
and "\<And>x. g (length vs + x) = f x"
shows "g = vs \<odot> f"
proof
fix x
show "g x = (vs \<odot> f) x"
proof (cases "x < length vs")
case True
then show ?thesis using assms by simp
next
case False
then show ?thesis
using assms prepend by (metis add_diff_inverse_nat)
qed
qed
fun r_prepend :: "nat list \<Rightarrow> recf \<Rightarrow> recf" where
"r_prepend [] r = r"
| "r_prepend (v # vs) r =
Cn 1 (r_lifz (r_const v) (Cn 1 (r_prepend vs r) [r_dec])) [Id 1 0, Id 1 0]"
lemma r_prepend_recfn:
assumes "recfn 1 r"
shows "recfn 1 (r_prepend vs r)"
using assms by (induction vs) simp_all
lemma r_prepend:
assumes "recfn 1 r"
shows "eval (r_prepend vs r) [x] =
(if x < length vs then Some (vs ! x) else eval r [x - length vs])"
proof (induction vs arbitrary: x)
case Nil
then show ?case using assms by simp
next
case (Cons v vs)
show ?case
using assms Cons by (cases "x = 0") (auto simp add: r_prepend_recfn)
qed
lemma r_prepend_total:
assumes "recfn 1 r" and "total r"
shows "eval (r_prepend vs r) [x] \<down>=
(if x < length vs then vs ! x else the (eval r [x - length vs]))"
proof (induction vs arbitrary: x)
case Nil
then show ?case using assms by simp
next
case (Cons v vs)
show ?case
using assms Cons by (cases "x = 0") (auto simp add: r_prepend_recfn)
qed
lemma prepend_in_P1:
assumes "f \<in> \<P>"
shows "vs \<odot> f \<in> \<P>"
proof -
obtain r where r: "recfn 1 r" "\<And>x. eval r [x] = f x"
using assms by auto
moreover have "recfn 1 (r_prepend vs r)"
using r r_prepend_recfn by simp
moreover have "eval (r_prepend vs r) [x] = (vs \<odot> f) x" for x
using r r_prepend by simp
ultimately show ?thesis by blast
qed
lemma prepend_in_R1:
assumes "f \<in> \<R>"
shows "vs \<odot> f \<in> \<R>"
proof -
obtain r where r: "recfn 1 r" "total r" "\<And>x. eval r [x] = f x"
using assms by auto
then have "total1 f"
using R1_imp_total1[OF assms] by simp
have "total (r_prepend vs r)"
using r r_prepend_total r_prepend_recfn totalI1[of "r_prepend vs r"] by simp
with r have "total (r_prepend vs r)" by simp
moreover have "recfn 1 (r_prepend vs r)"
using r r_prepend_recfn by simp
moreover have "eval (r_prepend vs r) [x] = (vs \<odot> f) x" for x
- using r r_prepend `total1 f` total1E by simp
+ using r r_prepend \<open>total1 f\<close> total1E by simp
ultimately show ?thesis by auto
qed
lemma prepend_associative: "(us @ vs) \<odot> f = us \<odot> vs \<odot> f" (is "?lhs = ?rhs")
proof
fix x
consider
"x < length us"
| "x \<ge> length us \<and> x < length (us @ vs)"
| "x \<ge> length (us @ vs)"
by linarith
then show "?lhs x = ?rhs x"
proof (cases)
case 1
then show ?thesis
by (metis le_add1 length_append less_le_trans nth_append prepend_at_less)
next
case 2
then show ?thesis
by (smt add_diff_inverse_nat add_less_cancel_left length_append nth_append prepend)
next
case 3
then show ?thesis
using prepend_at_ge by auto
qed
qed
abbreviation constant_divergent :: partial1 ("\<up>\<^sup>\<infinity>") where
"\<up>\<^sup>\<infinity> \<equiv> \<lambda>_. None"
abbreviation constant_zero :: partial1 ("0\<^sup>\<infinity>") where
"0\<^sup>\<infinity> \<equiv> \<lambda>_. Some 0"
lemma almost0_in_R1: "vs \<odot> 0\<^sup>\<infinity> \<in> \<R>"
using RPred1_subseteq_R1 const0_in_RPred1 prepend_in_R1 by auto
text \<open>The class $U_0$ of all total recursive functions that are almost
everywhere zero will be used several times to construct
(counter-)examples.\<close>
definition U0 :: "partial1 set" ("U\<^sub>0") where
"U\<^sub>0 \<equiv> {vs \<odot> 0\<^sup>\<infinity> |vs. vs \<in> UNIV}"
text \<open>The class @{term U0} contains exactly the functions in the
numbering @{term prenum}.\<close>
lemma U0_altdef: "U\<^sub>0 = {prenum e| e. e \<in> UNIV}" (is "U\<^sub>0 = ?W")
proof
show "U\<^sub>0 \<subseteq> ?W"
proof
fix f
assume "f \<in> U\<^sub>0"
with U0_def obtain vs where "f = vs \<odot> 0\<^sup>\<infinity>"
by auto
then have "f = prenum (list_encode vs)"
using prenum_encode by auto
then show "f \<in> ?W" by auto
qed
show "?W \<subseteq> U\<^sub>0"
unfolding U0_def by fastforce
qed
lemma U0_in_NUM: "U\<^sub>0 \<in> NUM"
using prenum_in_R2 U0_altdef by (intro NUM_I[of prenum]; force)
text \<open>Every almost-zero function can be represented by $v0^\infty$ for
a list $v$ not ending in zero.\<close>
lemma almost0_canonical:
assumes "f = vs \<odot> 0\<^sup>\<infinity>" and "f \<noteq> 0\<^sup>\<infinity>"
obtains ws where "length ws > 0" and "last ws \<noteq> 0" and "f = ws \<odot> 0\<^sup>\<infinity>"
proof -
let ?P = "\<lambda>k. k < length vs \<and> vs ! k \<noteq> 0"
from assms have "vs \<noteq> []"
by auto
then have ex: "\<exists>k<length vs. vs ! k \<noteq> 0"
using assms by auto
define m where "m = Greatest ?P"
moreover have le: "\<forall>y. ?P y \<longrightarrow> y \<le> length vs"
by simp
ultimately have "?P m"
using ex GreatestI_ex_nat[of ?P "length vs"] by simp
have not_gr: "\<not> ?P k" if "k > m" for k
using Greatest_le_nat[of ?P _ "length vs"] m_def ex le not_less that by blast
let ?ws = "take (Suc m) vs"
have "vs \<odot> 0\<^sup>\<infinity> = ?ws \<odot> 0\<^sup>\<infinity>"
proof
fix x
show "(vs \<odot> 0\<^sup>\<infinity>) x = (?ws \<odot> 0\<^sup>\<infinity>) x"
proof (cases "x < Suc m")
case True
- then show ?thesis using `?P m` by simp
+ then show ?thesis using \<open>?P m\<close> by simp
next
case False
moreover from this have "(?ws \<odot> 0\<^sup>\<infinity>) x \<down>= 0"
by simp
ultimately show ?thesis
using not_gr by (cases "x < length vs") simp_all
qed
qed
then have "f = ?ws \<odot> 0\<^sup>\<infinity>"
using assms(1) by simp
moreover have "length ?ws > 0"
by (simp add: \<open>vs \<noteq> []\<close>)
moreover have "last ?ws \<noteq> 0"
by (simp add: \<open>?P m\<close> take_Suc_conv_app_nth)
ultimately show ?thesis using that by blast
qed
section \<open>Types of inference\label{s:inference_types}\<close>
text \<open>This section introduces all inference types that we are going to
consider together with some of their simple properties. All these inference
types share the following condition, which essentially says that everything
must be computable:\<close>
abbreviation environment :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"environment \<psi> U s \<equiv> \<psi> \<in> \<P>\<^sup>2 \<and> U \<subseteq> \<R> \<and> s \<in> \<P> \<and> (\<forall>f\<in>U. \<forall>n. s (f \<triangleright> n) \<down>)"
subsection \<open>LIM: Learning in the limit\<close>
text \<open>A strategy $S$ learns a class $U$ in the limit with respect to a
hypothesis space @{term "\<psi> \<in> \<P>\<^sup>2"} if for all $f\in U$, the
sequence $(S(f^n))_{n\in\mathbb{N}}$ converges to an $i$ with $\psi_i = f$.
Convergence for a sequence of natural numbers means that almost all elements
are the same. We express this with the following notation.\<close>
abbreviation Almost_All :: "(nat \<Rightarrow> bool) \<Rightarrow> bool" (binder "\<forall>\<^sup>\<infinity>" 10) where
"\<forall>\<^sup>\<infinity>n. P n \<equiv> \<exists>n\<^sub>0. \<forall>n\<ge>n\<^sub>0. P n"
definition learn_lim :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_lim \<psi> U s \<equiv>
environment \<psi> U s \<and>
(\<forall>f\<in>U. \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i))"
lemma learn_limE:
assumes "learn_lim \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)"
using assms learn_lim_def by auto
lemma learn_limI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)"
shows "learn_lim \<psi> U s"
using assms learn_lim_def by auto
definition LIM_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"LIM_wrt \<psi> \<equiv> {U. \<exists>s. learn_lim \<psi> U s}"
definition Lim :: "partial1 set set" ("LIM") where
"LIM \<equiv> {U. \<exists>\<psi> s. learn_lim \<psi> U s}"
text \<open>LIM is closed under the the subset relation.\<close>
lemma learn_lim_closed_subseteq:
assumes "learn_lim \<psi> U s" and "V \<subseteq> U"
shows "learn_lim \<psi> V s"
using assms learn_lim_def by auto
corollary LIM_closed_subseteq:
assumes "U \<in> LIM" and "V \<subseteq> U"
shows "V \<in> LIM"
using assms learn_lim_closed_subseteq by (smt Lim_def mem_Collect_eq)
text \<open>Changing the hypothesis infinitely often precludes learning in
the limit.\<close>
lemma infinite_hyp_changes_not_Lim:
assumes "f \<in> U" and "\<forall>n. \<exists>m\<^sub>1>n. \<exists>m\<^sub>2>n. s (f \<triangleright> m\<^sub>1) \<noteq> s (f \<triangleright> m\<^sub>2)"
shows "\<not> learn_lim \<psi> U s"
using assms learn_lim_def by (metis less_imp_le)
lemma always_hyp_change_not_Lim:
assumes "\<And>x. s (f \<triangleright> (Suc x)) \<noteq> s (f \<triangleright> x)"
shows "\<not> learn_lim \<psi> {f} s"
using assms learn_limE by (metis le_SucI order_refl singletonI)
text \<open>Guessing a wrong hypothesis infinitely often precludes learning
in the limit.\<close>
lemma infinite_hyp_wrong_not_Lim:
assumes "f \<in> U" and "\<forall>n. \<exists>m>n. \<psi> (the (s (f \<triangleright> m))) \<noteq> f"
shows "\<not> learn_lim \<psi> U s"
using assms learn_limE by (metis less_imp_le option.sel)
text \<open>Converging to the same hypothesis on two functions precludes
learning in the limit.\<close>
lemma same_hyp_for_two_not_Lim:
assumes "f\<^sub>1 \<in> U"
and "f\<^sub>2 \<in> U"
and "f\<^sub>1 \<noteq> f\<^sub>2"
and "\<forall>n\<ge>n\<^sub>1. s (f\<^sub>1 \<triangleright> n) = h"
and "\<forall>n\<ge>n\<^sub>2. s (f\<^sub>2 \<triangleright> n) = h"
shows "\<not> learn_lim \<psi> U s"
using assms learn_limE by (metis le_cases option.sel)
text \<open>Every class that can be learned in the limit can be learned in
the limit with respect to any Gödel numbering. We prove a generalization in
which hypotheses may have to satisfy an extra condition, so we can re-use it
for other inference types later.\<close>
lemma learn_lim_extra_wrt_goedel:
fixes extra :: "(partial1 set) \<Rightarrow> partial1 \<Rightarrow> nat \<Rightarrow> partial1 \<Rightarrow> bool"
assumes "goedel_numbering \<chi>"
and "learn_lim \<psi> U s"
and "\<And>f n. f \<in> U \<Longrightarrow> extra U f n (\<psi> (the (s (f \<triangleright> n))))"
shows "\<exists>t. learn_lim \<chi> U t \<and> (\<forall>f\<in>U. \<forall>n. extra U f n (\<chi> (the (t (f \<triangleright> n)))))"
proof -
have env: "environment \<psi> U s"
and lim: "learn_lim \<psi> U s"
and extra: "\<forall>f\<in>U. \<forall>n. extra U f n (\<psi> (the (s (f \<triangleright> n))))"
using assms learn_limE by auto
obtain c where c: "c \<in> \<R>" "\<forall>i. \<psi> i = \<chi> (the (c i))"
using env goedel_numberingE[OF assms(1), of \<psi>] by auto
define t where "t \<equiv>
(\<lambda>x. if s x \<down> \<and> c (the (s x)) \<down> then Some (the (c (the (s x)))) else None)"
have "t \<in> \<P>"
unfolding t_def using env c concat_P1_P1[of c s] by auto
have "t x = (if s x \<down> then Some (the (c (the (s x)))) else None)" for x
using t_def c(1) R1_imp_total1 by auto
then have t: "t (f \<triangleright> n) \<down>= the (c (the (s (f \<triangleright> n))))" if "f \<in> U" for f n
using lim learn_limE that by simp
have "learn_lim \<chi> U t"
proof (rule learn_limI)
show "environment \<chi> U t"
using t by (simp add: \<open>t \<in> \<P>\<close> env goedel_numbering_P2[OF assms(1)])
show "\<exists>i. \<chi> i = f \<and> (\<forall>\<^sup>\<infinity>n. t (f \<triangleright> n) \<down>= i)" if "f \<in> U" for f
proof -
from lim learn_limE(2) obtain i n\<^sub>0 where
i: "\<psi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= i)"
using \<open>f \<in> U\<close> by blast
let ?j = "the (c i)"
have "\<chi> ?j = f"
using c(2) i by simp
moreover have "t (f \<triangleright> n) \<down>= ?j" if "n \<ge> n\<^sub>0" for n
by (simp add: \<open>f \<in> U\<close> i t that)
ultimately show ?thesis by auto
qed
qed
moreover have "extra U f n (\<chi> (the (t (f \<triangleright> n))))" if "f \<in> U" for f n
proof -
from t have "the (t (f \<triangleright> n)) = the (c (the (s (f \<triangleright> n))))"
by (simp add: that)
then have "\<chi> (the (t (f \<triangleright> n))) = \<psi> (the (s (f \<triangleright> n)))"
using c(2) by simp
with extra show ?thesis using that by simp
qed
ultimately show ?thesis by auto
qed
lemma learn_lim_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_lim \<psi> U s"
shows "\<exists>t. learn_lim \<chi> U t"
using assms learn_lim_extra_wrt_goedel[where ?extra="\<lambda>U f n h. True"]
by simp
lemma LIM_wrt_phi_eq_Lim: "LIM_wrt \<phi> = LIM"
using LIM_wrt_def Lim_def learn_lim_wrt_goedel[OF goedel_numbering_phi]
by blast
subsection \<open>BC: Behaviorally correct learning in the limit\<close>
text \<open>Behaviorally correct learning in the limit relaxes LIM by
requiring that the strategy almost always output an index for the target
function, but not necessarily the same index. In other words convergence of
$(S(f^n))_{n\in\mathbb{N}}$ is replaced by convergence of
$(\psi_{S(f^n)})_{n\in\mathbb{N}}$.\<close>
definition learn_bc :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_bc \<psi> U s \<equiv>
environment \<psi> U s \<and>
(\<forall>f\<in>U. \<forall>\<^sup>\<infinity>n. \<psi> (the (s (f \<triangleright> n))) = f)"
lemma learn_bcE:
assumes "learn_bc \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<forall>\<^sup>\<infinity>n. \<psi> (the (s (f \<triangleright> n))) = f"
using assms learn_bc_def by auto
lemma learn_bcI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<forall>\<^sup>\<infinity>n. \<psi> (the (s (f \<triangleright> n))) = f"
shows "learn_bc \<psi> U s"
using assms learn_bc_def by auto
definition BC_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"BC_wrt \<psi> \<equiv> {U. \<exists>s. learn_bc \<psi> U s}"
definition BC :: "partial1 set set" where
"BC \<equiv> {U. \<exists>\<psi> s. learn_bc \<psi> U s}"
text \<open>BC is a superset of LIM and closed under the subset relation.\<close>
lemma learn_lim_imp_BC: "learn_lim \<psi> U s \<Longrightarrow> learn_bc \<psi> U s"
using learn_limE learn_bcI[of \<psi> U s] by fastforce
lemma Lim_subseteq_BC: "LIM \<subseteq> BC"
using learn_lim_imp_BC Lim_def BC_def by blast
lemma learn_bc_closed_subseteq:
assumes "learn_bc \<psi> U s" and "V \<subseteq> U"
shows "learn_bc \<psi> V s"
using assms learn_bc_def by auto
corollary BC_closed_subseteq:
assumes "U \<in> BC" and "V \<subseteq> U"
shows "V \<in> BC"
using assms by (smt BC_def learn_bc_closed_subseteq mem_Collect_eq)
text \<open>Just like with LIM, guessing a wrong hypothesis infinitely often
precludes BC-style learning.\<close>
lemma infinite_hyp_wrong_not_BC:
assumes "f \<in> U" and "\<forall>n. \<exists>m>n. \<psi> (the (s (f \<triangleright> m))) \<noteq> f"
shows "\<not> learn_bc \<psi> U s"
proof
assume "learn_bc \<psi> U s"
then obtain n\<^sub>0 where "\<forall>n\<ge>n\<^sub>0. \<psi> (the (s (f \<triangleright> n))) = f"
using learn_bcE assms(1) by metis
with assms(2) show False using less_imp_le by blast
qed
text \<open>The proof that Gödel numberings suffice as hypothesis spaces for
BC is similar to the one for @{thm[source] learn_lim_extra_wrt_goedel}. We do
not need the @{term extra} part for BC, but we get it for free.\<close>
lemma learn_bc_extra_wrt_goedel:
fixes extra :: "(partial1 set) \<Rightarrow> partial1 \<Rightarrow> nat \<Rightarrow> partial1 \<Rightarrow> bool"
assumes "goedel_numbering \<chi>"
and "learn_bc \<psi> U s"
and "\<And>f n. f \<in> U \<Longrightarrow> extra U f n (\<psi> (the (s (f \<triangleright> n))))"
shows "\<exists>t. learn_bc \<chi> U t \<and> (\<forall>f\<in>U. \<forall>n. extra U f n (\<chi> (the (t (f \<triangleright> n)))))"
proof -
have env: "environment \<psi> U s"
and lim: "learn_bc \<psi> U s"
and extra: "\<forall>f\<in>U. \<forall>n. extra U f n (\<psi> (the (s (f \<triangleright> n))))"
using assms learn_bc_def by auto
obtain c where c: "c \<in> \<R>" "\<forall>i. \<psi> i = \<chi> (the (c i))"
using env goedel_numberingE[OF assms(1), of \<psi>] by auto
define t where
"t = (\<lambda>x. if s x \<down> \<and> c (the (s x)) \<down> then Some (the (c (the (s x)))) else None)"
have "t \<in> \<P>"
unfolding t_def using env c concat_P1_P1[of c s] by auto
have "t x = (if s x \<down> then Some (the (c (the (s x)))) else None)" for x
using t_def c(1) R1_imp_total1 by auto
then have t: "t (f \<triangleright> n) \<down>= the (c (the (s (f \<triangleright> n))))" if "f \<in> U" for f n
using lim learn_bcE(1) that by simp
have "learn_bc \<chi> U t"
proof (rule learn_bcI)
show "environment \<chi> U t"
using t by (simp add: \<open>t \<in> \<P>\<close> env goedel_numbering_P2[OF assms(1)])
show "\<forall>\<^sup>\<infinity>n. \<chi> (the (t (f \<triangleright> n))) = f" if "f \<in> U" for f
proof -
obtain n\<^sub>0 where "\<forall>n\<ge>n\<^sub>0. \<psi> (the (s (f \<triangleright> n))) = f"
using lim learn_bcE(2) \<open>f \<in> U\<close> by blast
then show ?thesis using that t c(2) by auto
qed
qed
moreover have "extra U f n (\<chi> (the (t (f \<triangleright> n))))" if "f \<in> U" for f n
proof -
from t have "the (t (f \<triangleright> n)) = the (c (the (s (f \<triangleright> n))))"
by (simp add: that)
then have "\<chi> (the (t (f \<triangleright> n))) = \<psi> (the (s (f \<triangleright> n)))"
using c(2) by simp
with extra show ?thesis using that by simp
qed
ultimately show ?thesis by auto
qed
corollary learn_bc_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_bc \<psi> U s"
shows "\<exists>t. learn_bc \<chi> U t"
using assms learn_bc_extra_wrt_goedel[where ?extra="\<lambda>_ _ _ _. True"] by simp
corollary BC_wrt_phi_eq_BC: "BC_wrt \<phi> = BC"
using learn_bc_wrt_goedel goedel_numbering_phi BC_def BC_wrt_def by blast
subsection \<open>CONS: Learning in the limit with consistent hypotheses\<close>
text \<open>A hypothesis is \emph{consistent} if it matches all values in the
prefix given to the strategy. Consistent learning in the limit requires the
strategy to output only consistent hypotheses for prefixes from the class.\<close>
definition learn_cons :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_cons \<psi> U s \<equiv>
learn_lim \<psi> U s \<and>
(\<forall>f\<in>U. \<forall>n. \<forall>k\<le>n. \<psi> (the (s (f \<triangleright> n))) k = f k)"
definition CONS_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"CONS_wrt \<psi> \<equiv> {U. \<exists>s. learn_cons \<psi> U s}"
definition CONS :: "partial1 set set" where
"CONS \<equiv> {U. \<exists>\<psi> s. learn_cons \<psi> U s}"
lemma CONS_subseteq_Lim: "CONS \<subseteq> LIM"
using CONS_def Lim_def learn_cons_def by blast
lemma learn_consI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<forall>k\<le>n. \<psi> (the (s (f \<triangleright> n))) k = f k"
shows "learn_cons \<psi> U s"
using assms learn_lim_def learn_cons_def by simp
text \<open>If a consistent strategy converges, it automatically converges to
a correct hypothesis. Thus we can remove @{term "\<psi> i = f"} from the second
assumption in the previous lemma.\<close>
lemma learn_consI2:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i"
and "\<And>f n. f \<in> U \<Longrightarrow> \<forall>k\<le>n. \<psi> (the (s (f \<triangleright> n))) k = f k"
shows "learn_cons \<psi> U s"
proof (rule learn_consI)
show "environment \<psi> U s"
and cons: "\<And>f n. f \<in> U \<Longrightarrow> \<forall>k\<le>n. \<psi> (the (s (f \<triangleright> n))) k = f k"
using assms by simp_all
show "\<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)" if "f \<in> U" for f
proof -
from that assms(2) obtain i n\<^sub>0 where i_n0: "\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= i"
by blast
have "\<psi> i x = f x" for x
proof (cases "x \<le> n\<^sub>0")
case True
then show ?thesis
using i_n0 cons that by fastforce
next
case False
moreover have "\<forall>k\<le>x. \<psi> (the (s (f \<triangleright> x))) k = f k"
using cons that by simp
ultimately show ?thesis using i_n0 by simp
qed
with i_n0 show ?thesis by auto
qed
qed
lemma learn_consE:
assumes "learn_cons \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<forall>k\<le>n. \<psi> (the (s (f \<triangleright> n))) k = f k"
using assms learn_cons_def learn_lim_def by auto
lemma learn_cons_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_cons \<psi> U s"
shows "\<exists>t. learn_cons \<chi> U t"
using learn_cons_def assms
learn_lim_extra_wrt_goedel[where ?extra="\<lambda>U f n h. \<forall>k\<le>n. h k = f k"]
by auto
lemma CONS_wrt_phi_eq_CONS: "CONS_wrt \<phi> = CONS"
using CONS_wrt_def CONS_def learn_cons_wrt_goedel goedel_numbering_phi
by blast
lemma learn_cons_closed_subseteq:
assumes "learn_cons \<psi> U s" and "V \<subseteq> U"
shows "learn_cons \<psi> V s"
using assms learn_cons_def learn_lim_closed_subseteq by auto
lemma CONS_closed_subseteq:
assumes "U \<in> CONS" and "V \<subseteq> U"
shows "V \<in> CONS"
using assms learn_cons_closed_subseteq by (smt CONS_def mem_Collect_eq)
text \<open>A consistent strategy cannot output the same hypothesis for two
different prefixes from the class to be learned.\<close>
lemma same_hyp_different_init_not_cons:
assumes "f \<in> U"
and "g \<in> U"
and "f \<triangleright> n \<noteq> g \<triangleright> n"
and "s (f \<triangleright> n) = s (g \<triangleright> n)"
shows "\<not> learn_cons \<phi> U s"
unfolding learn_cons_def by (auto, metis assms init_eqI)
subsection \<open>TOTAL: Learning in the limit with total hypotheses\<close>
text \<open>Total learning in the limit requires the strategy to hypothesize
only total functions for prefixes from the class.\<close>
definition learn_total :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_total \<psi> U s \<equiv>
learn_lim \<psi> U s \<and>
(\<forall>f\<in>U. \<forall>n. \<psi> (the (s (f \<triangleright> n))) \<in> \<R>)"
definition TOTAL_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"TOTAL_wrt \<psi> \<equiv> {U. \<exists>s. learn_total \<psi> U s}"
definition TOTAL :: "partial1 set set" where
"TOTAL \<equiv> {U. \<exists>\<psi> s. learn_total \<psi> U s}"
lemma TOTAL_subseteq_LIM: "TOTAL \<subseteq> LIM"
unfolding TOTAL_def Lim_def using learn_total_def by auto
lemma learn_totalI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<psi> (the (s (f \<triangleright> n))) \<in> \<R>"
shows "learn_total \<psi> U s"
using assms learn_lim_def learn_total_def by auto
lemma learn_totalE:
assumes "learn_total \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<psi> (the (s (f \<triangleright> n))) \<in> \<R>"
using assms learn_lim_def learn_total_def by auto
lemma learn_total_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_total \<psi> U s"
shows "\<exists>t. learn_total \<chi> U t"
using learn_total_def assms learn_lim_extra_wrt_goedel[where ?extra="\<lambda>U f n h. h \<in> \<R>"]
by auto
lemma TOTAL_wrt_phi_eq_TOTAL: "TOTAL_wrt \<phi> = TOTAL"
using TOTAL_wrt_def TOTAL_def learn_total_wrt_goedel goedel_numbering_phi
by blast
lemma learn_total_closed_subseteq:
assumes "learn_total \<psi> U s" and "V \<subseteq> U"
shows "learn_total \<psi> V s"
using assms learn_total_def learn_lim_closed_subseteq by auto
lemma TOTAL_closed_subseteq:
assumes "U \<in> TOTAL" and "V \<subseteq> U"
shows "V \<in> TOTAL"
using assms learn_total_closed_subseteq by (smt TOTAL_def mem_Collect_eq)
subsection \<open>CP: Learning in the limit with class-preserving hypotheses\<close>
text \<open>Class-preserving learning in the limit requires all hypotheses
for prefixes from the class to be functions from the class.\<close>
definition learn_cp :: "partial2 \<Rightarrow> (partial1 set) \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_cp \<psi> U s \<equiv>
learn_lim \<psi> U s \<and>
(\<forall>f\<in>U. \<forall>n. \<psi> (the (s (f \<triangleright> n))) \<in> U)"
definition CP_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"CP_wrt \<psi> \<equiv> {U. \<exists>s. learn_cp \<psi> U s}"
definition CP :: "partial1 set set" where
"CP \<equiv> {U. \<exists>\<psi> s. learn_cp \<psi> U s}"
lemma learn_cp_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_cp \<psi> U s"
shows "\<exists>t. learn_cp \<chi> U t"
using learn_cp_def assms learn_lim_extra_wrt_goedel[where ?extra="\<lambda>U f n h. h \<in> U"]
by auto
corollary CP_wrt_phi: "CP = CP_wrt \<phi>"
using learn_cp_wrt_goedel[OF goedel_numbering_phi]
by (smt CP_def CP_wrt_def Collect_cong)
lemma learn_cpI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. \<psi> i = f \<and> (\<forall>\<^sup>\<infinity>n. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<psi> (the (s (f \<triangleright> n))) \<in> U"
shows "learn_cp \<psi> U s"
using assms learn_cp_def learn_lim_def by auto
lemma learn_cpE:
assumes "learn_cp \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow> \<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> U \<Longrightarrow> \<psi> (the (s (f \<triangleright> n))) \<in> U"
using assms learn_lim_def learn_cp_def by auto
text \<open>Since classes contain only total functions, a CP strategy is also
a TOTAL strategy.\<close>
lemma learn_cp_imp_total: "learn_cp \<psi> U s \<Longrightarrow> learn_total \<psi> U s"
using learn_cp_def learn_total_def learn_lim_def by auto
lemma CP_subseteq_TOTAL: "CP \<subseteq> TOTAL"
using learn_cp_imp_total CP_def TOTAL_def by blast
subsection \<open>FIN: Finite learning\<close>
text \<open>In general it is undecidable whether a LIM strategy has reached
its final hypothesis. By contrast, in finite learning (also called ``one-shot
learning'') the strategy signals when it is ready to output a hypothesis. Up
until then it outputs a ``don't know yet'' value. This value is represented
by zero and the actual hypothesis $i$ by $i + 1$.\<close>
definition learn_fin :: "partial2 \<Rightarrow> partial1 set \<Rightarrow> partial1 \<Rightarrow> bool" where
"learn_fin \<psi> U s \<equiv>
environment \<psi> U s \<and>
(\<forall>f \<in> U. \<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i))"
definition FIN_wrt :: "partial2 \<Rightarrow> partial1 set set" where
"FIN_wrt \<psi> \<equiv> {U. \<exists>s. learn_fin \<psi> U s}"
definition FIN :: "partial1 set set" where
"FIN \<equiv> {U. \<exists>\<psi> s. learn_fin \<psi> U s}"
lemma learn_finI:
assumes "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow>
\<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
shows "learn_fin \<psi> U s"
using assms learn_fin_def by auto
lemma learn_finE:
assumes "learn_fin \<psi> U s"
shows "environment \<psi> U s"
and "\<And>f. f \<in> U \<Longrightarrow>
\<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
using assms learn_fin_def by auto
lemma learn_fin_closed_subseteq:
assumes "learn_fin \<psi> U s" and "V \<subseteq> U"
shows "learn_fin \<psi> V s"
using assms learn_fin_def by auto
lemma learn_fin_wrt_goedel:
assumes "goedel_numbering \<chi>" and "learn_fin \<psi> U s"
shows "\<exists>t. learn_fin \<chi> U t"
proof -
have env: "environment \<psi> U s"
and fin: "\<And>f. f \<in> U \<Longrightarrow>
\<exists>i n\<^sub>0. \<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
using assms(2) learn_finE by auto
obtain c where c: "c \<in> \<R>" "\<forall>i. \<psi> i = \<chi> (the (c i))"
using env goedel_numberingE[OF assms(1), of \<psi>] by auto
define t where "t \<equiv>
\<lambda>x. if s x \<up> then None
else if s x = Some 0 then Some 0
else Some (Suc (the (c (the (s x) - 1))))"
have "t \<in> \<P>"
proof -
from c obtain rc where rc:
"recfn 1 rc"
"total rc"
"\<forall>x. c x = eval rc [x]"
by auto
from env obtain rs where rs: "recfn 1 rs" "\<forall>x. s x = eval rs [x]"
by auto
then have "eval rs [f \<triangleright> n] \<down>" if "f \<in> U" for f n
using env that by simp
define rt where "rt = Cn 1 r_ifz [rs, Z, Cn 1 S [Cn 1 rc [Cn 1 r_dec [rs]]]]"
then have "recfn 1 rt"
using rc(1) rs(1) by simp
have "eval rt [x] \<up>" if "eval rs [x] \<up>" for x
using rc(1) rs(1) rt_def that by auto
moreover have "eval rt [x] \<down>= 0" if "eval rs [x] \<down>= 0" for x
using rt_def that rc(1,2) rs(1) by simp
moreover have "eval rt [x] \<down>= Suc (the (c (the (s x) - 1)))" if "eval rs [x] \<down>\<noteq> 0" for x
using rt_def that rc rs by auto
ultimately have "eval rt [x] = t x" for x
by (simp add: rs(2) t_def)
- with `recfn 1 rt` show ?thesis by auto
+ with \<open>recfn 1 rt\<close> show ?thesis by auto
qed
have t: "t (f \<triangleright> n) \<down>=
(if s (f \<triangleright> n) = Some 0 then 0 else Suc (the (c (the (s (f \<triangleright> n)) - 1))))"
if "f \<in> U" for f n
using that env by (simp add: t_def)
have "learn_fin \<chi> U t"
proof (rule learn_finI)
show "environment \<chi> U t"
using t by (simp add: \<open>t \<in> \<P>\<close> env goedel_numbering_P2[OF assms(1)])
show "\<exists>i n\<^sub>0. \<chi> i = f \<and> (\<forall>n<n\<^sub>0. t (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. t (f \<triangleright> n) \<down>= Suc i)"
if "f \<in> U" for f
proof -
from fin obtain i n\<^sub>0 where
i: "\<psi> i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
using \<open>f \<in> U\<close> by blast
let ?j = "the (c i)"
have "\<chi> ?j = f"
using c(2) i by simp
moreover have "\<forall>n<n\<^sub>0. t (f \<triangleright> n) \<down>= 0"
using t[OF that] i by simp
moreover have "t (f \<triangleright> n) \<down>= Suc ?j" if "n \<ge> n\<^sub>0" for n
- using that i t[OF `f \<in> U`] by simp
+ using that i t[OF \<open>f \<in> U\<close>] by simp
ultimately show ?thesis by auto
qed
qed
then show ?thesis by auto
qed
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/LIM_BC.thy b/thys/Inductive_Inference/LIM_BC.thy
--- a/thys/Inductive_Inference/LIM_BC.thy
+++ b/thys/Inductive_Inference/LIM_BC.thy
@@ -1,1329 +1,1329 @@
section \<open>LIM is a proper subset of BC\label{s:lim_bc}\<close>
theory LIM_BC
imports Lemma_R
begin
text \<open>The proper inclusion of LIM in BC has been proved by
Barzdin~\cite{b-ttlsf-74} (see also Case and Smith~\cite{cs-cicmii-83}). The
proof constructs a class $V \in \mathrm{BC} - \mathrm{LIM}$ by
diagonalization against all LIM strategies. Exploiting Lemma~R for LIM, we
can assume that all such strategies are total functions. From the effective
version of this lemma we derive a numbering @{term "\<sigma> \<in>
\<R>\<^sup>2"} such that for all $U \in \mathrm{LIM}$ there is an $i$ with
$U\in \mathrm{LIM}_\varphi(\sigma_i)$. The idea behind $V$
is for every $i$ to construct a class $V_i$ of cardinality one or two such
that $V_i \notin \mathrm{LIM}_\varphi(\sigma_i)$. It then follows that the
union $V := \bigcup_i V_i$ cannot be learned by any $\sigma_i$ and thus $V
\notin \mathrm{LIM}$. At the same time, the construction ensures that the
functions in $V$ are ``predictable enough'' to be learnable in the BC sense.
At the core is a process that maintains a state $(b, k)$ of a list $b$ of
numbers and an index $k < |b|$ into this list. We imagine $b$ to be the
prefix of the function being constructed, except for position $k$ where
we imagine $b$ to have a ``gap''; that is, $b_k$ is not defined yet.
Technically, we will always have $b_k = 0$, so $b$ also represents the prefix
after the ``gap is filled'' with 0, whereas $b_{k:=1}$ represents the prefix
where the gap is filled with 1. For every $i \in \mathbb{N}$, the process
starts in state $(i0, 1)$ and computes the next state from a given state
$(b,k)$ as follows:
\begin{enumerate}
\item if $ \sigma_i(b_{<k}) \neq \sigma_i(b)$ then the next state is $(b0, |b|)$,
\item else if $\sigma_i(b_{<k}) \neq \sigma_i(b_{k:=1})$ then the next state is $(b_{k:=1}0, |b|)$,
\item else the next state is $(b0, k)$.
\end{enumerate}
In other words, if $\sigma_i$ changes its hypothesis when the gap in $b$ is
filled with 0 or 1, then the process fills the gap with 0 or 1, respectively, and
appends a gap to $b$. If, however, a hypothesis change cannot be enforced at
this point, the process appends a 0 to $b$ and leaves the gap alone. Now there
are two cases:
\begin{itemize}
\item[Case 1.] Every gap gets filled eventually. Then the process generates
increasing prefixes of a total function $\tau_i$, on which $\sigma_i$ changes
its hypothesis infinitely often. We set $V_i := \{\tau_i\}$, and have $V_i \notin
\mathrm{LIM}_\varphi(\sigma_i)$.
\item[Case 2.] Some gap never gets filled. That means a state $(b, k)$ is
reached such that $\sigma_i(b0^t) = \sigma_i(b_{k:=1}0^t) = \sigma_i(b_{<k})$
for all $t$. Then the process describes a function $\tau_i = b_{<k}\uparrow0^\infty$,
where the value at the gap $k$ is undefined.
Replacing the value at $k$ by 0 and 1 yields two functions
$\tau_i^{(0)} = b0^\infty$ and $\tau_i^{(1)} = b_{k:=1}0^\infty$, which differ only at $k$ and on
which $\sigma_i$ converges to the same hypothesis. Thus $\sigma_i$ does not learn the
class $V_i := \{\tau_i^{(0)}, \tau_i^{(1)}\}$ in the limit.
\end{itemize}
Both cases combined imply $V \notin \mathrm{LIM}$.
A BC strategy $S$ for $V = \bigcup_i V_i$ works as follows. Let $f\in V$. On input $f^n$ the
strategy outputs a Gödel number of the function
\[
g_n(x) = \left\{\begin{array}{ll}
f(x) & \mbox{if } x \leq n,\\
\tau_{f(0)}(x) & \mbox{otherwise}.
\end{array}\right.
\]
By definition of $V$, $f$ is generated by the process running for $i = f(0)$.
If $f(0)$ leads to Case~1 then $f = \tau_{f(0)}$, and
$g_n$ equals $f$ for all $n$. If $f(0)$ leads to Case~2 with
a forever unfilled gap at $k$, then
$g_n$ will be equal to the correct one of $\tau_i^{(0)}$ or $\tau_i^{(1)}$ for all $n
\geq k$. Intuitively, the prefix received by $S$ eventually grows long enough to
reveal the value $f(k)$.
In both cases $S$ converges to $f$, but it outputs a different Gödel number
for every $f^n$ because $g_n$ contains the ``hard-coded'' values
$f(0),\dots,f(n)$. Therefore $S$ is a BC strategy but not a LIM
strategy for $V$.\<close>
subsection \<open>Enumerating enough total strategies\<close>
text \<open>For the construction of $\sigma$ we need the function @{term
r_limr} from the effective version of Lemma~R for LIM.\<close>
definition "r_sigma \<equiv> Cn 2 r_phi [Cn 2 r_limr [Id 2 0], Id 2 1]"
lemma r_sigma_recfn: "recfn 2 r_sigma"
unfolding r_sigma_def using r_limr_recfn by simp
lemma r_sigma: "eval r_sigma [i, x] = \<phi> (the (eval r_limr [i])) x"
unfolding r_sigma_def phi_def using r_sigma_recfn r_limr_total r_limr_recfn
by simp
lemma r_sigma_total: "total r_sigma"
using r_sigma r_limr r_sigma_recfn totalI2[of r_sigma] by simp
abbreviation sigma :: partial2 ("\<sigma>") where
"\<sigma> i x \<equiv> eval r_sigma [i, x]"
lemma sigma: "\<sigma> i = \<phi> (the (eval r_limr [i]))"
using r_sigma by simp
text \<open>The numbering @{term \<sigma>} does indeed enumerate enough total
strategies for every LIM learning problem.\<close>
lemma learn_lim_sigma:
assumes "learn_lim \<psi> U (\<phi> i)"
shows "learn_lim \<psi> U (\<sigma> i)"
using assms sigma r_limr by simp
subsection \<open>The diagonalization process\<close>
text \<open>The following function represents the process described above. It
computes the next state from a given state $(b, k)$.\<close>
definition "r_next \<equiv>
Cn 1 r_ifeq
[Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], r_pdec1],
Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_take [r_pdec2, r_pdec1]],
Cn 1 r_ifeq
[Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_update [r_pdec1, r_pdec2, r_const 1]],
Cn 1 r_sigma [Cn 1 r_hd [r_pdec1], Cn 1 r_take [r_pdec2, r_pdec1]],
Cn 1 r_prod_encode [Cn 1 r_snoc [r_pdec1, Z], r_pdec2],
Cn 1 r_prod_encode
[Cn 1 r_snoc
[Cn 1 r_update [r_pdec1, r_pdec2, r_const 1], Z], Cn 1 r_length [r_pdec1]]],
Cn 1 r_prod_encode [Cn 1 r_snoc [r_pdec1, Z], Cn 1 r_length [r_pdec1]]]"
lemma r_next_recfn: "recfn 1 r_next"
unfolding r_next_def using r_sigma_recfn by simp
text \<open>The three conditions distinguished in @{term r_next} correspond
to Steps 1, 2, and 3 of the process: hypothesis change when the gap is
filled with 0; hypothesis change when the gap is filled with 1; or
no hypothesis change either way.\<close>
abbreviation "change_on_0 b k \<equiv> \<sigma> (e_hd b) b \<noteq> \<sigma> (e_hd b) (e_take k b)"
abbreviation "change_on_1 b k \<equiv>
\<sigma> (e_hd b) b = \<sigma> (e_hd b) (e_take k b) \<and>
\<sigma> (e_hd b) (e_update b k 1) \<noteq> \<sigma> (e_hd b) (e_take k b)"
abbreviation "change_on_neither b k \<equiv>
\<sigma> (e_hd b) b = \<sigma> (e_hd b) (e_take k b) \<and>
\<sigma> (e_hd b) (e_update b k 1) = \<sigma> (e_hd b) (e_take k b)"
lemma change_conditions:
obtains
(on_0) "change_on_0 b k"
| (on_1) "change_on_1 b k"
| (neither) "change_on_neither b k"
by auto
lemma r_next:
assumes "arg = prod_encode (b, k)"
shows "change_on_0 b k \<Longrightarrow> eval r_next [arg] \<down>= prod_encode (e_snoc b 0, e_length b)"
and "change_on_1 b k \<Longrightarrow>
eval r_next [arg] \<down>= prod_encode (e_snoc (e_update b k 1) 0, e_length b)"
and "change_on_neither b k \<Longrightarrow> eval r_next [arg] \<down>= prod_encode (e_snoc b 0, k)"
proof -
let ?bhd = "Cn 1 r_hd [r_pdec1]"
let ?bup = "Cn 1 r_update [r_pdec1, r_pdec2, r_const 1]"
let ?bk = "Cn 1 r_take [r_pdec2, r_pdec1]"
let ?bap = "Cn 1 r_snoc [r_pdec1, Z]"
let ?len = "Cn 1 r_length [r_pdec1]"
let ?thenthen = "Cn 1 r_prod_encode [?bap, r_pdec2]"
let ?thenelse = "Cn 1 r_prod_encode [Cn 1 r_snoc [?bup, Z], ?len]"
let ?else = "Cn 1 r_prod_encode [?bap, ?len]"
have bhd: "eval ?bhd [arg] \<down>= e_hd b"
using assms by simp
have bup: "eval ?bup [arg] \<down>= e_update b k 1"
using assms by simp
have bk: "eval ?bk [arg] \<down>= e_take k b"
using assms by simp
have bap: "eval ?bap [arg] \<down>= e_snoc b 0"
using assms by simp
have len: "eval ?len [arg] \<down>= e_length b"
using assms by simp
have else_: "eval ?else [arg] \<down>= prod_encode (e_snoc b 0, e_length b)"
using bap len by simp
have thenthen: "eval ?thenthen [arg] \<down>= prod_encode (e_snoc b 0, k)"
using bap assms by simp
have thenelse: "eval ?thenelse [arg] \<down>= prod_encode (e_snoc (e_update b k 1) 0, e_length b)"
using bup len by simp
have then_:
"eval
(Cn 1 r_ifeq [Cn 1 r_sigma [?bhd, ?bup], Cn 1 r_sigma [?bhd, ?bk], ?thenthen, ?thenelse])
[arg] \<down>=
(if the (\<sigma> (e_hd b) (e_update b k 1)) = the (\<sigma> (e_hd b) (e_take k b))
then prod_encode (e_snoc b 0, k)
else prod_encode (e_snoc (e_update b k 1) 0, e_length b))"
(is "eval ?then [arg] \<down>= ?then_eval")
using bhd bup bk thenthen thenelse r_sigma r_sigma_recfn r_limr R1_imp_total1 by simp
have *: "eval r_next [arg] \<down>=
(if the (\<sigma> (e_hd b) b) = the (\<sigma> (e_hd b) (e_take k b))
then ?then_eval
else prod_encode (e_snoc b 0, e_length b))"
unfolding r_next_def
using bhd bk then_ else_ r_sigma r_sigma_recfn r_limr R1_imp_total1 assms
by simp
have r_sigma_neq: "eval r_sigma [x\<^sub>1, y\<^sub>1] \<noteq> eval r_sigma [x\<^sub>2, y\<^sub>2] \<longleftrightarrow>
the (eval r_sigma [x\<^sub>1, y\<^sub>1]) \<noteq> the (eval r_sigma [x\<^sub>2, y\<^sub>2])"
for x\<^sub>1 y\<^sub>1 x\<^sub>2 y\<^sub>2
using r_sigma r_limr totalE[OF r_sigma_total r_sigma_recfn] r_sigma_recfn r_sigma_total
by (metis One_nat_def Suc_1 length_Cons list.size(3) option.expand)
{
assume "change_on_0 b k"
then show "eval r_next [arg] \<down>= prod_encode (e_snoc b 0, e_length b)"
using * r_sigma_neq by simp
next
assume "change_on_1 b k"
then show "eval r_next [arg] \<down>= prod_encode (e_snoc (e_update b k 1) 0, e_length b)"
using * r_sigma_neq by simp
next
assume "change_on_neither b k"
then show "eval r_next [arg] \<down>= prod_encode (e_snoc b 0, k)"
using * r_sigma_neq by simp
}
qed
lemma r_next_total: "total r_next"
proof (rule totalI1)
show "recfn 1 r_next"
using r_next_recfn by simp
show "eval r_next [x] \<down>" for x
proof -
obtain b k where "x = prod_encode (b, k)"
using prod_encode_pdec'[of x] by metis
then show ?thesis using r_next by fast
qed
qed
text \<open>The next function computes the state of the process after
any number of iterations.\<close>
definition "r_state \<equiv>
Pr 1
(Cn 1 r_prod_encode [Cn 1 r_snoc [Cn 1 r_singleton_encode [Id 1 0], Z], r_const 1])
(Cn 3 r_next [Id 3 1])"
lemma r_state_recfn: "recfn 2 r_state"
unfolding r_state_def using r_next_recfn by simp
lemma r_state_at_0: "eval r_state [0, i] \<down>= prod_encode (list_encode [i, 0], 1)"
proof -
let ?f = "Cn 1 r_prod_encode [Cn 1 r_snoc [Cn 1 r_singleton_encode [Id 1 0], Z], r_const 1]"
have "eval r_state [0, i] = eval ?f [i]"
unfolding r_state_def using r_next_recfn by simp
also have "... \<down>= prod_encode (list_encode [i, 0], 1)"
by (simp add: list_decode_singleton)
finally show ?thesis .
qed
lemma r_state_total: "total r_state"
unfolding r_state_def
using r_next_recfn totalE[OF r_next_total r_next_recfn] totalI3[of "Cn 3 r_next [Id 3 1]"]
by (intro Pr_total) auto
text \<open>We call the components of a state $(b, k)$ the \emph{block} $b$
and the \emph{gap} $k$.\<close>
definition block :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"block i t \<equiv> pdec1 (the (eval r_state [t, i]))"
definition gap :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"gap i t \<equiv> pdec2 (the (eval r_state [t, i]))"
lemma state_at_0:
"block i 0 = list_encode [i, 0]"
"gap i 0 = 1"
unfolding block_def gap_def r_state_at_0 by simp_all
text \<open>Some lemmas describing the behavior of blocks and gaps in
one iteration of the process:\<close>
lemma state_Suc:
assumes "b = block i t" and "k = gap i t"
shows "block i (Suc t) = pdec1 (the (eval r_next [prod_encode (b, k)]))"
and "gap i (Suc t) = pdec2 (the (eval r_next [prod_encode (b, k)]))"
proof -
have "eval r_state [Suc t, i] =
eval (Cn 3 r_next [Id 3 1]) [t, the (eval r_state [t, i]), i]"
using r_state_recfn r_next_recfn totalE[OF r_state_total r_state_recfn, of "[t, i]"]
by (simp add: r_state_def)
also have "... = eval r_next [the (eval r_state [t, i])]"
using r_next_recfn by simp
also have "... = eval r_next [prod_encode (b, k)]"
using assms block_def gap_def by simp
finally have "eval r_state [Suc t, i] = eval r_next [prod_encode (b, k)]" .
then show
"block i (Suc t) = pdec1 (the (eval r_next [prod_encode (b, k)]))"
"gap i (Suc t) = pdec2 (the (eval r_next [prod_encode (b, k)]))"
by (simp add: block_def, simp add: gap_def)
qed
lemma gap_Suc:
assumes "b = block i t" and "k = gap i t"
shows "change_on_0 b k \<Longrightarrow> gap i (Suc t) = e_length b"
and "change_on_1 b k \<Longrightarrow> gap i (Suc t) = e_length b"
and "change_on_neither b k\<Longrightarrow> gap i (Suc t) = k"
using assms r_next state_Suc by simp_all
lemma block_Suc:
assumes "b = block i t" and "k = gap i t"
shows "change_on_0 b k \<Longrightarrow> block i (Suc t) = e_snoc b 0"
and "change_on_1 b k \<Longrightarrow> block i (Suc t) = e_snoc (e_update b k 1) 0"
and "change_on_neither b k\<Longrightarrow> block i (Suc t) = e_snoc b 0"
using assms r_next state_Suc by simp_all
text \<open>Non-gap positions in the block remain unchanged after an
iteration.\<close>
lemma block_stable:
assumes "j < e_length (block i t)" and "j \<noteq> gap i t"
shows "e_nth (block i t) j = e_nth (block i (Suc t)) j"
proof -
from change_conditions[of "block i t" "gap i t"] show ?thesis
using assms block_Suc gap_Suc
by (cases, (simp_all add: nth_append))
qed
text \<open>Next are some properties of @{term block} and @{term gap}.\<close>
lemma gap_in_block: "gap i t < e_length (block i t)"
proof (induction t)
case 0
then show ?case by (simp add: state_at_0)
next
case (Suc t)
with change_conditions[of "block i t" "gap i t"] show ?case
proof (cases)
case on_0
then show ?thesis by (simp add: block_Suc(1) gap_Suc(1))
next
case on_1
then show ?thesis by (simp add: block_Suc(2) gap_Suc(2))
next
case neither
then show ?thesis using Suc.IH block_Suc(3) gap_Suc(3) by force
qed
qed
lemma length_block: "e_length (block i t) = Suc (Suc t)"
proof (induction t)
case 0
then show ?case by (simp add: state_at_0)
next
case (Suc t)
with change_conditions[of "block i t" "gap i t"] show ?case
by (cases, simp_all add: block_Suc gap_Suc)
qed
lemma gap_gr0: "gap i t > 0"
proof (induction t)
case 0
then show ?case by (simp add: state_at_0)
next
case (Suc t)
with change_conditions[of "block i t" "gap i t"] show ?case
using length_block by (cases, simp_all add: block_Suc gap_Suc)
qed
lemma hd_block: "e_hd (block i t) = i"
proof (induction t)
case 0
then show ?case by (simp add: state_at_0)
next
case (Suc t)
from change_conditions[of "block i t" "gap i t"] show ?case
proof (cases)
case on_0
then show ?thesis
using Suc block_Suc(1) length_block by (metis e_hd_snoc gap_Suc(1) gap_gr0)
next
case on_1
let ?b = "block i t" and ?k = "gap i t"
have "?k > 0"
using gap_gr0 Suc by simp
then have "e_nth (e_update ?b ?k 1) 0 = e_nth ?b 0"
by simp
then have *: "e_hd (e_update ?b ?k 1) = e_hd ?b"
using e_hd_nth0 gap_Suc(2)[of _ i t] gap_gr0 on_1 by (metis e_length_update)
from on_1 have "block i (Suc t) = e_snoc (e_update ?b ?k 1) 0"
by (simp add: block_Suc(2))
then show ?thesis
- using e_hd_0 e_hd_snoc Suc length_block `?k > 0` *
+ using e_hd_0 e_hd_snoc Suc length_block \<open>?k > 0\<close> *
by (metis e_length_update gap_Suc(2) gap_gr0 on_1)
next
case neither
then show ?thesis using Suc block_Suc(3) length_block by simp
qed
qed
text \<open>Formally, a block always ends in zero, even if it ends in a gap.\<close>
lemma last_block: "e_nth (block i t) (gap i t) = 0"
proof (induction t)
case 0
then show ?case by (simp add: state_at_0)
next
case (Suc t)
from change_conditions[of "block i t" "gap i t"] show ?case
proof cases
case on_0
then show ?thesis using Suc by (simp add: block_Suc(1) gap_Suc(1))
next
case on_1
then show ?thesis using Suc by (simp add: block_Suc(2) gap_Suc(2) nth_append)
next
case neither
then have
"block i (Suc t) = e_snoc (block i t) 0"
"gap i (Suc t) = gap i t"
by (simp_all add: gap_Suc(3) block_Suc(3))
then show ?thesis
using Suc gap_in_block by (simp add: nth_append)
qed
qed
lemma gap_le_Suc: "gap i t \<le> gap i (Suc t)"
using change_conditions[of "block i t" "gap i t"]
gap_Suc gap_in_block less_imp_le[of "gap i t" "e_length (block i t)"]
by (cases) simp_all
lemma gap_monotone:
assumes "t\<^sub>1 \<le> t\<^sub>2"
shows "gap i t\<^sub>1 \<le> gap i t\<^sub>2"
proof -
have "gap i t\<^sub>1 \<le> gap i (t\<^sub>1 + j)" for j
proof (induction j)
case 0
then show ?case by simp
next
case (Suc j)
then show ?case using gap_le_Suc dual_order.trans by fastforce
qed
then show ?thesis using assms le_Suc_ex by blast
qed
text \<open>We need some lemmas relating the shape of the next state
to the hypothesis change conditions in Steps 1, 2, and 3.\<close>
lemma state_change_on_neither:
assumes "gap i (Suc t) = gap i t"
shows "change_on_neither (block i t) (gap i t)"
and "block i (Suc t) = e_snoc (block i t) 0"
proof -
let ?b = "block i t" and ?k = "gap i t"
have "?k < e_length ?b"
using gap_in_block by simp
from change_conditions[of ?b ?k] show "change_on_neither (block i t) (gap i t)"
proof (cases)
case on_0
then show ?thesis
using \<open>?k < e_length ?b\<close> assms gap_Suc(1) by auto
next
case on_1
then show ?thesis using assms gap_Suc(2) by auto
next
case neither
then show ?thesis by simp
qed
then show "block i (Suc t) = e_snoc (block i t) 0"
using block_Suc(3) by simp
qed
lemma state_change_on_either:
assumes "gap i (Suc t) \<noteq> gap i t"
shows "\<not> change_on_neither (block i t) (gap i t)"
and "gap i (Suc t) = e_length (block i t)"
proof -
let ?b = "block i t" and ?k = "gap i t"
show "\<not> change_on_neither (block i t) (gap i t)"
proof
assume "change_on_neither (block i t) (gap i t)"
then have "gap i (Suc t) = ?k"
by (simp add: gap_Suc(3))
with assms show False by simp
qed
then show "gap i (Suc t) = e_length (block i t)"
using gap_Suc(1) gap_Suc(2) by blast
qed
text \<open>Next up is the definition of $\tau$. In every iteration the
process determines $\tau_i(x)$ for some $x$ either by appending 0 to the
current block $b$, or by filling the current gap $k$. In the former case,
the value is determined for $x = |b|$, in the latter for $x = k$.\<close>
text \<open>For $i$ and $x$ the function @{term r_dettime} computes in which
iteration the process for $i$ determines the value $\tau_i(x)$. This is the
first iteration in which the block is long enough to contain position $x$ and
in which $x$ is not the gap. If $\tau_i(x)$ is never determined, because Case~2 is
reached with $k = x$, then @{term r_dettime} diverges.\<close>
abbreviation determined :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
"determined i x \<equiv> \<exists>t. x < e_length (block i t) \<and> x \<noteq> gap i t"
lemma determined_0: "determined i 0"
using gap_gr0[of i 0] gap_in_block[of i 0] by force
definition "r_dettime \<equiv>
Mn 2
(Cn 3 r_and
[Cn 3 r_less
[Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]],
Cn 3 r_neq
[Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]]])"
lemma r_dettime_recfn: "recfn 2 r_dettime"
unfolding r_dettime_def using r_state_recfn by simp
abbreviation dettime :: partial2 where
"dettime i x \<equiv> eval r_dettime [i, x]"
lemma r_dettime:
shows "determined i x \<Longrightarrow> dettime i x \<down>= (LEAST t. x < e_length (block i t) \<and> x \<noteq> gap i t)"
and "\<not> determined i x \<Longrightarrow> dettime i x \<up>"
proof -
define f where "f =
(Cn 3 r_and
[Cn 3 r_less
[Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]],
Cn 3 r_neq
[Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]]])"
then have "r_dettime = Mn 2 f"
unfolding f_def r_dettime_def by simp
have "recfn 3 f"
unfolding f_def using r_state_recfn by simp
then have "total f"
unfolding f_def using Cn_total r_state_total Mn_free_imp_total by simp
have f: "eval f [t, i, x] \<down>= (if x < e_length (block i t) \<and> x \<noteq> gap i t then 0 else 1)" for t
proof -
let ?b = "Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]"
let ?k = "Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]"
have "eval ?b [t, i, x] \<down>= pdec1 (the (eval r_state [t, i]))"
using r_state_recfn r_state_total by simp
then have b: "eval ?b [t, i, x] \<down>= block i t"
using block_def by simp
have "eval ?k [t, i, x] \<down>= pdec2 (the (eval r_state [t, i]))"
using r_state_recfn r_state_total by simp
then have k: "eval ?k [t, i, x] \<down>= gap i t"
using gap_def by simp
have "eval
(Cn 3 r_neq [Id 3 2, Cn 3 r_pdec2 [Cn 3 r_state [Id 3 0, Id 3 1]]])
[t, i, x] \<down>=
(if x \<noteq> gap i t then 0 else 1)"
using b k r_state_recfn r_state_total by simp
moreover have "eval
(Cn 3 r_less
[Id 3 2, Cn 3 r_length [Cn 3 r_pdec1 [Cn 3 r_state [Id 3 0, Id 3 1]]]])
[t, i, x] \<down>=
(if x < e_length (block i t) then 0 else 1)"
using b k r_state_recfn r_state_total by simp
ultimately show ?thesis
unfolding f_def using b k r_state_recfn r_state_total by simp
qed
{
assume "determined i x"
with f have "\<exists>t. eval f [t, i, x] \<down>= 0" by simp
then have "dettime i x \<down>= (LEAST t. eval f [t, i, x] \<down>= 0)"
- using `total f` `r_dettime = Mn 2 f` r_dettime_recfn `recfn 3 f`
+ using \<open>total f\<close> \<open>r_dettime = Mn 2 f\<close> r_dettime_recfn \<open>recfn 3 f\<close>
eval_Mn_total[of 2 f "[i, x]"]
by simp
then show "dettime i x \<down>= (LEAST t. x < e_length (block i t) \<and> x \<noteq> gap i t)"
using f by simp
next
assume "\<not> determined i x"
with f have "\<not> (\<exists>t. eval f [t, i, x] \<down>= 0)" by simp
then have "dettime i x \<up>"
- using `total f` `r_dettime = Mn 2 f` r_dettime_recfn `recfn 3 f`
+ using \<open>total f\<close> \<open>r_dettime = Mn 2 f\<close> r_dettime_recfn \<open>recfn 3 f\<close>
eval_Mn_total[of 2 f "[i, x]"]
by simp
with f show "dettime i x \<up>" by simp
}
qed
lemma r_dettimeI:
assumes "x < e_length (block i t) \<and> x \<noteq> gap i t"
and "\<And>T. x < e_length (block i T) \<and> x \<noteq> gap i T \<Longrightarrow> t \<le> T"
shows "dettime i x \<down>= t"
proof -
let ?P = "\<lambda>T. x < e_length (block i T) \<and> x \<noteq> gap i T"
have "determined i x"
using assms(1) by auto
moreover have "Least ?P = t"
using assms Least_equality[of ?P t] by simp
ultimately show ?thesis using r_dettime by simp
qed
lemma r_dettime_0: "dettime i 0 \<down>= 0"
using r_dettimeI[of _ i 0] determined_0 gap_gr0[of i 0] gap_in_block[of i 0]
by fastforce
text \<open>Computing the value of $\tau_i(x)$ works by running the process
@{term r_state} for @{term "dettime i x"} iterations and taking the value at
index $x$ of the resulting block.\<close>
definition "r_tau \<equiv> Cn 2 r_nth [Cn 2 r_pdec1 [Cn 2 r_state [r_dettime, Id 2 0]], Id 2 1]"
lemma r_tau_recfn: "recfn 2 r_tau"
unfolding r_tau_def using r_dettime_recfn r_state_recfn by simp
abbreviation tau :: partial2 ("\<tau>") where
"\<tau> i x \<equiv> eval r_tau [i, x]"
lemma tau_in_P2: "\<tau> \<in> \<P>\<^sup>2"
using r_tau_recfn by auto
lemma tau_diverg:
assumes "\<not> determined i x"
shows "\<tau> i x \<up>"
unfolding r_tau_def using assms r_dettime r_dettime_recfn r_state_recfn by simp
lemma tau_converg:
assumes "determined i x"
shows "\<tau> i x \<down>= e_nth (block i (the (dettime i x))) x"
proof -
from assms obtain t where t: "dettime i x \<down>= t"
using r_dettime(1) by blast
then have "eval (Cn 2 r_state [r_dettime, Id 2 0]) [i, x] = eval r_state [t, i]"
using r_state_recfn r_dettime_recfn by simp
moreover have "eval r_state [t, i] \<down>"
using r_state_total r_state_recfn by simp
ultimately have "eval (Cn 2 r_pdec1 [Cn 2 r_state [r_dettime, Id 2 0]]) [i, x] =
eval r_pdec1 [the (eval r_state [t, i])]"
using r_state_recfn r_dettime_recfn by simp
then show ?thesis
unfolding r_tau_def using r_state_recfn r_dettime_recfn t block_def by simp
qed
lemma tau_converg':
assumes "dettime i x \<down>= t"
shows "\<tau> i x \<down>= e_nth (block i t) x"
using assms tau_converg[of x i] r_dettime(2)[of x i] by fastforce
lemma tau_at_0: "\<tau> i 0 \<down>= i"
proof -
have "\<tau> i 0 \<down>= e_nth (block i 0) 0"
using tau_converg'[OF r_dettime_0] by simp
then show ?thesis using block_def by (simp add: r_state_at_0)
qed
lemma state_unchanged:
assumes "gap i t - 1 \<le> y" and "y \<le> t"
shows "gap i t = gap i y"
proof -
have "gap i t = gap i (gap i t - 1)"
proof (induction t)
case 0
then show ?case by (simp add: gap_def r_state_at_0)
next
case (Suc t)
show ?case
proof (cases "gap i (Suc t) = t + 2")
case True
then show ?thesis by simp
next
case False
then show ?thesis
using Suc state_change_on_either(2) length_block by force
qed
qed
moreover have "gap i (gap i t - 1) \<le> gap i y"
using assms(1) gap_monotone by simp
moreover have "gap i y \<le> gap i t"
using assms(2) gap_monotone by simp
ultimately show ?thesis by simp
qed
text \<open>The values of the non-gap indices $x$ of every block created in
the diagonalization process equal $\tau_i(x)$.\<close>
lemma tau_eq_state:
assumes "j < e_length (block i t)" and "j \<noteq> gap i t"
shows "\<tau> i j \<down>= e_nth (block i t) j"
using assms
proof (induction t)
case 0
then have "j = 0"
using gap_gr0[of i 0] gap_in_block[of i 0] length_block[of i 0] by simp
then have "\<tau> (e_hd (block i t)) j \<down>= e_nth (block i (the (dettime i 0))) 0"
using determined_0 tau_converg hd_block by simp
then have "\<tau> (e_hd (block i t)) j \<down>= e_nth (block i 0) 0"
using r_dettime_0 by simp
then show ?case using \<open>j = 0\<close> r_dettime_0 tau_converg' by simp
next
case (Suc t)
let ?b = "block i t"
let ?bb = "block i (Suc t)"
let ?k = "gap i t"
let ?kk = "gap i (Suc t)"
show ?case
proof (cases "?kk = ?k")
case kk_eq_k: True
then have bb_b0: "?bb = e_snoc ?b 0"
using state_change_on_neither by simp
show "\<tau> i j \<down>= e_nth ?bb j"
proof (cases "j < e_length ?b")
case True
then have "e_nth ?bb j = e_nth ?b j"
using bb_b0 by (simp add: nth_append)
moreover have "j \<noteq> ?k"
using Suc kk_eq_k by simp
ultimately show ?thesis using Suc True by simp
next
case False
then have j: "j = e_length ?b"
using Suc.prems(1) length_block by auto
then have "e_nth ?bb j = 0"
using bb_b0 by simp
have "dettime i j \<down>= Suc t"
proof (rule r_dettimeI)
show "j < e_length ?bb \<and> j \<noteq> ?kk"
using Suc.prems(1,2) by linarith
show "\<And>T. j < e_length (block i T) \<and> j \<noteq> gap i T \<Longrightarrow> Suc t \<le> T"
using length_block j by simp
qed
with tau_converg' show ?thesis by simp
qed
next
case False
then have kk_lenb: "?kk = e_length ?b"
using state_change_on_either by simp
then show ?thesis
proof (cases "j = ?k")
case j_eq_k: True
have "dettime i j \<down>= Suc t"
proof (rule r_dettimeI)
show "j < e_length ?bb \<and> j \<noteq> ?kk"
using Suc.prems(1,2) by simp
show "Suc t \<le> T" if "j < e_length (block i T) \<and> j \<noteq> gap i T" for T
proof (rule ccontr)
assume "\<not> (Suc t \<le> T)"
then have "T < Suc t" by simp
then show False
proof (cases "T < ?k - 1")
case True
then have "e_length (block i T) = T + 2"
using length_block by simp
then have "e_length (block i T) < ?k + 1"
using True by simp
then have "e_length (block i T) \<le> ?k" by simp
then have "e_length (block i T) \<le> j"
using j_eq_k by simp
then show False
using that by simp
next
case False
then have "?k - 1 \<le> T" and "T \<le> t"
- using `T < Suc t` by simp_all
+ using \<open>T < Suc t\<close> by simp_all
with state_unchanged have "gap i t = gap i T" by blast
then show False
using j_eq_k that by simp
qed
qed
qed
then show ?thesis using tau_converg' by simp
next
case False
then have "j < e_length ?b"
using kk_lenb Suc.prems(1,2) length_block by auto
then show ?thesis using Suc False block_stable by fastforce
qed
qed
qed
lemma tau_eq_state':
assumes "j < t + 2" and " j \<noteq> gap i t"
shows "\<tau> i j \<down>= e_nth (block i t) j"
using assms tau_eq_state length_block by simp
text \<open>We now consider the two cases described in the proof sketch.
In Case~2 there is a gap that never gets filled, or equivalently there is
a rightmost gap.\<close>
abbreviation "case_two i \<equiv> (\<exists>t. \<forall>T. gap i T \<le> gap i t)"
abbreviation "case_one i \<equiv> \<not> case_two i"
text \<open>Another characterization of Case~2 is that from some iteration on
only @{term change_on_neither} holds.\<close>
lemma case_two_iff_forever_neither:
"case_two i \<longleftrightarrow> (\<exists>t. \<forall>T\<ge>t. change_on_neither (block i T) (gap i T))"
proof
assume "\<exists>t. \<forall>T\<ge>t. change_on_neither (block i T) (gap i T)"
then obtain t where t: "\<forall>T\<ge>t. change_on_neither (block i T) (gap i T)"
by auto
have "(gap i T) \<le> (gap i t)" for T
proof (cases "T \<le> t")
case True
then show ?thesis using gap_monotone by simp
next
case False
then show ?thesis
proof (induction T)
case 0
then show ?case by simp
next
case (Suc T)
with t have "change_on_neither ((block i T)) ((gap i T))"
by simp
then show ?case
using Suc.IH state_change_on_either(1)[of i T] gap_monotone[of T t i]
by metis
qed
qed
then show "\<exists>t. \<forall>T. gap i T \<le> gap i t"
by auto
next
assume "\<exists>t. \<forall>T. gap i T \<le> gap i t"
then obtain t where t: "\<forall>T. gap i T \<le> gap i t"
by auto
have "change_on_neither (block i T) (gap i T)" if "T\<ge>t" for T
proof -
have T: "(gap i T) \<ge> (gap i t)"
using gap_monotone that by simp
show ?thesis
proof (rule ccontr)
assume "\<not> change_on_neither (block i T) (gap i T)"
then have "change_on_0 (block i T) (gap i T) \<or> change_on_1 (block i T) (gap i T)"
by simp
then have "gap i (Suc T) > gap i T"
using gap_le_Suc[of i] state_change_on_either(2)[of i] state_change_on_neither(1)[of i]
dual_order.strict_iff_order
by blast
with T have "gap i (Suc T) > gap i t" by simp
with t show False
using not_le by auto
qed
qed
then show "\<exists>t. \<forall>T\<ge>t. change_on_neither (block i T) (gap i T)"
by auto
qed
text \<open>In Case~1, $\tau_i$ is total.\<close>
lemma case_one_tau_total:
assumes "case_one i"
shows "\<tau> i x \<down>"
proof (cases "x = gap i x")
case True
from assms have "\<forall>t. \<exists>T. gap i T > gap i t"
using le_less_linear gap_def[of i x] by blast
then obtain T where T: "gap i T > gap i x"
by auto
then have "T > x"
using gap_monotone leD le_less_linear by blast
then have "x < T + 2" by simp
moreover from T True have "x \<noteq> gap i T" by simp
ultimately show ?thesis using tau_eq_state' by simp
next
case False
moreover have "x < x + 2" by simp
ultimately show ?thesis using tau_eq_state' by blast
qed
text \<open>In Case~2, $\tau_i$ is undefined only at the gap that never gets filled.\<close>
lemma case_two_tau_not_quite_total:
assumes "\<forall>T. gap i T \<le> gap i t"
shows "\<tau> i (gap i t) \<up>"
and "x \<noteq> gap i t \<Longrightarrow> \<tau> i x \<down>"
proof -
let ?k = "gap i t"
have "\<not> determined i ?k"
proof
assume "determined i ?k"
then obtain T where T: "?k < e_length (block i T) \<and> ?k \<noteq> gap i T"
by auto
with assms have snd_le: "gap i T < ?k"
by (simp add: dual_order.strict_iff_order)
then have "T < t"
using gap_monotone by (metis leD le_less_linear)
from T length_block have "?k < T + 2" by simp
moreover have "?k \<noteq> T + 1"
using T state_change_on_either(2) \<open>T < t\<close> state_unchanged
by (metis Suc_eq_plus1 Suc_leI add_diff_cancel_right' le_add1 nat_neq_iff)
ultimately have "?k \<le> T" by simp
then have "gap i T = gap i ?k"
using state_unchanged[of i T "?k"] \<open>?k < T + 2\<close> snd_le by simp
then show False
by (metis diff_le_self state_unchanged leD nat_le_linear gap_monotone snd_le)
qed
with tau_diverg show "\<tau> i ?k \<up>" by simp
assume "x \<noteq> ?k"
show "\<tau> i x \<down>"
proof (cases "x < t + 2")
case True
- with `x \<noteq> ?k` tau_eq_state' show ?thesis by simp
+ with \<open>x \<noteq> ?k\<close> tau_eq_state' show ?thesis by simp
next
case False
then have "gap i x = ?k"
using assms by (simp add: dual_order.antisym gap_monotone)
- with `x \<noteq> ?k` have "x \<noteq> gap i x" by simp
+ with \<open>x \<noteq> ?k\<close> have "x \<noteq> gap i x" by simp
then show ?thesis using tau_eq_state'[of x x] by simp
qed
qed
lemma case_two_tau_almost_total:
assumes "\<exists>t. \<forall>T. gap i T \<le> gap i t" (is "\<exists>t. ?P t")
shows "\<tau> i (gap i (Least ?P)) \<up>"
and "x \<noteq> gap i (Least ?P) \<Longrightarrow> \<tau> i x \<down>"
proof -
from assms have "?P (Least ?P)"
using LeastI_ex[of ?P] by simp
then show "\<tau> i (gap i (Least ?P)) \<up>" and "x \<noteq> gap i (Least ?P) \<Longrightarrow> \<tau> i x \<down>"
using case_two_tau_not_quite_total by simp_all
qed
text \<open>Some more properties of $\tau$.\<close>
lemma init_tau_gap: "(\<tau> i) \<triangleright> (gap i t - 1) = e_take (gap i t) (block i t)"
proof (intro initI')
show 1: "e_length (e_take (gap i t) (block i t)) = Suc (gap i t - 1)"
proof -
have "gap i t > 0"
using gap_gr0 by simp
moreover have "gap i t < e_length (block i t)"
using gap_in_block by simp
ultimately have "e_length (e_take (gap i t) (block i t)) = gap i t"
by simp
then show ?thesis using gap_gr0 by simp
qed
show "\<tau> i x \<down>= e_nth (e_take (gap i t) (block i t)) x" if "x < Suc (gap i t - 1)" for x
proof -
have x_le: "x < gap i t"
using that gap_gr0 by simp
then have "x < e_length (block i t)"
using gap_in_block less_trans by blast
then have *: "\<tau> i x \<down>= e_nth (block i t) x"
using x_le tau_eq_state by auto
have "x < e_length (e_take (gap i t) (block i t))"
using x_le 1 by simp
then have "e_nth (block i t) x = e_nth (e_take (gap i t) (block i t)) x"
using x_le by simp
then show ?thesis using * by simp
qed
qed
lemma change_on_0_init_tau:
assumes "change_on_0 (block i t) (gap i t)"
shows "(\<tau> i) \<triangleright> (t + 1) = block i t"
proof (intro initI')
let ?b = "block i t" and ?k = "gap i t"
show "e_length (block i t) = Suc (t + 1)"
using length_block by simp
show "(\<tau> i) x \<down>= e_nth (block i t) x" if "x < Suc (t + 1)" for x
proof (cases "x = ?k")
case True
have "gap i (Suc t) = e_length ?b" and b: "block i (Suc t) = e_snoc ?b 0"
using gap_Suc(1) block_Suc(1) assms by simp_all
then have "x < e_length (block i (Suc t))" "x \<noteq> gap i (Suc t)"
using that length_block by simp_all
then have "\<tau> i x \<down>= e_nth (block i (Suc t)) x"
using tau_eq_state by simp
then show ?thesis using that assms b by (simp add: nth_append)
next
case False
then show ?thesis using that assms tau_eq_state' by simp
qed
qed
lemma change_on_0_hyp_change:
assumes "change_on_0 (block i t) (gap i t)"
shows "\<sigma> i ((\<tau> i) \<triangleright> (t + 1)) \<noteq> \<sigma> i ((\<tau> i) \<triangleright> (gap i t - 1))"
using assms hd_block init_tau_gap change_on_0_init_tau by simp
lemma change_on_1_init_tau:
assumes "change_on_1 (block i t) (gap i t)"
shows "(\<tau> i) \<triangleright> (t + 1) = e_update (block i t) (gap i t) 1"
proof (intro initI')
let ?b = "block i t" and ?k = "gap i t"
show "e_length (e_update ?b ?k 1) = Suc (t + 1)"
using length_block by simp
show "(\<tau> i) x \<down>= e_nth (e_update ?b ?k 1) x" if "x < Suc (t + 1)" for x
proof (cases "x = ?k")
case True
have "gap i (Suc t) = e_length ?b" and b: "block i (Suc t) = e_snoc (e_update ?b ?k 1) 0"
using gap_Suc(2) block_Suc(2) assms by simp_all
then have "x < e_length (block i (Suc t))" "x \<noteq> gap i (Suc t)"
using that length_block by simp_all
then have "\<tau> i x \<down>= e_nth (block i (Suc t)) x"
using tau_eq_state by simp
then show ?thesis using that assms b nth_append by (simp add: nth_append)
next
case False
then show ?thesis using that assms tau_eq_state' by simp
qed
qed
lemma change_on_1_hyp_change:
assumes "change_on_1 (block i t) (gap i t)"
shows "\<sigma> i ((\<tau> i) \<triangleright> (t + 1)) \<noteq> \<sigma> i ((\<tau> i) \<triangleright> (gap i t - 1))"
using assms hd_block init_tau_gap change_on_1_init_tau by simp
lemma change_on_either_hyp_change:
assumes "\<not> change_on_neither (block i t) (gap i t)"
shows "\<sigma> i ((\<tau> i) \<triangleright> (t + 1)) \<noteq> \<sigma> i ((\<tau> i) \<triangleright> (gap i t - 1))"
using assms change_on_0_hyp_change change_on_1_hyp_change by auto
lemma filled_gap_0_init_tau:
assumes "f\<^sub>0 = (\<tau> i)((gap i t):=Some 0)"
shows "f\<^sub>0 \<triangleright> (t + 1) = block i t"
proof (intro initI')
show len: "e_length (block i t) = Suc (t + 1)"
using assms length_block by auto
show "f\<^sub>0 x \<down>= e_nth (block i t) x" if "x < Suc (t + 1)" for x
proof (cases "x = gap i t")
case True
then show ?thesis using assms last_block by auto
next
case False
then show ?thesis using assms len tau_eq_state that by auto
qed
qed
lemma filled_gap_1_init_tau:
assumes "f\<^sub>1 = (\<tau> i)((gap i t):=Some 1)"
shows "f\<^sub>1 \<triangleright> (t + 1) = e_update (block i t) (gap i t) 1"
proof (intro initI')
show len: "e_length (e_update (block i t) (gap i t) 1) = Suc (t + 1)"
using e_length_update length_block by simp
show "f\<^sub>1 x \<down>= e_nth (e_update (block i t) (gap i t) 1) x" if "x < Suc (t + 1)" for x
proof (cases "x = gap i t")
case True
moreover have "gap i t < e_length (block i t)"
using gap_in_block by simp
ultimately show ?thesis using assms by simp
next
case False
then show ?thesis using assms len tau_eq_state that by auto
qed
qed
subsection \<open>The separating class\<close>
text \<open>Next we define the sets $V_i$ from the introductory proof sketch
(page~\pageref{s:lim_bc}).\<close>
definition V_bclim :: "nat \<Rightarrow> partial1 set" where
"V_bclim i \<equiv>
if case_two i
then let k = gap i (LEAST t. \<forall>T. gap i T \<le> gap i t)
in {(\<tau> i)(k:=Some 0), (\<tau> i)(k:=Some 1)}
else {\<tau> i}"
lemma V_subseteq_R1: "V_bclim i \<subseteq> \<R>"
proof (cases "case_two i")
case True
define k where "k = gap i (LEAST t. \<forall>T. gap i T \<le> gap i t)"
have "\<tau> i \<in> \<P>"
using tau_in_P2 P2_proj_P1 by auto
then have "(\<tau> i)(k:=Some 0) \<in> \<P>" and "(\<tau> i)(k:=Some 1) \<in> \<P>"
using P1_update_P1 by simp_all
moreover have "total1 ((\<tau> i)(k:=Some v))" for v
using case_two_tau_almost_total(2)[OF True] k_def total1_def by simp
ultimately have "(\<tau> i)(k:=Some 0) \<in> \<R>" and "(\<tau> i)(k:=Some 1) \<in> \<R>"
using P1_total_imp_R1 by simp_all
moreover have "V_bclim i = {(\<tau> i)(k:=Some 0), (\<tau> i)(k:=Some 1)}"
using True V_bclim_def k_def by (simp add: Let_def)
ultimately show ?thesis by simp
next
case False
have "V_bclim i = {\<tau> i}"
unfolding V_bclim_def by (simp add: False)
moreover have "\<tau> i \<in> \<R>"
using total1I case_one_tau_total[OF False] tau_in_P2 P2_proj_P1[of \<tau>] P1_total_imp_R1
by simp
ultimately show ?thesis by simp
qed
lemma case_one_imp_gap_unbounded:
assumes "case_one i"
shows "\<exists>t. gap i t - 1 > n"
proof (induction n)
case 0
then show ?case
using assms gap_gr0[of i] state_at_0(2)[of i] by (metis diff_is_0_eq gr_zeroI)
next
case (Suc n)
then obtain t where t: "gap i t - 1 > n"
by auto
moreover from assms have "\<forall>t. \<exists>T. gap i T > gap i t"
using leI by blast
ultimately obtain T where "gap i T > gap i t"
by auto
then have "gap i T - 1 > gap i t - 1"
using gap_gr0[of i] by (simp add: Suc_le_eq diff_less_mono)
with t have "gap i T - 1 > Suc n" by simp
then show ?case by auto
qed
lemma case_one_imp_not_learn_lim_V:
assumes "case_one i"
shows "\<not> learn_lim \<phi> (V_bclim i) (\<sigma> i)"
proof -
have V_bclim: "V_bclim i = {\<tau> i}"
using assms V_bclim_def by (auto simp add: Let_def)
have "\<exists>m\<^sub>1>n. \<exists>m\<^sub>2>n. (\<sigma> i) ((\<tau> i) \<triangleright> m\<^sub>1) \<noteq> (\<sigma> i) ((\<tau> i) \<triangleright> m\<^sub>2)" for n
proof -
obtain t where t: "gap i t - 1 > n"
using case_one_imp_gap_unbounded[OF assms] by auto
moreover have "\<forall>t. \<exists>T\<ge>t. \<not> change_on_neither (block i T) (gap i T)"
using assms case_two_iff_forever_neither by blast
ultimately obtain T where T: "T \<ge> t" "\<not> change_on_neither (block i T) (gap i T)"
by auto
then have "(\<sigma> i) ((\<tau> i) \<triangleright> (T + 1)) \<noteq> (\<sigma> i) ((\<tau> i) \<triangleright> (gap i T - 1))"
using change_on_either_hyp_change by simp
moreover have "gap i T - 1 > n"
using t T(1) gap_monotone by (simp add: diff_le_mono less_le_trans)
moreover have "T + 1 > n"
proof -
have "gap i T - 1 \<le> T"
using gap_in_block length_block by (simp add: le_diff_conv less_Suc_eq_le)
- then show ?thesis using `gap i T - 1 > n` by simp
+ then show ?thesis using \<open>gap i T - 1 > n\<close> by simp
qed
ultimately show ?thesis by auto
qed
with infinite_hyp_changes_not_Lim V_bclim show ?thesis by simp
qed
lemma case_two_imp_not_learn_lim_V:
assumes "case_two i"
shows "\<not> learn_lim \<phi> (V_bclim i) (\<sigma> i)"
proof -
let ?P = "\<lambda>t. \<forall>T. (gap i T) \<le> (gap i t)"
let ?t = "LEAST t. ?P t"
let ?k = "gap i ?t"
let ?b = "e_take ?k (block i ?t)"
have t: "\<forall>T. gap i T \<le> gap i ?t"
using assms LeastI_ex[of ?P] by simp
then have neither: "\<forall>T\<ge>?t. change_on_neither (block i T) (gap i T)"
using gap_le_Suc gap_monotone state_change_on_neither(1)
by (metis (no_types, lifting) antisym)
have gap_T: "\<forall>T\<ge>?t. gap i T = ?k"
using t gap_monotone antisym_conv by blast
define f\<^sub>0 where "f\<^sub>0 = (\<tau> i)(?k:=Some 0)"
define f\<^sub>1 where "f\<^sub>1 = (\<tau> i)(?k:=Some 1)"
show ?thesis
proof (rule same_hyp_for_two_not_Lim)
show "f\<^sub>0 \<in> V_bclim i" and "f\<^sub>1 \<in> V_bclim i"
using assms V_bclim_def f\<^sub>0_def f\<^sub>1_def by (simp_all add: Let_def)
show "f\<^sub>0 \<noteq> f\<^sub>1" using f\<^sub>0_def f\<^sub>1_def by (meson map_upd_eqD1 zero_neq_one)
show "\<forall>n\<ge>Suc ?t. \<sigma> i (f\<^sub>0 \<triangleright> n) = \<sigma> i ?b"
proof -
have "\<sigma> i (block i T) = \<sigma> i (e_take ?k (block i T))" if "T \<ge> ?t" for T
using that gap_T neither hd_block by metis
then have "\<sigma> i (block i T) = \<sigma> i ?b" if "T \<ge> ?t" for T
by (metis (no_types, lifting) init_tau_gap gap_T that)
then have "\<sigma> i (f\<^sub>0 \<triangleright> (T + 1)) = \<sigma> i ?b" if "T \<ge> ?t" for T
using filled_gap_0_init_tau[of f\<^sub>0 i T] f\<^sub>0_def gap_T that
by (metis (no_types, lifting))
then have "\<sigma> i (f\<^sub>0 \<triangleright> T) = \<sigma> i ?b" if "T \<ge> Suc ?t" for T
using that by (metis (no_types, lifting) Suc_eq_plus1 Suc_le_D Suc_le_mono)
then show ?thesis by simp
qed
show "\<forall>n\<ge>Suc ?t. \<sigma> i (f\<^sub>1 \<triangleright> n) = \<sigma> i ?b"
proof -
have "\<sigma> i (e_update (block i T) ?k 1) = \<sigma> i (e_take ?k (block i T))" if "T \<ge> ?t" for T
using neither by (metis (no_types, lifting) hd_block gap_T that)
then have "\<sigma> i (e_update (block i T) ?k 1) = \<sigma> i ?b" if "T \<ge> ?t" for T
using that init_tau_gap[of i] gap_T by (metis (no_types, lifting))
then have "\<sigma> i (f\<^sub>1 \<triangleright> (T + 1)) = \<sigma> i ?b" if "T \<ge> ?t" for T
using filled_gap_1_init_tau[of f\<^sub>1 i T] f\<^sub>1_def gap_T that
by (metis (no_types, lifting))
then have "\<sigma> i (f\<^sub>1 \<triangleright> T) = \<sigma> i ?b" if "T \<ge> Suc ?t" for T
using that by (metis (no_types, lifting) Suc_eq_plus1 Suc_le_D Suc_le_mono)
then show ?thesis by simp
qed
qed
qed
corollary not_learn_lim_V: "\<not> learn_lim \<phi> (V_bclim i) (\<sigma> i)"
using case_one_imp_not_learn_lim_V case_two_imp_not_learn_lim_V
by (cases "case_two i") simp_all
text \<open>Next we define the separating class.\<close>
definition V_BCLIM :: "partial1 set" ("V\<^bsub>BC-LIM\<^esub>") where
"V\<^bsub>BC-LIM\<^esub> \<equiv> \<Union>i. V_bclim i"
lemma V_BCLIM_R1: "V\<^bsub>BC-LIM\<^esub> \<subseteq> \<R>"
using V_BCLIM_def V_subseteq_R1 by auto
lemma V_BCLIM_not_in_Lim: "V\<^bsub>BC-LIM\<^esub> \<notin> LIM"
proof
assume "V\<^bsub>BC-LIM\<^esub> \<in> LIM"
then obtain s where s: "learn_lim \<phi> V\<^bsub>BC-LIM\<^esub> s"
using learn_lim_wrt_goedel[OF goedel_numbering_phi] Lim_def by blast
moreover obtain i where "\<phi> i = s"
using s learn_limE(1) phi_universal by blast
ultimately have "learn_lim \<phi> V\<^bsub>BC-LIM\<^esub> (\<lambda>x. eval r_sigma [i, x])"
using learn_lim_sigma by simp
moreover have "V_bclim i \<subseteq> V\<^bsub>BC-LIM\<^esub>"
using V_BCLIM_def by auto
ultimately have "learn_lim \<phi> (V_bclim i) (\<lambda>x. eval r_sigma [i, x])"
using learn_lim_closed_subseteq by simp
then show False
using not_learn_lim_V by simp
qed
subsection \<open>The separating class is in BC\<close>
text \<open>In order to show @{term "V\<^bsub>BC-LIM\<^esub> \<in> BC"} we
define a hypothesis space that for every function $\tau_i$ and every list $b$
of numbers contains a copy of $\tau_i$ with the first $|b|$ values replaced
by $b$.\<close>
definition psitau :: partial2 ("\<psi>\<^sup>\<tau>") where
"\<psi>\<^sup>\<tau> b x \<equiv> (if x < e_length b then Some (e_nth b x) else \<tau> (e_hd b) x)"
lemma psitau_in_P2: "\<psi>\<^sup>\<tau> \<in> \<P>\<^sup>2"
proof -
define r where "r \<equiv>
Cn 2
(r_lifz r_nth (Cn 2 r_tau [Cn 2 r_hd [Id 2 0], Id 2 1]))
[Cn 2 r_less [Id 2 1, Cn 2 r_length [Id 2 0]], Id 2 0, Id 2 1]"
then have "recfn 2 r"
using r_tau_recfn by simp
moreover have "eval r [b, x] = \<psi>\<^sup>\<tau> b x" for b x
proof -
let ?f = "Cn 2 r_tau [Cn 2 r_hd [Id 2 0], Id 2 1]"
have "recfn 2 r_nth" "recfn 2 ?f"
using r_tau_recfn by simp_all
then have "eval (r_lifz r_nth ?f) [c, b, x] =
(if c = 0 then eval r_nth [b, x] else eval ?f [b, x])" for c
by simp
moreover have "eval r_nth [b, x] \<down>= e_nth b x"
by simp
moreover have "eval ?f [b, x] = \<tau> (e_hd b) x"
using r_tau_recfn by simp
ultimately have "eval (r_lifz r_nth ?f) [c, b, x] =
(if c = 0 then Some (e_nth b x) else \<tau> (e_hd b) x)" for c
by simp
moreover have "eval (Cn 2 r_less [Id 2 1, Cn 2 r_length [Id 2 0]]) [b, x] \<down>=
(if x < e_length b then 0 else 1)"
by simp
ultimately show ?thesis
unfolding r_def psitau_def using r_tau_recfn by simp
qed
ultimately show ?thesis by auto
qed
lemma psitau_init:
"\<psi>\<^sup>\<tau> (f \<triangleright> n) x = (if x < Suc n then Some (the (f x)) else \<tau> (the (f 0)) x)"
proof -
let ?e = "f \<triangleright> n"
have "e_length ?e = Suc n" by simp
moreover have "x < Suc n \<Longrightarrow> e_nth ?e x = the (f x)" by simp
moreover have "e_hd ?e = the (f 0)"
using hd_init by simp
ultimately show ?thesis using psitau_def by simp
qed
text \<open>The class @{term V_BCLIM} can be learned BC-style in the
hypothesis space @{term psitau} by the identity function.\<close>
lemma learn_bc_V_BCLIM: "learn_bc \<psi>\<^sup>\<tau> V\<^bsub>BC-LIM\<^esub> Some"
proof (rule learn_bcI)
show "environment \<psi>\<^sup>\<tau> V\<^bsub>BC-LIM\<^esub> Some"
using identity_in_R1 V_BCLIM_R1 psitau_in_P2 by auto
show "\<exists>n\<^sub>0. \<forall>n\<ge>n\<^sub>0. \<psi>\<^sup>\<tau> (the (Some (f \<triangleright> n))) = f" if "f \<in> V\<^bsub>BC-LIM\<^esub>" for f
proof -
from that V_BCLIM_def obtain i where i: "f \<in> V_bclim i"
by auto
show ?thesis
proof (cases "case_two i")
case True
let ?P = "\<lambda>t. \<forall>T. (gap i T) \<le> (gap i t)"
let ?lmin = "LEAST t. ?P t"
define k where "k \<equiv> gap i ?lmin"
have V_bclim: "V_bclim i = {(\<tau> i)(k:=Some 0), (\<tau> i)(k:=Some 1)}"
using True V_bclim_def k_def by (simp add: Let_def)
moreover have "0 < k"
using gap_gr0[of i] k_def by simp
ultimately have "f 0 \<down>= i"
using tau_at_0[of i] i by auto
have "\<psi>\<^sup>\<tau> (f \<triangleright> n) = f" if "n \<ge> k" for n
proof
fix x
show "\<psi>\<^sup>\<tau> (f \<triangleright> n) x = f x"
proof (cases "x \<le> n")
case True
then show ?thesis
using R1_imp_total1 V_subseteq_R1 i psitau_init by fastforce
next
case False
then have "\<psi>\<^sup>\<tau> (f \<triangleright> n) x = \<tau> (the (f 0)) x"
using psitau_init by simp
then have "\<psi>\<^sup>\<tau> (f \<triangleright> n) x = \<tau> i x"
- using `f 0 \<down>= i` by simp
+ using \<open>f 0 \<down>= i\<close> by simp
moreover have "f x = \<tau> i x"
using False V_bclim i that by auto
ultimately show ?thesis by simp
qed
qed
then show ?thesis by auto
next
case False
then have "V_bclim i = {\<tau> i}"
using V_bclim_def by (auto simp add: Let_def)
then have f: "f = \<tau> i"
using i by simp
have "\<psi>\<^sup>\<tau> (f \<triangleright> n) = f" for n
proof
fix x
show "\<psi>\<^sup>\<tau> (f \<triangleright> n) x = f x"
proof (cases "x \<le> n")
case True
then show ?thesis
using R1_imp_total1 V_BCLIM_R1 psitau_init that by auto
next
case False
then show ?thesis by (simp add: f psitau_init tau_at_0)
qed
qed
then show ?thesis by simp
qed
qed
qed
text \<open>Finally, the main result of this section:\<close>
theorem Lim_subset_BC: "LIM \<subset> BC"
using learn_bc_V_BCLIM BC_def Lim_subseteq_BC V_BCLIM_not_in_Lim by auto
end
diff --git a/thys/Inductive_Inference/Lemma_R.thy b/thys/Inductive_Inference/Lemma_R.thy
--- a/thys/Inductive_Inference/Lemma_R.thy
+++ b/thys/Inductive_Inference/Lemma_R.thy
@@ -1,2114 +1,2114 @@
section \<open>Lemma R\label{s:lemma_r}\<close>
theory Lemma_R
imports Inductive_Inference_Basics
begin
text \<open>A common technique for constructing a class that cannot be
learned is diagonalization against all strategies (see, for instance,
Section~\ref{s:lim_bc}). Similarly, the typical way of proving that a class
cannot be learned is by assuming there is a strategy and deriving a
contradiction. Both techniques are easier to carry out if one has to consider
only \emph{total} recursive strategies. This is not possible in general,
since after all the definitions of the inference types admit strictly partial
strategies. However, for many inference types one can show that for every
strategy there is a total strategy with at least the same ``learning power''.
Results to that effect are called Lemma~R.
Lemma~R comes in different strengths depending on how general the
construction of the total recursive strategy is. CONS is the only inference
type considered here for which not even a weak form of Lemma~R holds.\<close>
subsection \<open>Strong Lemma R for LIM, FIN, and BC\<close>
text \<open>In its strong form Lemma~R says that for any strategy $S$, there
is a total strategy $T$ that learns all classes $S$ learns regardless of
hypothesis space. The strategy $T$ can be derived from $S$ by a delayed
simulation of $S$. More precisely, for input $f^n$, $T$ simulates $S$ for
prefixes $f^0, f^1, \ldots, f^n$ for at most $n$ steps. If $S$ halts on none
of the prefixes, $T$ outputs an arbitrary hypothesis. Otherwise let $k \leq
n$ be maximal such that $S$ halts on $f^k$ in at most $n$ steps. Then $T$
outputs $S(f^k)$. \<close>
text \<open>We reformulate some lemmas for @{term r_result1} to make it easier
to use them with @{term "\<phi>"}.\<close>
lemma r_result1_converg_phi:
assumes "\<phi> i x \<down>= v"
shows "\<exists>t.
(\<forall>t'\<ge>t. eval r_result1 [t', i, x] \<down>= Suc v) \<and>
(\<forall>t'<t. eval r_result1 [t', i, x] \<down>= 0)"
using assms r_result1_converg' phi_def by simp_all
lemma r_result1_bivalent':
assumes "eval r_phi [i, x] \<down>= v"
shows "eval r_result1 [t, i, x] \<down>= Suc v \<or> eval r_result1 [t, i, x] \<down>= 0"
using assms r_result1 r_result_bivalent' r_phi'' by simp
lemma r_result1_bivalent_phi:
assumes "\<phi> i x \<down>= v"
shows "eval r_result1 [t, i, x] \<down>= Suc v \<or> eval r_result1 [t, i, x] \<down>= 0"
using assms r_result1_bivalent' phi_def by simp_all
lemma r_result1_diverg_phi:
assumes "\<phi> i x \<up>"
shows "eval r_result1 [t, i, x] \<down>= 0"
using assms phi_def r_result1_diverg' by simp
lemma r_result1_some_phi:
assumes "eval r_result1 [t, i, x] \<down>= Suc v"
shows "\<phi> i x \<down>= v"
using assms phi_def r_result1_Some' by simp
lemma r_result1_saturating':
assumes "eval r_result1 [t, i, x] \<down>= Suc v"
shows "eval r_result1 [t + d, i, x] \<down>= Suc v"
using assms r_result1 r_result_saturating r_phi'' by simp
lemma r_result1_saturating_the:
assumes "the (eval r_result1 [t, i, x]) > 0" and "t' \<ge> t"
shows "the (eval r_result1 [t', i, x]) > 0"
proof -
from assms(1) obtain v where "eval r_result1 [t, i, x] \<down>= Suc v"
using r_result1_bivalent_phi r_result1_diverg_phi
by (metis inc_induct le_0_eq not_less_zero option.discI option.expand option.sel)
with assms have "eval r_result1 [t', i, x] \<down>= Suc v"
using r_result1_saturating' le_Suc_ex by blast
then show ?thesis by simp
qed
lemma Greatest_bounded_Suc:
fixes P :: "nat \<Rightarrow> nat"
shows "(if P n > 0 then Suc n
else if \<exists>j<n. P j > 0 then Suc (GREATEST j. j < n \<and> P j > 0) else 0) =
(if \<exists>j<Suc n. P j > 0 then Suc (GREATEST j. j < Suc n \<and> P j > 0) else 0)"
(is "?lhs = ?rhs")
proof (cases "\<exists>j<Suc n. P j > 0")
case 1: True
show ?thesis
proof (cases "P n > 0")
case True
then have "(GREATEST j. j < Suc n \<and> P j > 0) = n"
using Greatest_equality[of "\<lambda>j. j < Suc n \<and> P j > 0"] by simp
moreover have "?rhs = Suc (GREATEST j. j < Suc n \<and> P j > 0)"
using 1 by simp
ultimately have "?rhs = Suc n" by simp
then show ?thesis using True by simp
next
case False
then have "?lhs = Suc (GREATEST j. j < n \<and> P j > 0)"
using 1 by (metis less_SucE)
moreover have "?rhs = Suc (GREATEST j. j < Suc n \<and> P j > 0)"
using 1 by simp
moreover have "(GREATEST j. j < n \<and> P j > 0) =
(GREATEST j. j < Suc n \<and> P j > 0)"
using 1 False by (metis less_SucI less_Suc_eq)
ultimately show ?thesis by simp
qed
next
case False
then show ?thesis by auto
qed
text \<open>For $n$, $i$, $x$, the next function simulates $\varphi_i$ on all
non-empty prefixes of at most length $n$ of the list $x$ for at most $n$
steps. It returns the length of the longest such prefix for which $\varphi_i$
halts, or zero if $\varphi_i$ does not halt for any prefix.\<close>
definition "r_delay_aux \<equiv>
Pr 2 (r_constn 1 0)
(Cn 4 r_ifz
[Cn 4 r_result1
[Cn 4 r_length [Id 4 3], Id 4 2,
Cn 4 r_take [Cn 4 S [Id 4 0], Id 4 3]],
Id 4 1, Cn 4 S [Id 4 0]])"
lemma r_delay_aux_prim: "prim_recfn 3 r_delay_aux"
unfolding r_delay_aux_def by simp_all
lemma r_delay_aux_total: "total r_delay_aux"
using prim_recfn_total[OF r_delay_aux_prim] .
lemma r_delay_aux:
assumes "n \<le> e_length x"
shows "eval r_delay_aux [n, i, x] \<down>=
(if \<exists>j<n. the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0
then Suc (GREATEST j.
j < n \<and>
the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0)
else 0)"
proof -
define z where "z \<equiv>
Cn 4 r_result1
[Cn 4 r_length [Id 4 3], Id 4 2, Cn 4 r_take [Cn 4 S [Id 4 0], Id 4 3]]"
then have z_recfn: "recfn 4 z" by simp
have z: "eval z [j, r, i, x] = eval r_result1 [e_length x, i, e_take (Suc j) x]"
if "j < e_length x" for j r i x
unfolding z_def using that by simp
define g where "g \<equiv> Cn 4 r_ifz [z, Id 4 1, Cn 4 S [Id 4 0]]"
then have g: "eval g [j, r, i, x] \<down>=
(if the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0 then Suc j else r)"
if "j < e_length x" for j r i x
using that z prim_recfn_total z_recfn by simp
show ?thesis
using assms
proof (induction n)
case 0
moreover have "eval r_delay_aux [0, i, x] \<down>= 0"
using eval_Pr_0 r_delay_aux_def r_delay_aux_prim r_constn
by (simp add: r_delay_aux_def)
ultimately show ?case by simp
next
case (Suc n)
let ?P = "\<lambda>j. the (eval r_result1 [e_length x, i, e_take (Suc j) x])"
have "eval r_delay_aux [n, i, x] \<down>"
using Suc by simp
moreover have "eval r_delay_aux [Suc n, i, x] =
eval (Pr 2 (r_constn 1 0) g) [Suc n, i, x]"
unfolding r_delay_aux_def g_def z_def by simp
ultimately have "eval r_delay_aux [Suc n, i, x] =
eval g [n, the (eval r_delay_aux [n, i, x]), i, x]"
using r_delay_aux_prim Suc eval_Pr_converg_Suc
by (simp add: r_delay_aux_def g_def z_def numeral_3_eq_3)
then have "eval r_delay_aux [Suc n, i, x] \<down>=
(if ?P n > 0 then Suc n
else if \<exists>j<n. ?P j > 0 then Suc (GREATEST j. j < n \<and> ?P j > 0) else 0)"
using g Suc by simp
then have "eval r_delay_aux [Suc n, i, x] \<down>=
(if \<exists>j<Suc n. ?P j > 0 then Suc (GREATEST j. j < Suc n \<and> ?P j > 0) else 0)"
using Greatest_bounded_Suc[where ?P="?P"] by simp
then show ?case by simp
qed
qed
text \<open>The next function simulates $\varphi_i$ on all non-empty prefixes
of a list $x$ of length $n$ for at most $n$ steps and outputs the length of
the longest prefix for which $\varphi_i$ halts, or zero if $\varphi_i$ does
not halt for any such prefix.\<close>
definition "r_delay \<equiv> Cn 2 r_delay_aux [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]"
lemma r_delay_recfn [simp]: "recfn 2 r_delay"
unfolding r_delay_def by (simp add: r_delay_aux_prim)
lemma r_delay:
"eval r_delay [i, x] \<down>=
(if \<exists>j<e_length x. the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0
then Suc (GREATEST j.
j < e_length x \<and> the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0)
else 0)"
unfolding r_delay_def using r_delay_aux r_delay_aux_prim by simp
definition "delay i x \<equiv> Some
(if \<exists>j<e_length x. the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0
then Suc (GREATEST j.
j < e_length x \<and> the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0)
else 0)"
lemma delay_in_R2: "delay \<in> \<R>\<^sup>2"
using r_delay totalI2 R2I delay_def r_delay_recfn
by (metis (no_types, lifting) numeral_2_eq_2 option.simps(3))
lemma delay_le_length: "the (delay i x) \<le> e_length x"
proof (cases "\<exists>j<e_length x. the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0")
case True
let ?P = "\<lambda>j. j < e_length x \<and> the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0"
from True have "\<exists>j. ?P j" by simp
moreover have "\<And>y. ?P y \<Longrightarrow> y \<le> e_length x" by simp
ultimately have "?P (Greatest ?P)"
using GreatestI_ex_nat[where ?P="?P"] by blast
then have "Greatest ?P < e_length x" by simp
moreover have "delay i x \<down>= Suc (Greatest ?P)"
using delay_def True by simp
ultimately show ?thesis by auto
next
case False
then show ?thesis using delay_def by auto
qed
lemma e_take_delay_init:
assumes "f \<in> \<R>" and "the (delay i (f \<triangleright> n)) > 0"
shows "e_take (the (delay i (f \<triangleright> n))) (f \<triangleright> n) = f \<triangleright> (the (delay i (f \<triangleright> n)) - 1)"
using assms e_take_init[of f _ n] length_init[of f n] delay_le_length[of i "f \<triangleright> n"]
by (metis One_nat_def Suc_le_lessD Suc_pred)
lemma delay_gr0_converg:
assumes "the (delay i x) > 0"
shows "\<phi> i (e_take (the (delay i x)) x) \<down>"
proof -
let ?P = "\<lambda>j. j < e_length x \<and> the (eval r_result1 [e_length x, i, e_take (Suc j) x]) > 0"
have "\<exists>j. ?P j"
proof (rule ccontr)
assume "\<not> (\<exists>j. ?P j)"
then have "delay i x \<down>= 0"
using delay_def by simp
with assms show False by simp
qed
then have d: "the (delay i x) = Suc (Greatest ?P)"
using delay_def by simp
moreover have "\<And>y. ?P y \<Longrightarrow> y \<le> e_length x" by simp
ultimately have "?P (Greatest ?P)"
- using `\<exists>j. ?P j` GreatestI_ex_nat[where ?P="?P"] by blast
+ using \<open>\<exists>j. ?P j\<close> GreatestI_ex_nat[where ?P="?P"] by blast
then have "the (eval r_result1 [e_length x, i, e_take (Suc (Greatest ?P)) x]) > 0"
by simp
then have "the (eval r_result1 [e_length x, i, e_take (the (delay i x)) x]) > 0"
using d by simp
then show ?thesis using r_result1_diverg_phi by fastforce
qed
lemma delay_unbounded:
fixes n :: nat
assumes "f \<in> \<R>" and "\<forall>n. \<phi> i (f \<triangleright> n) \<down>"
shows "\<exists>m. the (delay i (f \<triangleright> m)) > n"
proof -
from assms have "\<exists>t. the (eval r_result1 [t, i, f \<triangleright> n]) > 0"
using r_result1_converg_phi
by (metis le_refl option.exhaust_sel option.sel zero_less_Suc)
then obtain t where t: "the (eval r_result1 [t, i, f \<triangleright> n]) > 0"
by auto
let ?m = "max n t"
have "Suc ?m \<ge> t" by simp
have m: "the (eval r_result1 [Suc ?m, i, f \<triangleright> n]) > 0"
proof -
let ?w = "eval r_result1 [t, i, f \<triangleright> n]"
obtain v where v: "?w \<down>= Suc v"
using t assms(2) r_result1_bivalent_phi by fastforce
have "eval r_result1 [Suc ?m, i, f \<triangleright> n] = ?w"
- using v t r_result1_saturating' `Suc ?m \<ge> t` le_Suc_ex by fastforce
+ using v t r_result1_saturating' \<open>Suc ?m \<ge> t\<close> le_Suc_ex by fastforce
then show ?thesis using t by simp
qed
let ?x = "f \<triangleright> ?m"
have "the (delay i ?x) > n"
proof -
let ?P = "\<lambda>j. j < e_length ?x \<and> the (eval r_result1 [e_length ?x, i, e_take (Suc j) ?x]) > 0"
have "e_length ?x = Suc ?m" by simp
moreover have "e_take (Suc n) ?x = f \<triangleright> n"
using assms(1) e_take_init by auto
ultimately have "?P n"
using m by simp
have "\<And>y. ?P y \<Longrightarrow> y \<le> e_length ?x" by simp
- with `?P n` have "n \<le> (Greatest ?P)"
+ with \<open>?P n\<close> have "n \<le> (Greatest ?P)"
using Greatest_le_nat[of ?P n "e_length ?x"] by simp
moreover have "the (delay i ?x) = Suc (Greatest ?P)"
- using delay_def `?P n` by auto
+ using delay_def \<open>?P n\<close> by auto
ultimately show ?thesis by simp
qed
then show ?thesis by auto
qed
lemma delay_monotone:
assumes "f \<in> \<R>" and "n\<^sub>1 \<le> n\<^sub>2"
shows "the (delay i (f \<triangleright> n\<^sub>1)) \<le> the (delay i (f \<triangleright> n\<^sub>2))"
(is "the (delay i ?x1) \<le> the (delay i ?x2)")
proof (cases "the (delay i (f \<triangleright> n\<^sub>1)) = 0")
case True
then show ?thesis by simp
next
case False
let ?P1 = "\<lambda>j. j < e_length ?x1 \<and> the (eval r_result1 [e_length ?x1, i, e_take (Suc j) ?x1]) > 0"
let ?P2 = "\<lambda>j. j < e_length ?x2 \<and> the (eval r_result1 [e_length ?x2, i, e_take (Suc j) ?x2]) > 0"
from False have d1: "the (delay i ?x1) = Suc (Greatest ?P1)" "\<exists>j. ?P1 j"
using delay_def option.collapse by fastforce+
moreover have "\<And>y. ?P1 y \<Longrightarrow> y \<le> e_length ?x1" by simp
ultimately have *: "?P1 (Greatest ?P1)" using GreatestI_ex_nat[of ?P1] by blast
let ?j = "Greatest ?P1"
from * have "?j < e_length ?x1" by auto
then have 1: "e_take (Suc ?j) ?x1 = e_take (Suc ?j) ?x2"
using assms e_take_init by auto
from * have 2: "?j < e_length ?x2" using assms(2) by auto
with 1 * have "the (eval r_result1 [e_length ?x1, i, e_take (Suc ?j) ?x2]) > 0"
by simp
moreover have "e_length ?x1 \<le> e_length ?x2"
using assms(2) by auto
ultimately have "the (eval r_result1 [e_length ?x2, i, e_take (Suc ?j) ?x2]) > 0"
using r_result1_saturating_the by simp
with 2 have "?P2 ?j" by simp
then have d2: "the (delay i ?x2) = Suc (Greatest ?P2)"
using delay_def by auto
have "\<And>y. ?P2 y \<Longrightarrow> y \<le> e_length ?x2" by simp
- with `?P2 ?j` have "?j \<le> (Greatest ?P2)" using Greatest_le_nat[of ?P2] by blast
+ with \<open>?P2 ?j\<close> have "?j \<le> (Greatest ?P2)" using Greatest_le_nat[of ?P2] by blast
with d1 d2 show ?thesis by simp
qed
lemma delay_unbounded_monotone:
fixes n :: nat
assumes "f \<in> \<R>" and "\<forall>n. \<phi> i (f \<triangleright> n) \<down>"
shows "\<exists>m\<^sub>0. \<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n"
proof -
from assms delay_unbounded obtain m\<^sub>0 where "the (delay i (f \<triangleright> m\<^sub>0)) > n"
by blast
then have "\<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n"
using assms(1) delay_monotone order.strict_trans2 by blast
then show ?thesis by auto
qed
text \<open>Now we can define a function that simulates an arbitrary strategy
$\varphi_i$ in a delayed way. The parameter $d$ is the default hypothesis for
when $\varphi_i$ does not halt within the time bound for any prefix.\<close>
definition r_totalizer :: "nat \<Rightarrow> recf" where
"r_totalizer d \<equiv>
Cn 2
(r_lifz
(r_constn 1 d)
(Cn 2 r_phi
[Id 2 0, Cn 2 r_take [Cn 2 r_delay [Id 2 0, Id 2 1], Id 2 1]]))
[Cn 2 r_delay [Id 2 0, Id 2 1], Id 2 0, Id 2 1]"
lemma r_totalizer_recfn: "recfn 2 (r_totalizer d)"
unfolding r_totalizer_def by simp
lemma r_totalizer:
"eval (r_totalizer d) [i, x] =
(if the (delay i x) = 0 then Some d else \<phi> i (e_take (the (delay i x)) x))"
proof -
let ?i = "Cn 2 r_delay [Id 2 0, Id 2 1]"
have "eval ?i [i, x] = eval r_delay [i, x]" for i x
using r_delay_recfn by simp
then have i: "eval ?i [i, x] = delay i x" for i x
using r_delay by (simp add: delay_def)
let ?t = "r_constn 1 d"
have t: "eval ?t [i, x] \<down>= d" for i x by simp
let ?e1 = "Cn 2 r_take [?i, Id 2 1]"
let ?e = "Cn 2 r_phi [Id 2 0, ?e1]"
have "eval ?e1 [i, x] = eval r_take [the (delay i x), x]" for i x
using r_delay i delay_def by simp
then have "eval ?e1 [i, x] \<down>= e_take (the (delay i x)) x" for i x
using delay_le_length by simp
then have e: "eval ?e [i, x] = \<phi> i (e_take (the (delay i x)) x)"
using phi_def by simp
let ?z = "r_lifz ?t ?e"
have recfn_te: "recfn 2 ?t" "recfn 2 ?e"
by simp_all
then have "eval (r_totalizer d) [i, x] = eval (r_lifz ?t ?e) [the (delay i x), i, x]"
for i x
unfolding r_totalizer_def using i r_totalizer_recfn delay_def by simp
then have "eval (r_totalizer d) [i, x] =
(if the (delay i x) = 0 then eval ?t [i, x] else eval ?e [i, x])"
for i x
using recfn_te by simp
then show ?thesis using t e by simp
qed
lemma r_totalizer_total: "total (r_totalizer d)"
proof (rule totalI2)
show "recfn 2 (r_totalizer d)" using r_totalizer_recfn by simp
show "\<And>x y. eval (r_totalizer d) [x, y] \<down>"
using r_totalizer delay_gr0_converg by simp
qed
definition totalizer :: "nat \<Rightarrow> partial2" where
"totalizer d i x \<equiv>
if the (delay i x) = 0 then Some d else \<phi> i (e_take (the (delay i x)) x)"
lemma totalizer_init:
assumes "f \<in> \<R>"
shows "totalizer d i (f \<triangleright> n) =
(if the (delay i (f \<triangleright> n)) = 0 then Some d
else \<phi> i (f \<triangleright> (the (delay i (f \<triangleright> n)) - 1)))"
using assms e_take_delay_init by (simp add: totalizer_def)
lemma totalizer_in_R2: "totalizer d \<in> \<R>\<^sup>2"
using totalizer_def r_totalizer r_totalizer_total R2I r_totalizer_recfn
by metis
text \<open>For LIM, @{term totalizer} works with every default hypothesis
$d$.\<close>
lemma lemma_R_for_Lim:
assumes "learn_lim \<psi> U (\<phi> i)"
shows "learn_lim \<psi> U (totalizer d i)"
proof (rule learn_limI)
show env: "environment \<psi> U (totalizer d i)"
using assms learn_limE(1) totalizer_in_R2 by auto
show "\<exists>j. \<psi> j = f \<and> (\<forall>\<^sup>\<infinity>n. totalizer d i (f \<triangleright> n) \<down>= j)" if "f \<in> U" for f
proof -
have "f \<in> \<R>"
using assms env that by auto
from assms learn_limE obtain j n\<^sub>0 where
j: "\<psi> j = f" and
n0: "\<forall>n\<ge>n\<^sub>0. (\<phi> i) (f \<triangleright> n) \<down>= j"
- using `f \<in> U` by metis
+ using \<open>f \<in> U\<close> by metis
obtain m\<^sub>0 where m0: "\<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n\<^sub>0"
- using delay_unbounded_monotone `f \<in> \<R>` \<open>f \<in> U\<close> assms learn_limE(1)
+ using delay_unbounded_monotone \<open>f \<in> \<R>\<close> \<open>f \<in> U\<close> assms learn_limE(1)
by blast
then have "\<forall>m\<ge>m\<^sub>0. totalizer d i (f \<triangleright> m) = \<phi> i (e_take (the (delay i (f \<triangleright> m))) (f \<triangleright> m))"
using totalizer_def by auto
then have "\<forall>m\<ge>m\<^sub>0. totalizer d i (f \<triangleright> m) = \<phi> i (f \<triangleright> (the (delay i (f \<triangleright> m)) - 1))"
- using e_take_delay_init m0 `f \<in> \<R>` by auto
+ using e_take_delay_init m0 \<open>f \<in> \<R>\<close> by auto
with m0 n0 have "\<forall>m\<ge>m\<^sub>0. totalizer d i (f \<triangleright> m) \<down>= j"
by auto
with j show ?thesis by auto
qed
qed
text \<open>The effective version of Lemma~R for LIM states that there is a
total recursive function computing Gödel numbers of total strategies
from those of arbitrary strategies.\<close>
lemma lemma_R_for_Lim_effective:
"\<exists>g\<in>\<R>. \<forall>i.
\<phi> (the (g i)) \<in> \<R> \<and>
(\<forall>U \<psi>. learn_lim \<psi> U (\<phi> i) \<longrightarrow> learn_lim \<psi> U (\<phi> (the (g i))))"
proof -
have "totalizer 0 \<in> \<P>\<^sup>2" using totalizer_in_R2 by auto
then obtain g where g: "g \<in> \<R>" "\<forall>i. (totalizer 0) i = \<phi> (the (g i))"
using numbering_translation_for_phi by blast
with totalizer_in_R2 have "\<forall>i. \<phi> (the (g i)) \<in> \<R>"
by (metis R2_proj_R1)
moreover from g(2) lemma_R_for_Lim[where ?d=0] have
"\<forall>i U \<psi>. learn_lim \<psi> U (\<phi> i) \<longrightarrow> learn_lim \<psi> U (\<phi> (the (g i)))"
by simp
ultimately show ?thesis using g(1) by blast
qed
text \<open>In order for us to use the previous lemma, we need a function
that performs the actual computation:\<close>
definition "r_limr \<equiv>
SOME g.
recfn 1 g \<and>
total g \<and>
(\<forall>i. \<phi> (the (eval g [i])) \<in> \<R> \<and>
(\<forall>U \<psi>. learn_lim \<psi> U (\<phi> i) \<longrightarrow> learn_lim \<psi> U (\<phi> (the (eval g [i])))))"
lemma r_limr_recfn: "recfn 1 r_limr"
and r_limr_total: "total r_limr"
and r_limr:
"\<phi> (the (eval r_limr [i])) \<in> \<R>"
"learn_lim \<psi> U (\<phi> i) \<Longrightarrow> learn_lim \<psi> U (\<phi> (the (eval r_limr [i])))"
proof -
let ?P = "\<lambda>g.
g \<in> \<R> \<and>
(\<forall>i. \<phi> (the (g i)) \<in> \<R> \<and> (\<forall>U \<psi>. learn_lim \<psi> U (\<phi> i) \<longrightarrow> learn_lim \<psi> U (\<phi> (the (g i)))))"
let ?Q = "\<lambda>g.
recfn 1 g \<and>
total g \<and>
(\<forall>i. \<phi> (the (eval g [i])) \<in> \<R> \<and>
(\<forall>U \<psi>. learn_lim \<psi> U (\<phi> i) \<longrightarrow> learn_lim \<psi> U (\<phi> (the (eval g [i])))))"
have "\<exists>g. ?P g" using lemma_R_for_Lim_effective by auto
then obtain g where "?P g" by auto
then obtain g' where g': "recfn 1 g'" "total g'" "\<forall>i. eval g' [i] = g i"
by blast
- with `?P g` have "?Q g'" by simp
+ with \<open>?P g\<close> have "?Q g'" by simp
with r_limr_def someI_ex[of ?Q] show
"recfn 1 r_limr"
"total r_limr"
"\<phi> (the (eval r_limr [i])) \<in> \<R>"
"learn_lim \<psi> U (\<phi> i) \<Longrightarrow> learn_lim \<psi> U (\<phi> (the (eval r_limr [i])))"
by auto
qed
text \<open>For BC, too, @{term totalizer} works with every default
hypothesis $d$.\<close>
lemma lemma_R_for_BC:
assumes "learn_bc \<psi> U (\<phi> i)"
shows "learn_bc \<psi> U (totalizer d i)"
proof (rule learn_bcI)
show env: "environment \<psi> U (totalizer d i)"
using assms learn_bcE(1) totalizer_in_R2 by auto
show "\<exists>n\<^sub>0. \<forall>n\<ge>n\<^sub>0. \<psi> (the (totalizer d i (f \<triangleright> n))) = f" if "f \<in> U" for f
proof -
have "f \<in> \<R>"
using assms env that by auto
obtain n\<^sub>0 where n0: "\<forall>n\<ge>n\<^sub>0. \<psi> (the ((\<phi> i) (f \<triangleright> n))) = f"
- using assms learn_bcE `f \<in> U` by metis
+ using assms learn_bcE \<open>f \<in> U\<close> by metis
obtain m\<^sub>0 where m0: "\<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n\<^sub>0"
- using delay_unbounded_monotone `f \<in> \<R>` \<open>f \<in> U\<close> assms learn_bcE(1)
+ using delay_unbounded_monotone \<open>f \<in> \<R>\<close> \<open>f \<in> U\<close> assms learn_bcE(1)
by blast
then have "\<forall>m\<ge>m\<^sub>0. totalizer d i (f \<triangleright> m) = \<phi> i (e_take (the (delay i (f \<triangleright> m))) (f \<triangleright> m))"
using totalizer_def by auto
then have "\<forall>m\<ge>m\<^sub>0. totalizer d i (f \<triangleright> m) = \<phi> i (f \<triangleright> (the (delay i (f \<triangleright> m)) - 1))"
- using e_take_delay_init m0 `f \<in> \<R>` by auto
+ using e_take_delay_init m0 \<open>f \<in> \<R>\<close> by auto
with m0 n0 have "\<forall>m\<ge>m\<^sub>0. \<psi> (the (totalizer d i (f \<triangleright> m))) = f"
by auto
then show ?thesis by auto
qed
qed
corollary lemma_R_for_BC_simple:
assumes "learn_bc \<psi> U s"
shows "\<exists>s'\<in>\<R>. learn_bc \<psi> U s'"
using assms lemma_R_for_BC totalizer_in_R2 learn_bcE
by (metis R2_proj_R1 learn_bcE(1) phi_universal)
text \<open>For FIN the default hypothesis of @{term totalizer} must be
zero, signalling ``don't know yet''.\<close>
lemma lemma_R_for_FIN:
assumes "learn_fin \<psi> U (\<phi> i)"
shows "learn_fin \<psi> U (totalizer 0 i)"
proof (rule learn_finI)
show env: "environment \<psi> U (totalizer 0 i)"
using assms learn_finE(1) totalizer_in_R2 by auto
show "\<exists>j n\<^sub>0. \<psi> j = f \<and>
(\<forall>n<n\<^sub>0. totalizer 0 i (f \<triangleright> n) \<down>= 0) \<and>
(\<forall>n\<ge>n\<^sub>0. totalizer 0 i (f \<triangleright> n) \<down>= Suc j)"
if "f \<in> U" for f
proof -
have "f \<in> \<R>"
using assms env that by auto
from assms learn_finE[of \<psi> U "\<phi> i"] obtain j where
j: "\<psi> j = f" and
ex_n0: "\<exists>n\<^sub>0. (\<forall>n<n\<^sub>0. (\<phi> i) (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. (\<phi> i) (f \<triangleright> n) \<down>= Suc j)"
- using `f \<in> U` by blast
+ using \<open>f \<in> U\<close> by blast
let ?Q = "\<lambda>n\<^sub>0. (\<forall>n<n\<^sub>0. (\<phi> i) (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. (\<phi> i) (f \<triangleright> n) \<down>= Suc j)"
define n\<^sub>0 where "n\<^sub>0 = Least ?Q"
with ex_n0 have n0: "?Q n\<^sub>0" "\<forall>n<n\<^sub>0. \<not> ?Q n"
using LeastI_ex[of ?Q] not_less_Least[of _ ?Q] by blast+
define m\<^sub>0 where "m\<^sub>0 = (LEAST m\<^sub>0. \<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n\<^sub>0)"
(is "m\<^sub>0 = Least ?P")
moreover have "\<exists>m\<^sub>0. \<forall>m\<ge>m\<^sub>0. the (delay i (f \<triangleright> m)) > n\<^sub>0"
- using delay_unbounded_monotone `f\<in>\<R>` \<open>f \<in> U\<close> assms learn_finE(1)
+ using delay_unbounded_monotone \<open>f\<in>\<R>\<close> \<open>f \<in> U\<close> assms learn_finE(1)
by simp
ultimately have m0: "?P m\<^sub>0" "\<forall>m<m\<^sub>0. \<not> ?P m"
using LeastI_ex[of ?P] not_less_Least[of _ ?P] by blast+
then have "\<forall>m\<ge>m\<^sub>0. totalizer 0 i (f \<triangleright> m) = \<phi> i (e_take (the (delay i (f \<triangleright> m))) (f \<triangleright> m))"
using totalizer_def by auto
then have "\<forall>m\<ge>m\<^sub>0. totalizer 0 i (f \<triangleright> m) = \<phi> i (f \<triangleright> (the (delay i (f \<triangleright> m)) - 1))"
- using e_take_delay_init m0 `f\<in>\<R>` by auto
+ using e_take_delay_init m0 \<open>f\<in>\<R>\<close> by auto
with m0 n0 have "\<forall>m\<ge>m\<^sub>0. totalizer 0 i (f \<triangleright> m) \<down>= Suc j"
by auto
moreover have "totalizer 0 i (f \<triangleright> m) \<down>= 0" if "m < m\<^sub>0" for m
proof (cases "the (delay i (f \<triangleright> m)) = 0")
case True
then show ?thesis by (simp add: totalizer_def)
next
case False
then have "the (delay i (f \<triangleright> m)) \<le> n\<^sub>0"
- using m0 that `f \<in> \<R>` delay_monotone by (meson leI order.strict_trans2)
+ using m0 that \<open>f \<in> \<R>\<close> delay_monotone by (meson leI order.strict_trans2)
then show ?thesis
using \<open>f \<in> \<R>\<close> n0(1) totalizer_init by (simp add: Suc_le_lessD)
qed
ultimately show ?thesis using j by auto
qed
qed
subsection \<open>Weaker Lemma R for CP and TOTAL\<close>
text \<open>For TOTAL the default hypothesis used by @{term totalizer}
depends on the hypothesis space, because it must refer to a total function in
that space. Consequently the total strategy depends on the hypothesis space,
which makes this form of Lemma~R weaker than the ones in the previous
section.\<close>
lemma lemma_R_for_TOTAL:
fixes \<psi> :: partial2
shows "\<exists>d. \<forall>U. \<forall>i. learn_total \<psi> U (\<phi> i) \<longrightarrow> learn_total \<psi> U (totalizer d i)"
proof (cases "\<exists>d. \<psi> d \<in> \<R>")
case True
then obtain d where "\<psi> d \<in> \<R>" by auto
have "learn_total \<psi> U (totalizer d i)" if "learn_total \<psi> U (\<phi> i)" for U i
proof (rule learn_totalI)
show env: "environment \<psi> U (totalizer d i)"
using that learn_totalE(1) totalizer_in_R2 by auto
show "\<And>f. f \<in> U \<Longrightarrow> \<exists>j. \<psi> j = f \<and> (\<forall>\<^sup>\<infinity>n. totalizer d i (f \<triangleright> n) \<down>= j)"
using that learn_total_def lemma_R_for_Lim[where ?d=d] learn_limE(2) by metis
show "\<psi> (the (totalizer d i (f \<triangleright> n))) \<in> \<R>" if "f \<in> U" for f n
proof (cases "the (delay i (f \<triangleright> n)) = 0")
case True
- then show ?thesis using totalizer_def `\<psi> d \<in> \<R>` by simp
+ then show ?thesis using totalizer_def \<open>\<psi> d \<in> \<R>\<close> by simp
next
case False
have "f \<in> \<R>"
using that env by auto
then show ?thesis
- using False that `learn_total \<psi> U (\<phi> i)` totalizer_init learn_totalE(3)
+ using False that \<open>learn_total \<psi> U (\<phi> i)\<close> totalizer_init learn_totalE(3)
by simp
qed
qed
then show ?thesis by auto
next
case False
then show ?thesis using learn_total_def lemma_R_for_Lim by auto
qed
corollary lemma_R_for_TOTAL_simple:
assumes "learn_total \<psi> U s"
shows "\<exists>s'\<in>\<R>. learn_total \<psi> U s'"
using assms lemma_R_for_TOTAL totalizer_in_R2
by (metis R2_proj_R1 learn_totalE(1) phi_universal)
text \<open>For CP the default hypothesis used by @{term totalizer} depends
on both the hypothesis space and the class. Therefore the total strategy
depends on both the the hypothesis space and the class, which makes Lemma~R
for CP even weaker than the one for TOTAL.\<close>
lemma lemma_R_for_CP:
fixes \<psi> :: partial2 and U :: "partial1 set"
assumes "learn_cp \<psi> U (\<phi> i)"
shows "\<exists>d. learn_cp \<psi> U (totalizer d i)"
proof (cases "U = {}")
case True
then show ?thesis using assms learn_cp_def lemma_R_for_Lim by auto
next
case False
then obtain f where "f \<in> U" by auto
- from `f \<in> U` obtain d where "\<psi> d = f"
+ from \<open>f \<in> U\<close> obtain d where "\<psi> d = f"
using learn_cpE(2)[OF assms] by auto
- with `f \<in> U` have "\<psi> d \<in> U" by simp
+ with \<open>f \<in> U\<close> have "\<psi> d \<in> U" by simp
have "learn_cp \<psi> U (totalizer d i)"
proof (rule learn_cpI)
show env: "environment \<psi> U (totalizer d i)"
using assms learn_cpE(1) totalizer_in_R2 by auto
show "\<And>f. f \<in> U \<Longrightarrow> \<exists>j. \<psi> j = f \<and> (\<forall>\<^sup>\<infinity>n. totalizer d i (f \<triangleright> n) \<down>= j)"
using assms learn_cp_def lemma_R_for_Lim[where ?d=d] learn_limE(2) by metis
show "\<psi> (the (totalizer d i (f \<triangleright> n))) \<in> U" if "f \<in> U" for f n
proof (cases "the (delay i (f \<triangleright> n)) = 0")
case True
- then show ?thesis using totalizer_def `\<psi> d \<in> U` by simp
+ then show ?thesis using totalizer_def \<open>\<psi> d \<in> U\<close> by simp
next
case False
then show ?thesis
using that env assms totalizer_init learn_cpE(3) by auto
qed
qed
then show ?thesis by auto
qed
subsection \<open>No Lemma R for CONS\<close>
text \<open>This section demonstrates that the class $V_{01}$ of all total
recursive functions $f$ where $f(0)$ or $f(1)$ is a Gödel number of $f$ can
be consistently learned in the limit, but not by a total strategy. This implies
that Lemma~R does not hold for CONS.\<close>
definition V01 :: "partial1 set" ("V\<^sub>0\<^sub>1") where
"V\<^sub>0\<^sub>1 = {f. f \<in> \<R> \<and> (\<phi> (the (f 0)) = f \<or> \<phi> (the (f 1)) = f)}"
subsubsection \<open>No total CONS strategy for @{term "V\<^sub>0\<^sub>1"}\label{s:v01_not_total}\<close>
text \<open>In order to show that no total strategy can learn @{term
"V\<^sub>0\<^sub>1"} we construct, for each total strategy $S$, one or two
functions in @{term "V\<^sub>0\<^sub>1"} such that $S$ fails for at least one
of them. At the core of this construction is a process that given a total
recursive strategy $S$ and numbers $z, i, j \in \mathbb{N}$ builds a function
$f$ as follows: Set $f(0) = i$ and $f(1) = j$. For $x\geq1$:
\begin{enumerate}
\item[(a)] Check whether $S$ changes its hypothesis when $f^x$ is
extended by 0, that is, if $S(f^x) \neq S(f^x0)$. If so, set $f(x+1) = 0$.
\item[(b)] Otherwise check if $S$ changes its hypothesis when $f^x$ is extended
by $1$, that is, if $S(f^x) \neq S(f^x1)$. If so, set $f(x+1) = 1$.
\item[(c)] If neither happens, set $f(x+1) = z$.
\end{enumerate}
In other words, as long as we can force $S$ to change its hypothesis by
extending the function by 0 or 1, we do just that. Now there are two
cases:
\begin{enumerate}
\item[Case 1.] For all $x\geq1$ either (a) or (b) occurs; then $S$
changes its hypothesis on $f$ all the time and thus does not learn $f$ in
the limit (not to mention consistently). The value of $z$ makes no
difference in this case.
\item[Case 2.] For some minimal $x$, (c) occurs, that is,
there is an $f^x$ such that $h := S(f^x) = S(f^x0) = S(f^x1)$. But the
hypothesis $h$ cannot be consistent with both prefixes $f^x0$ and $f^x1$.
Running the process once with $z = 0$ and once with $z = 1$ yields two
functions starting with $f^x0$ and $f^x1$, respectively, such that $S$
outputs the same hypothesis, $h$, on both prefixes and thus cannot be
consistent for both functions.
\end{enumerate}
This process is computable because $S$ is total. The construction does not
work if we only assume $S$ to be a CONS strategy for $V_{01}$, because we
need to be able to apply $S$ to prefixes not in $V_{01}$.
The parameters $i$ and $j$ provide flexibility to find functions built by the
above process that are actually in $V_{01}$. To this end we will use
Smullyan's double fixed-point theorem.\<close>
context
fixes s :: partial1
assumes s_in_R1 [simp, intro]: "s \<in> \<R>"
begin
text \<open>The function @{term prefixes} constructs prefixes according to the
aforementioned process.\<close>
fun prefixes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list" where
"prefixes z i j 0 = [i]"
| "prefixes z i j (Suc x) = prefixes z i j x @
[if x = 0 then j
else if s (list_encode (prefixes z i j x @ [0])) \<noteq> s (list_encode (prefixes z i j x))
then 0
else if s (list_encode (prefixes z i j x @ [1])) \<noteq> s (list_encode (prefixes z i j x))
then 1
else z]"
lemma prefixes_length: "length (prefixes z i j x) = Suc x"
by (induction x) simp_all
text \<open>The functions @{term[names_short] "adverse z i j"} are the
functions constructed by @{term[names_short] "prefixes"}.\<close>
definition adverse :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat option" where
"adverse z i j x \<equiv> Some (last (prefixes z i j x))"
lemma init_adverse_eq_prefixes: "(adverse z i j) \<triangleright> n = list_encode (prefixes z i j n)"
proof -
have "prefix (adverse z i j) n = prefixes z i j n"
proof (induction n)
case 0
then show ?case using adverse_def prefixes_length prefixI' by fastforce
next
case (Suc n)
then show ?case using adverse_def by (simp add: prefix_Suc)
qed
then show ?thesis by (simp add: init_def)
qed
lemma adverse_at_01:
"adverse z i j 0 \<down>= i"
"adverse z i j 1 \<down>= j"
by (auto simp add: adverse_def)
text \<open>Had we introduced ternary partial recursive functions, the
@{term[names_short] "adverse z"} functions would be among them.\<close>
lemma adverse_in_R3: "\<exists>r. recfn 3 r \<and> total r \<and> (\<lambda>i j x. eval r [i, j, x]) = adverse z"
proof -
obtain rs where rs: "recfn 1 rs" "total rs" "(\<lambda>x. eval rs [x]) = s"
using R1E by auto
have s_total: "\<And>x. s x \<down>" by simp
define f where "f = Cn 2 r_singleton_encode [Id 2 0]"
then have "recfn 2 f" by simp
have f: "\<And>i j. eval f [i, j] \<down>= list_encode [i]"
unfolding f_def by simp
define ch1 where "ch1 = Cn 4 r_ifeq
[Cn 4 rs [Cn 4 r_snoc [Id 4 1, r_constn 3 1]],
Cn 4 rs [Id 4 1],
r_dummy 3 (r_const z),
r_constn 3 1]"
then have ch1: "recfn 4 ch1" "total ch1"
using Cn_total prim_recfn_total rs by auto
define ch0 where "ch0 = Cn 4 r_ifeq
[Cn 4 rs [Cn 4 r_snoc [Id 4 1, r_constn 3 0]],
Cn 4 rs [Id 4 1],
ch1,
r_constn 3 0]"
then have ch0_total: "total ch0" "recfn 4 ch0"
using Cn_total prim_recfn_total rs ch1 by auto
have "eval ch1 [l, v, i, j] \<down>= (if s (e_snoc v 1) = s v then z else 1)" for l v i j
proof -
have "eval ch1 [l, v, i, j] = eval r_ifeq [the (s (e_snoc v 1)), the (s v), z, 1]"
unfolding ch1_def using rs by auto
then show ?thesis by (simp add: s_total option.expand)
qed
moreover have "eval ch0 [l, v, i, j] \<down>=
(if s (e_snoc v 0) = s v then the (eval ch1 [l, v, i, j]) else 0)" for l v i j
proof -
have "eval ch0 [l, v, i, j] =
eval r_ifeq [the (s (e_snoc v 0)), the (s v), the (eval ch1 [l, v, i, j]), 0]"
unfolding ch0_def using rs ch1 by auto
then show ?thesis by (simp add: s_total option.expand)
qed
ultimately have ch0: "\<And>l v i j. eval ch0 [l, v, i, j] \<down>=
(if s (e_snoc v 0) \<noteq> s v then 0
else if s (e_snoc v 1) \<noteq> s v then 1 else z)"
by simp
define app where "app = Cn 4 r_ifz [Id 4 0, Id 4 3, ch0]"
then have "recfn 4 app" "total app"
using ch0_total totalI4 by auto
have "eval app [l, v, i, j] \<down>= (if l = 0 then j else the (eval ch0 [l, v, i, j]))" for l v i j
unfolding app_def using ch0_total by simp
with ch0 have app: "\<And>l v i j. eval app [l, v, i, j] \<down>=
(if l = 0 then j
else if s (e_snoc v 0) \<noteq> s v then 0
else if s (e_snoc v 1) \<noteq> s v then 1 else z)"
by simp
define g where "g = Cn 4 r_snoc [Id 4 1, app]"
with app have g: "\<And>l v i j. eval g [l, v, i, j] \<down>= e_snoc v
(if l = 0 then j
else if s (e_snoc v 0) \<noteq> s v then 0
else if s (e_snoc v 1) \<noteq> s v then 1 else z)"
- using `recfn 4 app` by auto
+ using \<open>recfn 4 app\<close> by auto
from g_def have "recfn 4 g" "total g"
- using `recfn 4 app` `total app` Cn_total Mn_free_imp_total by auto
+ using \<open>recfn 4 app\<close> \<open>total app\<close> Cn_total Mn_free_imp_total by auto
define b where "b = Pr 2 f g"
then have "recfn 3 b"
- using `recfn 2 f` `recfn 4 g` by simp
+ using \<open>recfn 2 f\<close> \<open>recfn 4 g\<close> by simp
have b: "eval b [x, i, j] \<down>= list_encode (prefixes z i j x)" for x i j
proof (induction x)
case 0
then show ?case
- unfolding b_def using f `recfn 2 f` \<open>recfn 4 g\<close> by simp
+ unfolding b_def using f \<open>recfn 2 f\<close> \<open>recfn 4 g\<close> by simp
next
case (Suc x)
then have "eval b [Suc x, i, j] = eval g [x, the (eval b [x, i, j]), i, j]"
- using b_def `recfn 3 b` by simp
+ using b_def \<open>recfn 3 b\<close> by simp
also have "... \<down>=
(let v = list_encode (prefixes z i j x)
in e_snoc v
(if x = 0 then j
else if s (e_snoc v 0) \<noteq> s v then 0
else if s (e_snoc v 1) \<noteq> s v then 1 else z))"
using g Suc by simp
also have "... \<down>=
(let v = list_encode (prefixes z i j x)
in e_snoc v
(if x = 0 then j
else if s (list_encode (prefixes z i j x @ [0])) \<noteq> s v then 0
else if s (list_encode (prefixes z i j x @ [1])) \<noteq> s v then 1 else z))"
using list_decode_encode by presburger
finally show ?case by simp
qed
define b' where "b' = Cn 3 b [Id 3 2, Id 3 0, Id 3 1]"
then have "recfn 3 b'"
- using `recfn 3 b` by simp
+ using \<open>recfn 3 b\<close> by simp
with b have b': "\<And>i j x. eval b' [i, j, x] \<down>= list_encode (prefixes z i j x)"
using b'_def by simp
define r where "r = Cn 3 r_last [b']"
then have "recfn 3 r"
- using `recfn 3 b'` by simp
+ using \<open>recfn 3 b'\<close> by simp
with b' have "\<And>i j x. eval r [i, j, x] \<down>= last (prefixes z i j x)"
using r_def prefixes_length by auto
moreover from this have "total r"
- using totalI3 `recfn 3 r` by simp
+ using totalI3 \<open>recfn 3 r\<close> by simp
ultimately have "(\<lambda>i j x. eval r [i, j, x]) = adverse z"
unfolding adverse_def by simp
- with `recfn 3 r` `total r` show ?thesis by auto
+ with \<open>recfn 3 r\<close> \<open>total r\<close> show ?thesis by auto
qed
lemma adverse_in_R1: "adverse z i j \<in> \<R>"
proof -
from adverse_in_R3 obtain r where
r: "recfn 3 r" "total r" "(\<lambda>i j x. eval r [i, j, x]) = adverse z"
by blast
define rij where "rij = Cn 1 r [r_const i, r_const j, Id 1 0]"
then have "recfn 1 rij" "total rij"
using r(1,2) Cn_total Mn_free_imp_total by auto
from rij_def have "\<And>x. eval rij [x] = eval r [i, j, x]"
using r(1) by auto
with r(3) have "\<And>x. eval rij [x] = adverse z i j x"
by metis
- with `recfn 1 rij` `total rij` show ?thesis by auto
+ with \<open>recfn 1 rij\<close> \<open>total rij\<close> show ?thesis by auto
qed
text \<open>Next we show that for every $z$ there are $i$, $j$ such that
@{term[names_short] "adverse z i j \<in> V\<^sub>0\<^sub>1"}. The first step is to show that for every
$z$, Gödel numbers for @{term[names_short] "adverse z i j"} can be computed
uniformly from $i$ and $j$.\<close>
lemma phi_translate_adverse: "\<exists>f\<in>\<R>\<^sup>2.\<forall>i j. \<phi> (the (f i j)) = adverse z i j"
proof -
obtain r where r: "recfn 3 r" "total r" "(\<lambda>i j x. eval r [i, j, x]) = adverse z"
using adverse_in_R3 by blast
let ?p = "encode r"
define rf where "rf = Cn 2 (r_smn 1 2) [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]"
then have "recfn 2 rf" and "total rf"
using Mn_free_imp_total by simp_all
define f where "f \<equiv> \<lambda>i j. eval rf [i, j]"
- with `recfn 2 rf` `total rf` have "f \<in> \<R>\<^sup>2" by auto
+ with \<open>recfn 2 rf\<close> \<open>total rf\<close> have "f \<in> \<R>\<^sup>2" by auto
have rf: "eval rf [i, j] = eval (r_smn 1 2) [?p, i, j]" for i j
unfolding rf_def by simp
{
fix i j x
have "\<phi> (the (f i j)) x = eval r_phi [the (f i j), x]"
using phi_def by simp
also have "... = eval r_phi [the (eval rf [i, j]), x]"
using f_def by simp
also have "... = eval (r_universal 1) [the (eval (r_smn 1 2) [?p, i, j]), x]"
using rf r_phi_def by simp
also have "... = eval (r_universal (2 + 1)) (?p # [i, j] @ [x])"
using smn_lemma[of 1 "[i, j]" 2 "[x]"] by simp
also have "... = eval (r_universal 3) [?p, i, j, x]"
by simp
also have "... = eval r [i, j, x]"
using r_universal r by force
also have "... = adverse z i j x"
using r(3) by metis
finally have "\<phi> (the (f i j)) x = adverse z i j x" .
}
- with `f \<in> \<R>\<^sup>2` show ?thesis by blast
+ with \<open>f \<in> \<R>\<^sup>2\<close> show ?thesis by blast
qed
text \<open>The second, and final, step is to apply Smullyan's double
fixed-point theorem to show the existence of @{term[names_short] adverse}
functions in @{term "V\<^sub>0\<^sub>1"}.\<close>
lemma adverse_in_V01: "\<exists>m n. adverse 0 m n \<in> V\<^sub>0\<^sub>1 \<and> adverse 1 m n \<in> V\<^sub>0\<^sub>1"
proof -
obtain f\<^sub>0 where f0: "f\<^sub>0 \<in> \<R>\<^sup>2" "\<forall>i j. \<phi> (the (f\<^sub>0 i j)) = adverse 0 i j"
using phi_translate_adverse[of 0] by auto
obtain f\<^sub>1 where f1: "f\<^sub>1 \<in> \<R>\<^sup>2" "\<forall>i j. \<phi> (the (f\<^sub>1 i j)) = adverse 1 i j"
using phi_translate_adverse[of 1] by auto
obtain m n where "\<phi> m = \<phi> (the (f\<^sub>0 m n))" and "\<phi> n = \<phi> (the (f\<^sub>1 m n))"
using smullyan_double_fixed_point[OF f0(1) f1(1)] by blast
with f0(2) f1(2) have "\<phi> m = adverse 0 m n" and "\<phi> n = adverse 1 m n"
by simp_all
moreover have "the (adverse 0 m n 0) = m" and "the (adverse 1 m n 1) = n"
using adverse_at_01 by simp_all
ultimately have
"\<phi> (the (adverse 0 m n 0)) = adverse 0 m n"
"\<phi> (the (adverse 1 m n 1)) = adverse 1 m n"
by simp_all
moreover have "adverse 0 m n \<in> \<R>" and "adverse 1 m n \<in> \<R>"
using adverse_in_R1 by simp_all
ultimately show ?thesis using V01_def by auto
qed
text \<open>Before we prove the main result of this section we need some
lemmas regarding the shape of the @{term[names_short] adverse} functions and
hypothesis changes of the strategy.\<close>
lemma adverse_Suc:
assumes "x > 0"
shows "adverse z i j (Suc x) \<down>=
(if s (e_snoc ((adverse z i j) \<triangleright> x) 0) \<noteq> s ((adverse z i j) \<triangleright> x)
then 0
else if s (e_snoc ((adverse z i j) \<triangleright> x) 1) \<noteq> s ((adverse z i j) \<triangleright> x)
then 1 else z)"
proof -
have "adverse z i j (Suc x) \<down>=
(if s (list_encode (prefixes z i j x @ [0])) \<noteq> s (list_encode (prefixes z i j x))
then 0
else if s (list_encode (prefixes z i j x @ [1])) \<noteq> s (list_encode (prefixes z i j x))
then 1 else z)"
using assms adverse_def by simp
then show ?thesis by (simp add: init_adverse_eq_prefixes)
qed
text \<open>The process in the proof sketch (page~\pageref{s:v01_not_total})
consists of steps (a), (b), and (c). The next abbreviation is true iff.\ step
(a) or (b) applies.\<close>
abbreviation "hyp_change z i j x \<equiv>
s (e_snoc ((adverse z i j) \<triangleright> x) 0) \<noteq> s ((adverse z i j) \<triangleright> x) \<or>
s (e_snoc ((adverse z i j) \<triangleright> x) 1) \<noteq> s ((adverse z i j) \<triangleright> x)"
text \<open>If step (c) applies, the process appends $z$.\<close>
lemma adverse_Suc_not_hyp_change:
assumes "x > 0" and "\<not> hyp_change z i j x"
shows "adverse z i j (Suc x) \<down>= z"
using assms adverse_Suc by simp
text \<open>While (a) or (b) applies, the process appends a value that
forces $S$ to change its hypothesis.\<close>
lemma while_hyp_change:
assumes "\<forall>x\<le>n. x > 0 \<longrightarrow> hyp_change z i j x"
shows "\<forall>x\<le>Suc n. adverse z i j x = adverse z' i j x"
using assms
proof (induction n)
case 0
then show ?case by (simp add: adverse_def le_Suc_eq)
next
case (Suc n)
then have "\<forall>x\<le>n. x > 0 \<longrightarrow> hyp_change z i j x" by simp
with Suc have "\<forall>x\<le>Suc n. x > 0 \<longrightarrow> adverse z i j x = adverse z' i j x"
by simp
moreover have "adverse z i j 0 = adverse z' i j 0"
using adverse_at_01 by simp
ultimately have zz': "\<forall>x\<le>Suc n. adverse z i j x = adverse z' i j x"
by auto
moreover have "adverse z i j \<in> \<R>" "adverse z' i j \<in> \<R>"
using adverse_in_R1 by simp_all
ultimately have init_zz': "(adverse z i j) \<triangleright> (Suc n) = (adverse z' i j) \<triangleright> (Suc n)"
using init_eqI by blast
have "adverse z i j (Suc (Suc n)) = adverse z' i j (Suc (Suc n))"
proof (cases "s (e_snoc ((adverse z i j) \<triangleright> (Suc n)) 0) \<noteq> s ((adverse z i j) \<triangleright> (Suc n))")
case True
then have "s (e_snoc ((adverse z' i j) \<triangleright> (Suc n)) 0) \<noteq> s ((adverse z' i j) \<triangleright> (Suc n))"
using init_zz' by simp
then have "adverse z' i j (Suc (Suc n)) \<down>= 0"
by (simp add: adverse_Suc)
moreover have "adverse z i j (Suc (Suc n)) \<down>= 0"
using True by (simp add: adverse_Suc)
ultimately show ?thesis by simp
next
case False
then have "s (e_snoc ((adverse z' i j) \<triangleright> (Suc n)) 0) = s ((adverse z' i j) \<triangleright> (Suc n))"
using init_zz' by simp
then have "adverse z' i j (Suc (Suc n)) \<down>= 1"
using init_zz' Suc.prems adverse_Suc by (smt le_refl zero_less_Suc)
moreover have "adverse z i j (Suc (Suc n)) \<down>= 1"
using False Suc.prems adverse_Suc by auto
ultimately show ?thesis by simp
qed
with zz' show ?case using le_SucE by blast
qed
text \<open>The next result corresponds to Case~1 from the proof sketch.\<close>
lemma always_hyp_change_no_lim:
assumes "\<forall>x>0. hyp_change z i j x"
shows "\<not> learn_lim \<phi> {adverse z i j} s"
proof (rule infinite_hyp_changes_not_Lim[of "adverse z i j"])
show "adverse z i j \<in> {adverse z i j}" by simp
show "\<forall>n. \<exists>m\<^sub>1>n. \<exists>m\<^sub>2>n. s (adverse z i j \<triangleright> m\<^sub>1) \<noteq> s (adverse z i j \<triangleright> m\<^sub>2)"
proof
fix n
from assms obtain m\<^sub>1 where m1: "m\<^sub>1 > n" "hyp_change z i j m\<^sub>1"
by auto
have "s (adverse z i j \<triangleright> m\<^sub>1) \<noteq> s (adverse z i j \<triangleright> (Suc m\<^sub>1))"
proof (cases "s (e_snoc ((adverse z i j) \<triangleright> m\<^sub>1) 0) \<noteq> s ((adverse z i j) \<triangleright> m\<^sub>1)")
case True
then have "adverse z i j (Suc m\<^sub>1) \<down>= 0"
using m1 adverse_Suc by simp
then have "(adverse z i j) \<triangleright> (Suc m\<^sub>1) = e_snoc ((adverse z i j) \<triangleright> m\<^sub>1) 0"
by (simp add: init_Suc_snoc)
with True show ?thesis by simp
next
case False
then have "adverse z i j (Suc m\<^sub>1) \<down>= 1"
using m1 adverse_Suc by simp
then have "(adverse z i j) \<triangleright> (Suc m\<^sub>1) = e_snoc ((adverse z i j) \<triangleright> m\<^sub>1) 1"
by (simp add: init_Suc_snoc)
with False m1(2) show ?thesis by simp
qed
then show "\<exists>m\<^sub>1>n. \<exists>m\<^sub>2>n. s (adverse z i j \<triangleright> m\<^sub>1) \<noteq> s (adverse z i j \<triangleright> m\<^sub>2)"
using less_SucI m1(1) by blast
qed
qed
text \<open>The next result corresponds to Case~2 from the proof sketch.\<close>
lemma no_hyp_change_no_cons:
assumes "x > 0" and "\<not> hyp_change z i j x"
shows "\<not> learn_cons \<phi> {adverse 0 i j, adverse 1 i j} s"
proof -
let ?P = "\<lambda>x. x > 0 \<and> \<not> hyp_change z i j x"
define xmin where "xmin = Least ?P"
with assms have xmin:
"?P xmin"
"\<And>x. x < xmin \<Longrightarrow> \<not> ?P x"
using LeastI[of ?P] not_less_Least[of _ ?P] by simp_all
then have "xmin > 0" by simp
have "\<forall>x\<le>xmin - 1. x > 0 \<longrightarrow> hyp_change z i j x"
using xmin by (metis One_nat_def Suc_pred le_imp_less_Suc)
then have
"\<forall>x\<le>xmin. adverse z i j x = adverse 0 i j x"
"\<forall>x\<le>xmin. adverse z i j x = adverse 1 i j x"
using while_hyp_change[of "xmin - 1" z i j 0]
using while_hyp_change[of "xmin - 1" z i j 1]
by simp_all
then have
init_z0: "(adverse z i j) \<triangleright> xmin = (adverse 0 i j) \<triangleright> xmin" and
init_z1: "(adverse z i j) \<triangleright> xmin = (adverse 1 i j) \<triangleright> xmin"
using adverse_in_R1 init_eqI by blast+
then have
a0: "adverse 0 i j (Suc xmin) \<down>= 0" and
a1: "adverse 1 i j (Suc xmin) \<down>= 1"
using adverse_Suc_not_hyp_change xmin(1) init_z1
by metis+
then have
i0: "(adverse 0 i j) \<triangleright> (Suc xmin) = e_snoc ((adverse z i j) \<triangleright> xmin) 0" and
i1: "(adverse 1 i j) \<triangleright> (Suc xmin) = e_snoc ((adverse z i j) \<triangleright> xmin) 1"
using init_z0 init_z1 by (simp_all add: init_Suc_snoc)
moreover have
"s (e_snoc ((adverse z i j) \<triangleright> xmin) 0) = s ((adverse z i j) \<triangleright> xmin)"
"s (e_snoc ((adverse z i j) \<triangleright> xmin) 1) = s ((adverse z i j) \<triangleright> xmin)"
using xmin by simp_all
ultimately have
"s ((adverse 0 i j) \<triangleright> (Suc xmin)) = s ((adverse z i j) \<triangleright> xmin)"
"s ((adverse 1 i j) \<triangleright> (Suc xmin)) = s ((adverse z i j) \<triangleright> xmin)"
by simp_all
then have
"s ((adverse 0 i j) \<triangleright> (Suc xmin)) = s ((adverse 1 i j) \<triangleright> (Suc xmin))"
by simp
moreover have "(adverse 0 i j) \<triangleright> (Suc xmin) \<noteq> (adverse 1 i j) \<triangleright> (Suc xmin)"
using a0 a1 i0 i1 by (metis append1_eq_conv list_decode_encode zero_neq_one)
ultimately show "\<not> learn_cons \<phi> {adverse 0 i j, adverse 1 i j} s"
using same_hyp_different_init_not_cons by blast
qed
text \<open>Combining the previous two lemmas shows that @{term
"V\<^sub>0\<^sub>1"} cannot be learned consistently in the limit by the total
strategy $S$.\<close>
lemma V01_not_in_R_cons: "\<not> learn_cons \<phi> V\<^sub>0\<^sub>1 s"
proof -
obtain m n where
mn0: "adverse 0 m n \<in> V\<^sub>0\<^sub>1" and
mn1: "adverse 1 m n \<in> V\<^sub>0\<^sub>1"
using adverse_in_V01 by auto
show "\<not> learn_cons \<phi> V\<^sub>0\<^sub>1 s"
proof (cases "\<forall>x>0. hyp_change 0 m n x")
case True
then have "\<not> learn_lim \<phi> {adverse 0 m n} s"
using always_hyp_change_no_lim by simp
with mn0 show ?thesis
using learn_cons_def learn_lim_closed_subseteq by auto
next
case False
then obtain x where x: "x > 0" "\<not> hyp_change 0 m n x" by auto
then have "\<not> learn_cons \<phi> {adverse 0 m n, adverse 1 m n} s"
using no_hyp_change_no_cons[OF x] by simp
with mn0 mn1 show ?thesis using learn_cons_closed_subseteq by auto
qed
qed
end
subsubsection \<open>@{term "V\<^sub>0\<^sub>1"} is in CONS\<close>
text \<open>At first glance, consistently learning @{term "V\<^sub>0\<^sub>1"} looks fairly
easy. After all every @{term "f \<in> V\<^sub>0\<^sub>1"} provides a Gödel number of itself
either at argument 0 or 1. A strategy only has to figure out which one is
right. However, the strategy $S$ we are going to devise does not always
converge to $f(0)$ or $f(1)$. Instead it uses a technique called
``amalgamation''. The amalgamation of two Gödel numbers $i$ and $j$ is a
function whose value at $x$ is determined by simulating $\varphi_i(x)$ and
$\varphi_j(x)$ in parallel and outputting the value of the first one to halt.
If neither halts the value is undefined. There is a function
$a\in\mathcal{R}^2$ such that $\varphi_{a(i,j)}$ is the amalgamation of $i$
and $j$.
If @{term "f \<in> V\<^sub>0\<^sub>1"} then $\varphi_{a(f(0), f(1))}$ is
total because by definition of @{term "V\<^sub>0\<^sub>1"} we have
$\varphi_{f(0)} = f$ or $\varphi_{f(1)} = f$ and $f$ is total.
Given a prefix $f^n$ of an @{term "f \<in> V\<^sub>0\<^sub>1"} the strategy
$S$ first computes $\varphi_{a(f(0), f(1))}(x)$ for $x = 0, \ldots, n$. For
the resulting prefix $\varphi^n_{a(f(0), f(1))}$ there are two cases:
\begin{enumerate}
\item[Case 1.] It differs from $f^n$, say at minimum index $x$. Then for
either $z = 0$ or $z = 1$ we have $\varphi_{f(z)}(x) \neq f(x)$ by
definition of amalgamation. This
implies $\varphi_{f(z)} \neq f$, and thus $\varphi_{f(1-z)} = f$ by
definition of @{term "V\<^sub>0\<^sub>1"}. We set $S(f^n) = f(1 - z)$. This
hypothesis is correct and hence consistent.
\item[Case 2.] It equals $f^n$. Then we set $S(f^n) = a(f(0), f(1))$. This
hypothesis is consistent by definition of this case.
\end{enumerate}
In both cases the hypothesis is consistent. If Case~1 holds for some $n$, the
same $x$ and $z$ will be found also for all larger values of $n$. Therefore
$S$ converges to the correct hypothesis $f(1 - z)$. If Case~2 holds for all
$n$, then $S$ always outputs the same hypothesis $a(f(0), f(1))$ and thus
also converges.
The above discussion tacitly assumes $n \geq 1$, such that both $f(0)$ and
$f(1)$ are available to $S$. For $n = 0$ the strategy outputs an arbitrary
consistent hypothesis.\<close>
text \<open>Amalgamation uses the concurrent simulation of functions.\<close>
definition parallel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat option" where
"parallel i j x \<equiv> eval r_parallel [i, j, x]"
lemma r_parallel': "eval r_parallel [i, j, x] = parallel i j x"
using parallel_def by simp
lemma r_parallel'':
shows "eval r_phi [i, x] \<up> \<and> eval r_phi [j, x] \<up> \<Longrightarrow> eval r_parallel [i, j, x] \<up>"
and "eval r_phi [i, x] \<down> \<and> eval r_phi [j, x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval r_phi [i, x]))"
and "eval r_phi [j, x] \<down> \<and> eval r_phi [i, x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval r_phi [j, x]))"
and "eval r_phi [i, x] \<down> \<and> eval r_phi [j, x] \<down> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval r_phi [i, x])) \<or>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval r_phi [j, x]))"
proof -
let ?f = "Cn 1 r_phi [r_const i, Id 1 0]"
let ?g = "Cn 1 r_phi [r_const j, Id 1 0]"
have *: "\<And>x. eval r_phi [i, x] = eval ?f [x]" "\<And>x. eval r_phi [j, x] = eval ?g [x]"
by simp_all
show "eval r_phi [i, x] \<up> \<and> eval r_phi [j, x] \<up> \<Longrightarrow> eval r_parallel [i, j, x] \<up>"
and "eval r_phi [i, x] \<down> \<and> eval r_phi [j, x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval r_phi [i, x]))"
and "eval r_phi [j, x] \<down> \<and> eval r_phi [i, x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval r_phi [j, x]))"
and "eval r_phi [i, x] \<down> \<and> eval r_phi [j, x] \<down> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval r_phi [i, x])) \<or>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval r_phi [j, x]))"
using r_parallel[OF *] by simp_all
qed
lemma parallel:
"\<phi> i x \<up> \<and> \<phi> j x \<up> \<Longrightarrow> parallel i j x \<up>"
"\<phi> i x \<down> \<and> \<phi> j x \<up> \<Longrightarrow> parallel i j x \<down>= prod_encode (0, the (\<phi> i x))"
"\<phi> j x \<down> \<and> \<phi> i x \<up> \<Longrightarrow> parallel i j x \<down>= prod_encode (1, the (\<phi> j x))"
"\<phi> i x \<down> \<and> \<phi> j x \<down> \<Longrightarrow>
parallel i j x \<down>= prod_encode (0, the (\<phi> i x)) \<or>
parallel i j x \<down>= prod_encode (1, the (\<phi> j x))"
using phi_def r_parallel'' r_parallel parallel_def by simp_all
lemma parallel_converg_pdec1_0_or_1:
assumes "parallel i j x \<down>"
shows "pdec1 (the (parallel i j x)) = 0 \<or> pdec1 (the (parallel i j x)) = 1"
using assms parallel[of i x j] parallel(3)[of j x i]
by (metis fst_eqD option.sel prod_encode_inverse)
lemma parallel_converg_either: "(\<phi> i x \<down> \<or> \<phi> j x \<down>) = (parallel i j x \<down>)"
using parallel by (metis option.simps(3))
lemma parallel_0:
assumes "parallel i j x \<down>= prod_encode (0, v)"
shows "\<phi> i x \<down>= v"
using parallel assms
by (smt option.collapse option.sel option.simps(3) prod.inject prod_encode_eq zero_neq_one)
lemma parallel_1:
assumes "parallel i j x \<down>= prod_encode (1, v)"
shows "\<phi> j x \<down>= v"
using parallel assms
by (smt option.collapse option.sel option.simps(3) prod.inject prod_encode_eq zero_neq_one)
lemma parallel_converg_V01:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "parallel (the (f 0)) (the (f 1)) x \<down>"
proof -
have "f \<in> \<R> \<and> (\<phi> (the (f 0)) = f \<or> \<phi> (the (f 1)) = f)"
using assms V01_def by auto
then have "\<phi> (the (f 0)) \<in> \<R> \<or> \<phi> (the (f 1)) \<in> \<R>"
by auto
then have "\<phi> (the (f 0)) x \<down> \<or> \<phi> (the (f 1)) x \<down>"
using R1_imp_total1 by auto
then show ?thesis using parallel_converg_either by simp
qed
text \<open>The amalgamation of two Gödel numbers can then be described
in terms of @{term "parallel"}.\<close>
definition amalgamation :: "nat \<Rightarrow> nat \<Rightarrow> partial1" where
"amalgamation i j x \<equiv>
if parallel i j x \<up> then None else Some (pdec2 (the (parallel i j x)))"
lemma amalgamation_diverg: "amalgamation i j x \<up> \<longleftrightarrow> \<phi> i x \<up> \<and> \<phi> j x \<up>"
using amalgamation_def parallel by (metis option.simps(3))
lemma amalgamation_total:
assumes "total1 (\<phi> i) \<or> total1 (\<phi> j)"
shows "total1 (amalgamation i j)"
using assms amalgamation_diverg[of i j] total_def by auto
lemma amalgamation_V01_total:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "total1 (amalgamation (the (f 0)) (the (f 1)))"
using assms V01_def amalgamation_total R1_imp_total1 total1_def
by (metis (mono_tags, lifting) mem_Collect_eq)
definition "r_amalgamation \<equiv> Cn 3 r_pdec2 [r_parallel]"
lemma r_amalgamation_recfn: "recfn 3 r_amalgamation"
unfolding r_amalgamation_def by simp
lemma r_amalgamation: "eval r_amalgamation [i, j, x] = amalgamation i j x"
proof (cases "parallel i j x \<up>")
case True
then have "eval r_parallel [i, j, x] \<up>"
by (simp add: r_parallel')
then have "eval r_amalgamation [i, j, x] \<up>"
unfolding r_amalgamation_def by simp
moreover from True have "amalgamation i j x \<up>"
using amalgamation_def by simp
ultimately show ?thesis by simp
next
case False
then have "eval r_parallel [i, j, x] \<down>"
by (simp add: r_parallel')
then have "eval r_amalgamation [i, j, x] = eval r_pdec2 [the (eval r_parallel [i, j, x])]"
unfolding r_amalgamation_def by simp
also have "... \<down>= pdec2 (the (eval r_parallel [i, j, x]))"
by simp
finally show ?thesis by (simp add: False amalgamation_def r_parallel')
qed
text \<open>The function @{term "amalgamate"} computes Gödel numbers of
amalgamations. It corresponds to the function $a$ from the proof sketch.\<close>
definition amalgamate :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"amalgamate i j \<equiv> smn 1 (encode r_amalgamation) [i, j]"
lemma amalgamate: "\<phi> (amalgamate i j) = amalgamation i j"
proof
fix x
have "\<phi> (amalgamate i j) x = eval r_phi [amalgamate i j, x]"
by (simp add: phi_def)
also have "... = eval r_phi [smn 1 (encode r_amalgamation) [i, j], x]"
using amalgamate_def by simp
also have "... = eval r_phi
[encode (Cn 1 (r_universal 3)
(r_constn 0 (encode r_amalgamation) # map (r_constn 0) [i, j] @ map (Id 1) [0])), x]"
using smn[of 1 "encode r_amalgamation" "[i, j]"] by (simp add: numeral_3_eq_3)
also have "... = eval r_phi
[encode (Cn 1 (r_universal 3)
(r_const (encode r_amalgamation) # [r_const i, r_const j, Id 1 0])), x]"
(is "... = eval r_phi [encode ?f, x]")
by (simp add: r_constn_def)
finally have "\<phi> (amalgamate i j) x = eval r_phi
[encode (Cn 1 (r_universal 3)
(r_const (encode r_amalgamation) # [r_const i, r_const j, Id 1 0])), x]" .
then have "\<phi> (amalgamate i j) x = eval (r_universal 3) [encode r_amalgamation, i, j, x]"
unfolding r_phi_def using r_universal[of ?f 1] r_amalgamation_recfn by simp
then show "\<phi> (amalgamate i j) x = amalgamation i j x"
using r_amalgamation by (simp add: r_amalgamation_recfn r_universal)
qed
lemma amalgamation_in_P1: "amalgamation i j \<in> \<P>"
using amalgamate by (metis P2_proj_P1 phi_in_P2)
lemma amalgamation_V01_R1:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "amalgamation (the (f 0)) (the (f 1)) \<in> \<R>"
using assms amalgamation_V01_total amalgamation_in_P1
by (simp add: P1_total_imp_R1)
definition "r_amalgamate \<equiv>
Cn 2 (r_smn 1 2) [r_dummy 1 (r_const (encode r_amalgamation)), Id 2 0, Id 2 1]"
lemma r_amalgamate_recfn: "recfn 2 r_amalgamate"
unfolding r_amalgamate_def by simp
lemma r_amalgamate: "eval r_amalgamate [i, j] \<down>= amalgamate i j"
proof -
let ?p = "encode r_amalgamation"
have rs21: "eval (r_smn 1 2) [?p, i, j] \<down>= smn 1 ?p [i, j]"
using r_smn by simp
moreover have "eval r_amalgamate [i, j] = eval (r_smn 1 2) [?p, i, j]"
unfolding r_amalgamate_def by auto
ultimately have "eval r_amalgamate [i, j] \<down>= smn 1 ?p [i, j]"
by simp
then show ?thesis using amalgamate_def by simp
qed
text \<open>The strategy $S$ distinguishes the two cases from the proof
sketch with the help of the next function, which checks if a hypothesis
$\varphi_i$ is inconsistent with a prefix $e$. If so, it returns the least $x
< |e|$ witnessing the inconsistency; otherwise it returns the length $|e|$.
If $\varphi_i$ diverges for some $x < |e|$, so does the function.\<close>
definition inconsist :: partial2 where
"inconsist i e \<equiv>
(if \<exists>x<e_length e. \<phi> i x \<up> then None
else if \<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x
then Some (LEAST x. x < e_length e \<and> \<phi> i x \<down>\<noteq> e_nth e x)
else Some (e_length e))"
lemma inconsist_converg:
assumes "inconsist i e \<down>"
shows "inconsist i e =
(if \<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x
then Some (LEAST x. x < e_length e \<and> \<phi> i x \<down>\<noteq> e_nth e x)
else Some (e_length e))"
and "\<forall>x<e_length e. \<phi> i x \<down>"
using inconsist_def assms by (presburger, meson)
lemma inconsist_bounded:
assumes "inconsist i e \<down>"
shows "the (inconsist i e) \<le> e_length e"
proof (cases "\<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x")
case True
then show ?thesis
using inconsist_converg[OF assms]
by (smt Least_le dual_order.strict_implies_order dual_order.strict_trans2 option.sel)
next
case False
then show ?thesis using inconsist_converg[OF assms] by auto
qed
lemma inconsist_consistent:
assumes "inconsist i e \<down>"
shows "inconsist i e \<down>= e_length e \<longleftrightarrow> (\<forall>x<e_length e. \<phi> i x \<down>= e_nth e x)"
proof
show "\<forall>x<e_length e. \<phi> i x \<down>= e_nth e x" if "inconsist i e \<down>= e_length e"
proof (cases "\<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x")
case True
then show ?thesis
using that inconsist_converg[OF assms]
by (metis (mono_tags, lifting) not_less_Least option.inject)
next
case False
then show ?thesis
using that inconsist_converg[OF assms] by simp
qed
show "\<forall>x<e_length e. \<phi> i x \<down>= e_nth e x \<Longrightarrow> inconsist i e \<down>= e_length e"
unfolding inconsist_def using assms by auto
qed
lemma inconsist_converg_eq:
assumes "inconsist i e \<down>= e_length e"
shows "\<forall>x<e_length e. \<phi> i x \<down>= e_nth e x"
using assms inconsist_consistent by auto
lemma inconsist_converg_less:
assumes "inconsist i e \<down>" and "the (inconsist i e) < e_length e"
shows "\<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x"
and "inconsist i e \<down>= (LEAST x. x < e_length e \<and> \<phi> i x \<down>\<noteq> e_nth e x)"
proof -
show "\<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x"
using assms by (metis (no_types, lifting) inconsist_converg(1) nat_neq_iff option.sel)
then show "inconsist i e \<down>= (LEAST x. x < e_length e \<and> \<phi> i x \<down>\<noteq> e_nth e x)"
using assms inconsist_converg by presburger
qed
lemma least_bounded_Suc:
assumes "\<exists>x. x < upper \<and> P x"
shows "(LEAST x. x < upper \<and> P x) = (LEAST x. x < Suc upper \<and> P x)"
proof -
let ?Q = "\<lambda>x. x < upper \<and> P x"
let ?x = "Least ?Q"
from assms have "?x < upper \<and> P ?x"
using LeastI_ex[of ?Q] by simp
then have 1: "?x < Suc upper \<and> P ?x" by simp
from assms have 2: "\<forall>y<?x. \<not> P y"
using Least_le[of ?Q] not_less_Least by fastforce
have "(LEAST x. x < Suc upper \<and> P x) = ?x"
proof (rule Least_equality)
show "?x < Suc upper \<and> P ?x" using 1 2 by blast
show "\<And>y. y < Suc upper \<and> P y \<Longrightarrow> ?x \<le> y"
using 1 2 leI by blast
qed
then show ?thesis ..
qed
lemma least_bounded_gr:
fixes P :: "nat \<Rightarrow> bool" and m :: nat
assumes "\<exists>x. x < upper \<and> P x"
shows "(LEAST x. x < upper \<and> P x) = (LEAST x. x < upper + m \<and> P x)"
proof (induction m)
case 0
then show ?case by simp
next
case (Suc m)
moreover have "\<exists>x. x < upper + m \<and> P x"
using assms trans_less_add1 by blast
ultimately show ?case using least_bounded_Suc by simp
qed
lemma inconsist_init_converg_less:
assumes "f \<in> \<R>"
and "\<phi> i \<in> \<R>"
and "inconsist i (f \<triangleright> n) \<down>"
and "the (inconsist i (f \<triangleright> n)) < Suc n"
shows "inconsist i (f \<triangleright> (n + m)) = inconsist i (f \<triangleright> n)"
proof -
have phi_i_total: "\<phi> i x \<down>" for x
using assms by simp
moreover have f_nth: "f x \<down>= e_nth (f \<triangleright> n) x" if "x < Suc n" for x n
using that assms(1) by simp
ultimately have "(\<phi> i x \<noteq> f x) = (\<phi> i x \<down>\<noteq> e_nth (f \<triangleright> n) x)" if "x < Suc n" for x n
using that by simp
then have cond: "(x < Suc n \<and> \<phi> i x \<noteq> f x) =
(x < e_length (f \<triangleright> n) \<and> \<phi> i x \<down>\<noteq> e_nth (f \<triangleright> n) x)" for x n
using length_init by metis
then have
1: "\<exists>x<Suc n. \<phi> i x \<noteq> f x" and
2: "inconsist i (f \<triangleright> n) \<down>= (LEAST x. x < Suc n \<and> \<phi> i x \<noteq> f x)"
using assms(3,4) inconsist_converg_less[of i "f \<triangleright> n"] by simp_all
then have 3: "\<exists>x<Suc (n + m). \<phi> i x \<noteq> f x"
using not_add_less1 by fastforce
then have "\<exists>x<Suc (n + m). \<phi> i x \<down>\<noteq> e_nth (f \<triangleright> (n + m)) x"
using cond by blast
then have "\<exists>x<e_length (f \<triangleright> (n + m)). \<phi> i x \<down>\<noteq> e_nth (f \<triangleright> (n + m)) x"
by simp
moreover have 4: "inconsist i (f \<triangleright> (n + m)) \<down>"
using assms(2) R1_imp_total1 inconsist_def by simp
ultimately have "inconsist i (f \<triangleright> (n + m)) \<down>=
(LEAST x. x < e_length (f \<triangleright> (n + m)) \<and> \<phi> i x \<down>\<noteq> e_nth (f \<triangleright> (n + m)) x)"
using inconsist_converg[OF 4] by simp
then have 5: "inconsist i (f \<triangleright> (n + m)) \<down>= (LEAST x. x < Suc (n + m) \<and> \<phi> i x \<noteq> f x)"
using cond[of _ "n + m"] by simp
then have "(LEAST x. x < Suc n \<and> \<phi> i x \<noteq> f x) =
(LEAST x. x < Suc n + m \<and> \<phi> i x \<noteq> f x)"
using least_bounded_gr[where ?upper="Suc n"] 1 3 by simp
then show ?thesis using 2 5 by simp
qed
definition "r_inconsist \<equiv>
let
f = Cn 2 r_length [Id 2 1];
g = Cn 4 r_ifless
[Id 4 1,
Cn 4 r_length [Id 4 3],
Id 4 1,
Cn 4 r_ifeq
[Cn 4 r_phi [Id 4 2, Id 4 0],
Cn 4 r_nth [Id 4 3, Id 4 0],
Id 4 1,
Id 4 0]]
in Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]"
lemma r_inconsist_recfn: "recfn 2 r_inconsist"
unfolding r_inconsist_def by simp
lemma r_inconsist: "eval r_inconsist [i, e] = inconsist i e"
proof -
define f where "f = Cn 2 r_length [Id 2 1]"
define len where "len = Cn 4 r_length [Id 4 3]"
define nth where "nth = Cn 4 r_nth [Id 4 3, Id 4 0]"
define ph where "ph = Cn 4 r_phi [Id 4 2, Id 4 0]"
define g where
"g = Cn 4 r_ifless [Id 4 1, len, Id 4 1, Cn 4 r_ifeq [ph, nth, Id 4 1, Id 4 0]]"
have "recfn 2 f"
unfolding f_def by simp
have f: "eval f [i, e] \<down>= e_length e"
unfolding f_def by simp
have "recfn 4 len"
unfolding len_def by simp
have len: "eval len [j, v, i, e] \<down>= e_length e" for j v
unfolding len_def by simp
have "recfn 4 nth"
unfolding nth_def by simp
have nth: "eval nth [j, v, i, e] \<down>= e_nth e j" for j v
unfolding nth_def by simp
have "recfn 4 ph"
unfolding ph_def by simp
have ph: "eval ph [j, v, i, e] = \<phi> i j" for j v
unfolding ph_def using phi_def by simp
have "recfn 4 g"
- unfolding g_def using `recfn 4 nth` `recfn 4 ph` `recfn 4 len` by simp
+ unfolding g_def using \<open>recfn 4 nth\<close> \<open>recfn 4 ph\<close> \<open>recfn 4 len\<close> by simp
have g_diverg: "eval g [j, v, i, e] \<up>" if "eval ph [j, v, i, e] \<up>" for j v
- unfolding g_def using that `recfn 4 nth` `recfn 4 ph` `recfn 4 len` by simp
+ unfolding g_def using that \<open>recfn 4 nth\<close> \<open>recfn 4 ph\<close> \<open>recfn 4 len\<close> by simp
have g_converg: "eval g [j, v, i, e] \<down>=
(if v < e_length e then v else if \<phi> i j \<down>= e_nth e j then v else j)"
if "eval ph [j, v, i, e] \<down>" for j v
- unfolding g_def using that `recfn 4 nth` `recfn 4 ph` `recfn 4 len` len nth ph
+ unfolding g_def using that \<open>recfn 4 nth\<close> \<open>recfn 4 ph\<close> \<open>recfn 4 len\<close> len nth ph
by auto
define h where "h \<equiv> Pr 2 f g"
then have "recfn 3 h"
by (simp add: \<open>recfn 2 f\<close> \<open>recfn 4 g\<close>)
let ?invariant = "\<lambda>j i e.
(if \<exists>x<j. \<phi> i x \<up> then None
else if \<exists>x<j. \<phi> i x \<down>\<noteq> e_nth e x
then Some (LEAST x. x < j \<and> \<phi> i x \<down>\<noteq> e_nth e x)
else Some (e_length e))"
have "eval h [j, i, e] = ?invariant j i e" if "j \<le> e_length e" for j
using that
proof (induction j)
case 0
- then show ?case unfolding h_def using `recfn 2 f` f `recfn 4 g` by simp
+ then show ?case unfolding h_def using \<open>recfn 2 f\<close> f \<open>recfn 4 g\<close> by simp
next
case (Suc j)
then have j_less: "j < e_length e" by simp
then have j_le: "j \<le> e_length e" by simp
show ?case
proof (cases "eval h [j, i, e] \<up>")
case True
then have "\<exists>x<j. \<phi> i x \<up>"
using j_le Suc.IH by (metis option.simps(3))
then have "\<exists>x<Suc j. \<phi> i x \<up>"
using less_SucI by blast
moreover have h: "eval h [Suc j, i, e] \<up>"
- using True h_def `recfn 3 h` by simp
+ using True h_def \<open>recfn 3 h\<close> by simp
ultimately show ?thesis by simp
next
case False
with Suc.IH j_le have h_j: "eval h [j, i, e] =
(if \<exists>x<j. \<phi> i x \<down>\<noteq> e_nth e x
then Some (LEAST x. x < j \<and> \<phi> i x \<down>\<noteq> e_nth e x)
else Some (e_length e))"
by presburger
then have the_h_j: "the (eval h [j, i, e]) =
(if \<exists>x<j. \<phi> i x \<down>\<noteq> e_nth e x
then LEAST x. x < j \<and> \<phi> i x \<down>\<noteq> e_nth e x
else e_length e)"
(is "_ = ?v")
by auto
have h_Suc: "eval h [Suc j, i, e] = eval g [j, the (eval h [j, i, e]), i, e]"
- using False h_def `recfn 4 g` `recfn 2 f` by auto
+ using False h_def \<open>recfn 4 g\<close> \<open>recfn 2 f\<close> by auto
show ?thesis
proof (cases "\<phi> i j \<up>")
case True
with ph g_diverg h_Suc show ?thesis by auto
next
case False
with h_Suc have "eval h [Suc j, i, e] \<down>=
(if ?v < e_length e then ?v
else if \<phi> i j \<down>= e_nth e j then ?v else j)"
(is "_ \<down>= ?lhs")
using g_converg ph the_h_j by simp
moreover have "?invariant (Suc j) i e \<down>=
(if \<exists>x<Suc j. \<phi> i x \<down>\<noteq> e_nth e x
then LEAST x. x < Suc j \<and> \<phi> i x \<down>\<noteq> e_nth e x
else e_length e)"
(is "_ \<down>= ?rhs")
proof -
from False have "\<phi> i j \<down>" by simp
moreover have "\<not> (\<exists>x<j. \<phi> i x \<up>)"
by (metis (no_types, lifting) Suc.IH h_j j_le option.simps(3))
ultimately have "\<not> (\<exists>x<Suc j. \<phi> i x \<up>)"
using less_Suc_eq by auto
then show ?thesis by auto
qed
moreover have "?lhs = ?rhs"
proof (cases "?v < e_length e")
case True
then have
ex_j: "\<exists>x<j. \<phi> i x \<down>\<noteq> e_nth e x" and
v_eq: "?v = (LEAST x. x < j \<and> \<phi> i x \<down>\<noteq> e_nth e x)"
by presburger+
with True have "?lhs = ?v" by simp
from ex_j have "\<exists>x<Suc j. \<phi> i x \<down>\<noteq> e_nth e x"
using less_SucI by blast
then have "?rhs = (LEAST x. x < Suc j \<and> \<phi> i x \<down>\<noteq> e_nth e x)" by simp
with True v_eq ex_j show ?thesis
using least_bounded_Suc[of j "\<lambda>x. \<phi> i x \<down>\<noteq> e_nth e x"] by simp
next
case False
then have not_ex: "\<not> (\<exists>x<j. \<phi> i x \<down>\<noteq> e_nth e x)"
using Least_le[of "\<lambda>x. x < j \<and> \<phi> i x \<down>\<noteq> e_nth e x"] j_le
by (smt leD le_less_linear le_trans)
then have "?v = e_length e" by argo
with False have lhs: "?lhs = (if \<phi> i j \<down>= e_nth e j then e_length e else j)"
by simp
show ?thesis
proof (cases "\<phi> i j \<down>= e_nth e j")
case True
then have "\<not> (\<exists>x<Suc j. \<phi> i x \<down>\<noteq> e_nth e x)"
using less_SucE not_ex by blast
then have "?rhs = e_length e" by argo
moreover from True have "?lhs = e_length e"
using lhs by simp
ultimately show ?thesis by simp
next case False
then have "\<phi> i j \<down>\<noteq> e_nth e j"
- using `\<phi> i j \<down>` by simp
+ using \<open>\<phi> i j \<down>\<close> by simp
with not_ex have "(LEAST x. x<Suc j \<and> \<phi> i x \<down>\<noteq> e_nth e x) = j"
using LeastI[of "\<lambda>x. x<Suc j \<and> \<phi> i x \<down>\<noteq> e_nth e x" j] less_Suc_eq
by blast
then have "?rhs = j"
using \<open>\<phi> i j \<down>\<noteq> e_nth e j\<close> by (meson lessI)
moreover from False lhs have "?lhs = j" by simp
ultimately show ?thesis by simp
qed
qed
ultimately show ?thesis by simp
qed
qed
qed
then have "eval h [e_length e, i, e] = ?invariant (e_length e) i e"
by auto
then have "eval h [e_length e, i, e] = inconsist i e"
using inconsist_def by simp
moreover have "eval (Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]) [i, e] =
eval h [e_length e, i, e]"
- using `recfn 4 g` `recfn 2 f` h_def by auto
+ using \<open>recfn 4 g\<close> \<open>recfn 2 f\<close> h_def by auto
ultimately show ?thesis
unfolding r_inconsist_def by (simp add: f_def g_def len_def nth_def ph_def)
qed
lemma inconsist_for_total:
assumes "total1 (\<phi> i)"
shows "inconsist i e \<down>=
(if \<exists>x<e_length e. \<phi> i x \<down>\<noteq> e_nth e x
then LEAST x. x < e_length e \<and> \<phi> i x \<down>\<noteq> e_nth e x
else e_length e)"
unfolding inconsist_def using assms total1_def by (auto; blast)
lemma inconsist_for_V01:
assumes "f \<in> V\<^sub>0\<^sub>1" and "k = amalgamate (the (f 0)) (the (f 1))"
shows "inconsist k e \<down>=
(if \<exists>x<e_length e. \<phi> k x \<down>\<noteq> e_nth e x
then LEAST x. x < e_length e \<and> \<phi> k x \<down>\<noteq> e_nth e x
else e_length e)"
proof -
have "\<phi> k \<in> \<R>"
using amalgamation_V01_R1[OF assms(1)] assms(2) amalgamate by simp
then have "total1 (\<phi> k)" by simp
with inconsist_for_total[of k] show ?thesis by simp
qed
text \<open>The next function computes Gödel numbers of functions consistent
with a given prefix. The strategy will use these as consistent auxiliary
hypotheses when receiving a prefix of length one.\<close>
definition "r_auxhyp \<equiv> Cn 1 (r_smn 1 1) [r_const (encode r_prenum), Id 1 0]"
lemma r_auxhyp_prim: "prim_recfn 1 r_auxhyp"
unfolding r_auxhyp_def by simp
lemma r_auxhyp: "\<phi> (the (eval r_auxhyp [e])) = prenum e"
proof
fix x
let ?p = "encode r_prenum"
let ?p = "encode r_prenum"
have "eval r_auxhyp [e] = eval (r_smn 1 1) [?p, e]"
unfolding r_auxhyp_def by simp
then have "eval r_auxhyp [e] \<down>= smn 1 ?p [e]"
by (simp add: r_smn)
also have "... \<down>= encode (Cn 1 (r_universal (1 + length [e]))
(r_constn (1 - 1) ?p #
map (r_constn (1 - 1)) [e] @ map (recf.Id 1) [0..<1]))"
using smn[of 1 ?p "[e]"] by simp
also have "... \<down>= encode (Cn 1 (r_universal (1 + 1))
(r_constn 0 ?p # map (r_constn 0) [e] @ [Id 1 0]))"
by simp
also have "... \<down>= encode (Cn 1 (r_universal 2)
(r_constn 0 ?p # map (r_constn 0) [e] @ [Id 1 0]))"
by (metis one_add_one)
also have "... \<down>= encode (Cn 1 (r_universal 2) [r_constn 0 ?p, r_constn 0 e, Id 1 0])"
by simp
also have "... \<down>= encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0])"
using r_constn_def by simp
finally have "eval r_auxhyp [e] \<down>=
encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0])" .
moreover have "\<phi> (the (eval r_auxhyp [e])) x = eval r_phi [the (eval r_auxhyp [e]), x]"
by (simp add: phi_def)
ultimately have "\<phi> (the (eval r_auxhyp [e])) x =
eval r_phi [encode (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0]), x]"
(is "_ = eval r_phi [encode ?f, x]")
by simp
then have "\<phi> (the (eval r_auxhyp [e])) x =
eval (Cn 1 (r_universal 2) [r_const ?p, r_const e, Id 1 0]) [x]"
using r_phi_def r_universal[of ?f 1 "[x]"] by simp
then have "\<phi> (the (eval r_auxhyp [e])) x = eval (r_universal 2) [?p, e, x]"
by simp
then have "\<phi> (the (eval r_auxhyp [e])) x = eval r_prenum [e, x]"
using r_universal by simp
then show "\<phi> (the (eval r_auxhyp [e])) x = prenum e x" by simp
qed
definition auxhyp :: partial1 where
"auxhyp e \<equiv> eval r_auxhyp [e]"
lemma auxhyp_prenum: "\<phi> (the (auxhyp e)) = prenum e"
using auxhyp_def r_auxhyp by metis
lemma auxhyp_in_R1: "auxhyp \<in> \<R>"
using auxhyp_def Mn_free_imp_total R1I r_auxhyp_prim by metis
text \<open>Now we can define our consistent learning strategy for @{term "V\<^sub>0\<^sub>1"}.\<close>
definition "r_sv01 \<equiv>
let
at0 = Cn 1 r_nth [Id 1 0, Z];
at1 = Cn 1 r_nth [Id 1 0, r_const 1];
m = Cn 1 r_amalgamate [at0, at1];
c = Cn 1 r_inconsist [m, Id 1 0];
p = Cn 1 r_pdec1 [Cn 1 r_parallel [at0, at1, c]];
g = Cn 1 r_ifeq [c, r_length, m, Cn 1 r_ifz [p, at1, at0]]
in Cn 1 (r_lifz r_auxhyp g) [Cn 1 r_eq [r_length, r_const 1], Id 1 0]"
lemma r_sv01_recfn: "recfn 1 r_sv01"
unfolding r_sv01_def using r_auxhyp_prim r_inconsist_recfn r_amalgamate_recfn
by (simp add: Let_def)
definition sv01 :: partial1 ("s\<^bsub>01\<^esub>")where
"sv01 e \<equiv> eval r_sv01 [e]"
lemma sv01_in_P1: "s\<^bsub>01\<^esub> \<in> \<P>"
using sv01_def r_sv01_recfn P1I by presburger
text \<open>We are interested in the behavior of @{term "s\<^bsub>01\<^esub>"} only on
prefixes of functions in @{term "V\<^sub>0\<^sub>1"}. This behavior is linked
to the amalgamation of $f(0)$ and $f(1)$, where $f$ is the function
to be learned.\<close>
abbreviation amalg01 :: "partial1 \<Rightarrow> nat" where
"amalg01 f \<equiv> amalgamate (the (f 0)) (the (f 1))"
lemma sv01:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "s\<^bsub>01\<^esub> (f \<triangleright> 0) = auxhyp (f \<triangleright> 0)"
and "n \<noteq> 0 \<Longrightarrow>
inconsist (amalg01 f) (f \<triangleright> n) \<down>= Suc n \<Longrightarrow>
s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= amalg01 f"
and "n \<noteq> 0 \<Longrightarrow>
the (inconsist (amalg01 f) (f \<triangleright> n)) < Suc n \<Longrightarrow>
pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist (amalg01 f) (f \<triangleright> n))))) = 0 \<Longrightarrow>
s\<^bsub>01\<^esub> (f \<triangleright> n) = f 1"
and "n \<noteq> 0 \<Longrightarrow>
the (inconsist (amalg01 f) (f \<triangleright> n)) < Suc n \<Longrightarrow>
pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist (amalg01 f) (f \<triangleright> n))))) \<noteq> 0 \<Longrightarrow>
s\<^bsub>01\<^esub> (f \<triangleright> n) = f 0"
proof -
have f_total: "\<And>x. f x \<down>"
using assms V01_def R1_imp_total1 by blast
define at0 where "at0 = Cn 1 r_nth [Id 1 0, Z]"
define at1 where "at1 = Cn 1 r_nth [Id 1 0, r_const 1]"
define m where "m = Cn 1 r_amalgamate [at0, at1]"
define c where "c = Cn 1 r_inconsist [m, Id 1 0]"
define p where "p = Cn 1 r_pdec1 [Cn 1 r_parallel [at0, at1, c]]"
define g where "g = Cn 1 r_ifeq [c, r_length, m, Cn 1 r_ifz [p, at1, at0]]"
have "recfn 1 g"
unfolding g_def p_def c_def m_def at1_def at0_def
using r_auxhyp_prim r_inconsist_recfn r_amalgamate_recfn
by simp
have "eval (Cn 1 r_eq [r_length, r_const 1]) [f \<triangleright> 0] \<down>= 0"
by simp
then have "eval r_sv01 [f \<triangleright> 0] = eval r_auxhyp [f \<triangleright> 0]"
- unfolding r_sv01_def using `recfn 1 g` c_def g_def m_def p_def r_auxhyp_prim
+ unfolding r_sv01_def using \<open>recfn 1 g\<close> c_def g_def m_def p_def r_auxhyp_prim
by (auto simp add: Let_def)
then show "s\<^bsub>01\<^esub> (f \<triangleright> 0) = auxhyp (f \<triangleright> 0)"
by (simp add: auxhyp_def sv01_def)
have sv01: "s\<^bsub>01\<^esub> (f \<triangleright> n) = eval g [f \<triangleright> n]" if "n \<noteq> 0"
proof -
have *: "eval (Cn 1 r_eq [r_length, r_const 1]) [f \<triangleright> n] \<down>\<noteq> 0"
(is "?r_eq \<down>\<noteq> 0")
using that by simp
moreover have "recfn 2 (r_lifz r_auxhyp g)"
- using `recfn 1 g` r_auxhyp_prim by simp
+ using \<open>recfn 1 g\<close> r_auxhyp_prim by simp
moreover have "eval r_sv01 [f \<triangleright> n] =
eval (Cn 1 (r_lifz r_auxhyp g) [Cn 1 r_eq [r_length, r_const 1], Id 1 0]) [f \<triangleright> n]"
using r_sv01_def by (metis at0_def at1_def c_def g_def m_def p_def)
ultimately have "eval r_sv01 [f \<triangleright> n] = eval (r_lifz r_auxhyp g) [the ?r_eq, f \<triangleright> n]"
by simp
then have "eval r_sv01 [f \<triangleright> n] = eval g [f \<triangleright> n]"
using "*" \<open>recfn 1 g\<close> r_auxhyp_prim by auto
then show ?thesis by (simp add: sv01_def that)
qed
have "recfn 1 at0"
unfolding at0_def by simp
have at0: "eval at0 [f \<triangleright> n] \<down>= the (f 0)"
unfolding at0_def by simp
have "recfn 1 at1"
unfolding at1_def by simp
have at1: "n \<noteq> 0 \<Longrightarrow> eval at1 [f \<triangleright> n] \<down>= the (f 1)"
unfolding at1_def by simp
have "recfn 1 m"
unfolding m_def at0_def at1_def using r_amalgamate_recfn by simp
have m: "n \<noteq> 0 \<Longrightarrow> eval m [f \<triangleright> n] \<down>= amalg01 f"
(is "_ \<Longrightarrow> _ \<down>= ?m")
unfolding m_def at0_def at1_def
using at0 at1 amalgamate r_amalgamate r_amalgamate_recfn by simp
then have c: "n \<noteq> 0 \<Longrightarrow> eval c [f \<triangleright> n] = inconsist (amalg01 f) (f \<triangleright> n)"
(is "_ \<Longrightarrow> _ = ?c")
- unfolding c_def using r_inconsist_recfn `recfn 1 m` r_inconsist by auto
+ unfolding c_def using r_inconsist_recfn \<open>recfn 1 m\<close> r_inconsist by auto
then have c_converg: "n \<noteq> 0 \<Longrightarrow> eval c [f \<triangleright> n] \<down>"
using inconsist_for_V01[OF assms] by simp
have "recfn 1 c"
- unfolding c_def using `recfn 1 m` r_inconsist_recfn by simp
+ unfolding c_def using \<open>recfn 1 m\<close> r_inconsist_recfn by simp
have par: "n \<noteq> 0 \<Longrightarrow>
eval (Cn 1 r_parallel [at0, at1, c]) [f \<triangleright> n] = parallel (the (f 0)) (the (f 1)) (the ?c)"
(is "_ \<Longrightarrow> _ = ?par")
- using at0 at1 c c_converg m r_parallel' `recfn 1 at0` `recfn 1 at1` `recfn 1 c`
+ using at0 at1 c c_converg m r_parallel' \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 c\<close>
by simp
with parallel_converg_V01[OF assms] have
par_converg: "n \<noteq> 0 \<Longrightarrow> eval (Cn 1 r_parallel [at0, at1, c]) [f \<triangleright> n] \<down>"
by simp
then have p_converg: "n \<noteq> 0 \<Longrightarrow> eval p [f \<triangleright> n] \<down>"
- unfolding p_def using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c`
+ unfolding p_def using at0 at1 c_converg \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 c\<close>
by simp
have p: "n \<noteq> 0 \<Longrightarrow> eval p [f \<triangleright> n] \<down>= pdec1 (the ?par)"
unfolding p_def
- using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c` par par_converg
+ using at0 at1 c_converg \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 c\<close> par par_converg
by simp
have "recfn 1 p"
- unfolding p_def using `recfn 1 at0` `recfn 1 at1` `recfn 1 m` `recfn 1 c`
+ unfolding p_def using \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 m\<close> \<open>recfn 1 c\<close>
by simp
let ?r = "Cn 1 r_ifz [p, at1, at0]"
have r: "n \<noteq> 0 \<Longrightarrow> eval ?r [f \<triangleright> n] = (if pdec1 (the ?par) = 0 then f 1 else f 0)"
- using at0 at1 c_converg `recfn 1 at0` `recfn 1 at1` `recfn 1 c`
- `recfn 1 m` `recfn 1 p` p f_total
+ using at0 at1 c_converg \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 c\<close>
+ \<open>recfn 1 m\<close> \<open>recfn 1 p\<close> p f_total
by fastforce
have g: "n \<noteq> 0 \<Longrightarrow>
eval g [f \<triangleright> n] \<down>=
(if the ?c = e_length (f \<triangleright> n)
then ?m else the (eval (Cn 1 r_ifz [p, at1, at0]) [f \<triangleright> n]))"
unfolding g_def
- using `recfn 1 p` `recfn 1 at0` `recfn 1 at1` `recfn 1 c` `recfn 1 m`
+ using \<open>recfn 1 p\<close> \<open>recfn 1 at0\<close> \<open>recfn 1 at1\<close> \<open>recfn 1 c\<close> \<open>recfn 1 m\<close>
p_converg at1 at0 c c_converg m
by simp
{
assume "n \<noteq> 0" and "?c \<down>= Suc n"
moreover have "e_length (f \<triangleright> n) = Suc n" by simp
ultimately have "eval g [f \<triangleright> n] \<down>= ?m" using g by simp
then show "s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= amalg01 f"
- using sv01[OF `n \<noteq> 0`] by simp
+ using sv01[OF \<open>n \<noteq> 0\<close>] by simp
next
assume "n \<noteq> 0" and "the ?c < Suc n" and "pdec1 (the ?par) = 0"
with g r f_total have "eval g [f \<triangleright> n] = f 1" by simp
then show "s\<^bsub>01\<^esub> (f \<triangleright> n) = f 1"
- using sv01[OF `n \<noteq> 0`] by simp
+ using sv01[OF \<open>n \<noteq> 0\<close>] by simp
next
assume "n \<noteq> 0" and "the ?c < Suc n" and "pdec1 (the ?par) \<noteq> 0"
with g r f_total have "eval g [f \<triangleright> n] = f 0" by simp
then show "s\<^bsub>01\<^esub> (f \<triangleright> n) = f 0"
- using sv01[OF `n \<noteq> 0`] by simp
+ using sv01[OF \<open>n \<noteq> 0\<close>] by simp
}
qed
text \<open>Part of the correctness of @{term sv01} is convergence on
prefixes of functions in @{term "V\<^sub>0\<^sub>1"}.\<close>
lemma sv01_converg_V01:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>"
proof (cases "n = 0")
case True
then show ?thesis
using assms sv01 R1_imp_total1 auxhyp_in_R1 by simp
next
case n_gr_0: False
show ?thesis
proof (cases "inconsist (amalg01 f) (f \<triangleright> n) \<down>= Suc n")
case True
then show ?thesis
using n_gr_0 assms sv01 by simp
next
case False
then have "the (inconsist (amalg01 f) (f \<triangleright> n)) < Suc n"
using assms inconsist_bounded inconsist_for_V01 length_init
by (metis (no_types, lifting) le_neq_implies_less option.collapse option.simps(3))
then show ?thesis
using n_gr_0 assms sv01 R1_imp_total1 total1E V01_def
by (metis (no_types, lifting) mem_Collect_eq)
qed
qed
text \<open>Another part of the correctness of @{term sv01} is its hypotheses
being consistent on prefixes of functions in @{term "V\<^sub>0\<^sub>1"}.\<close>
lemma sv01_consistent_V01:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "\<forall>x\<le>n. \<phi> (the (s\<^bsub>01\<^esub> (f \<triangleright> n))) x = f x"
proof (cases "n = 0")
case True
then have "s\<^bsub>01\<^esub> (f \<triangleright> n) = auxhyp (f \<triangleright> n)"
using sv01[OF assms] by simp
then have "\<phi> (the (s\<^bsub>01\<^esub> (f \<triangleright> n))) = prenum (f \<triangleright> n)"
using auxhyp_prenum by simp
then show ?thesis
using R1_imp_total1 total1E assms by (simp add: V01_def)
next
case n_gr_0: False
let ?m = "amalg01 f"
let ?e = "f \<triangleright> n"
let ?c = "the (inconsist ?m ?e)"
have c: "inconsist ?m ?e \<down>"
using assms inconsist_for_V01 by blast
show ?thesis
proof (cases "inconsist ?m ?e \<down>= Suc n")
case True
then show ?thesis
using assms n_gr_0 sv01 R1_imp_total1 total1E V01_def is_init_of_def
inconsist_consistent not_initial_imp_not_eq length_init inconsist_converg_eq
by (metis (no_types, lifting) le_imp_less_Suc mem_Collect_eq option.sel)
next
case False
then have less: "the (inconsist ?m ?e) < Suc n"
using c assms inconsist_bounded inconsist_for_V01 length_init
by (metis le_neq_implies_less option.collapse)
then have "the (inconsist ?m ?e) < e_length ?e"
by auto
then have
"\<exists>x<e_length ?e. \<phi> ?m x \<down>\<noteq> e_nth ?e x"
"inconsist ?m ?e \<down>= (LEAST x. x < e_length ?e \<and> \<phi> ?m x \<down>\<noteq> e_nth ?e x)"
(is "_ \<down>= Least ?P")
using inconsist_converg_less[OF c] by simp_all
then have "?P ?c" and "\<And>x. x < ?c \<Longrightarrow> \<not> ?P x"
using LeastI_ex[of ?P] not_less_Least[of _ ?P] by (auto simp del: e_nth)
then have "\<phi> ?m ?c \<noteq> f ?c" by auto
then have "amalgamation (the (f 0)) (the (f 1)) ?c \<noteq> f ?c"
using amalgamate by simp
then have *: "Some (pdec2 (the (parallel (the (f 0)) (the (f 1)) ?c))) \<noteq> f ?c"
using amalgamation_def by (metis assms parallel_converg_V01)
let ?p = "parallel (the (f 0)) (the (f 1)) ?c"
show ?thesis
proof (cases "pdec1 (the ?p) = 0")
case True
then have "\<phi> (the (f 0)) ?c \<down>= pdec2 (the ?p)"
using assms parallel_0 parallel_converg_V01
by (metis option.collapse prod.collapse prod_decode_inverse)
then have "\<phi> (the (f 0)) ?c \<noteq> f ?c"
using * by simp
then have "\<phi> (the (f 0)) \<noteq> f" by auto
then have "\<phi> (the (f 1)) = f"
using assms V01_def by auto
moreover have "s\<^bsub>01\<^esub> (f \<triangleright> n) = f 1"
using True less n_gr_0 sv01 assms by simp
ultimately show ?thesis by simp
next
case False
then have "pdec1 (the ?p) = 1"
by (meson assms parallel_converg_V01 parallel_converg_pdec1_0_or_1)
then have "\<phi> (the (f 1)) ?c \<down>= pdec2 (the ?p)"
using assms parallel_1 parallel_converg_V01
by (metis option.collapse prod.collapse prod_decode_inverse)
then have "\<phi> (the (f 1)) ?c \<noteq> f ?c"
using * by simp
then have "\<phi> (the (f 1)) \<noteq> f" by auto
then have "\<phi> (the (f 0)) = f"
using assms V01_def by auto
moreover from False less n_gr_0 sv01 assms have "s\<^bsub>01\<^esub> (f \<triangleright> n) = f 0"
by simp
ultimately show ?thesis by simp
qed
qed
qed
text \<open>The final part of the correctness is @{term "sv01"} converging
for all functions in @{term "V\<^sub>0\<^sub>1"}.\<close>
lemma sv01_limit_V01:
assumes "f \<in> V\<^sub>0\<^sub>1"
shows "\<exists>i. \<forall>\<^sup>\<infinity>n. s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= i"
proof (cases "\<forall>n>0. s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= amalgamate (the (f 0)) (the (f 1))")
case True
then show ?thesis by (meson less_le_trans zero_less_one)
next
case False
then obtain n\<^sub>0 where n0:
"n\<^sub>0 \<noteq> 0"
"s\<^bsub>01\<^esub> (f \<triangleright> n\<^sub>0) \<down>\<noteq> amalg01 f"
using \<open>f \<in> V\<^sub>0\<^sub>1\<close> sv01_converg_V01 by blast
then have *: "the (inconsist (amalg01 f) (f \<triangleright> n\<^sub>0)) < Suc n\<^sub>0"
(is "the (inconsist ?m (f \<triangleright> n\<^sub>0)) < Suc n\<^sub>0")
- using assms `n\<^sub>0 \<noteq> 0` sv01(2) inconsist_bounded inconsist_for_V01 length_init
+ using assms \<open>n\<^sub>0 \<noteq> 0\<close> sv01(2) inconsist_bounded inconsist_for_V01 length_init
by (metis (no_types, lifting) le_neq_implies_less option.collapse option.simps(3))
moreover have "f \<in> \<R>"
using assms V01_def by auto
moreover have "\<phi> ?m \<in> \<R>"
using amalgamate amalgamation_V01_R1 assms by auto
moreover have "inconsist ?m (f \<triangleright> n\<^sub>0) \<down>"
using inconsist_for_V01 assms by blast
ultimately have **: "inconsist ?m (f \<triangleright> (n\<^sub>0 + m)) = inconsist ?m (f \<triangleright> n\<^sub>0)" for m
using inconsist_init_converg_less[of f ?m] by simp
then have "the (inconsist ?m (f \<triangleright> (n\<^sub>0 + m))) < Suc n\<^sub>0 + m" for m
using * by auto
moreover have
"pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist ?m (f \<triangleright> (n\<^sub>0 + m)))))) =
pdec1 (the (parallel (the (f 0)) (the (f 1)) (the (inconsist ?m (f \<triangleright> n\<^sub>0)))))"
for m
using ** by auto
moreover have "n\<^sub>0 + m \<noteq> 0" for m
- using `n\<^sub>0 \<noteq> 0` by simp
+ using \<open>n\<^sub>0 \<noteq> 0\<close> by simp
ultimately have "s\<^bsub>01\<^esub> (f \<triangleright> (n\<^sub>0 + m)) = s\<^bsub>01\<^esub> (f \<triangleright> n\<^sub>0)" for m
using assms sv01 * \<open>n\<^sub>0 \<noteq> 0\<close> by (metis add_Suc)
moreover define i where "i = s\<^bsub>01\<^esub> (f \<triangleright> n\<^sub>0)"
ultimately have "\<forall>n\<ge>n\<^sub>0. s\<^bsub>01\<^esub> (f \<triangleright> n) = i"
using nat_le_iff_add by auto
then have "\<forall>n\<ge>n\<^sub>0. s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= the i"
using n0(2) by simp
then show ?thesis by auto
qed
lemma V01_learn_cons: "learn_cons \<phi> V\<^sub>0\<^sub>1 s\<^bsub>01\<^esub>"
proof (rule learn_consI2)
show "environment \<phi> V\<^sub>0\<^sub>1 s\<^bsub>01\<^esub>"
by (simp add: Collect_mono V01_def phi_in_P2 sv01_in_P1 sv01_converg_V01)
show "\<And>f n. f \<in> V\<^sub>0\<^sub>1 \<Longrightarrow> \<forall>k\<le>n. \<phi> (the (s\<^bsub>01\<^esub> (f \<triangleright> n))) k = f k"
using sv01_consistent_V01 .
show "\<exists>i n\<^sub>0. \<forall>n\<ge>n\<^sub>0. s\<^bsub>01\<^esub> (f \<triangleright> n) \<down>= i" if "f \<in> V\<^sub>0\<^sub>1" for f
using sv01_limit_V01 that by simp
qed
corollary V01_in_CONS: "V\<^sub>0\<^sub>1 \<in> CONS"
using V01_learn_cons CONS_def by auto
text \<open>Now we can show the main result of this section, namely that
there is a consistently learnable class that cannot be learned consistently
by a total strategy. In other words, there is no Lemma~R for CONS.\<close>
lemma no_lemma_R_for_CONS: "\<exists>U. U \<in> CONS \<and> (\<not> (\<exists>s. s \<in> \<R> \<and> learn_cons \<phi> U s))"
using V01_in_CONS V01_not_in_R_cons by auto
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/Partial_Recursive.thy b/thys/Inductive_Inference/Partial_Recursive.thy
--- a/thys/Inductive_Inference/Partial_Recursive.thy
+++ b/thys/Inductive_Inference/Partial_Recursive.thy
@@ -1,1914 +1,1914 @@
chapter \<open>Partial recursive functions\<close>
theory Partial_Recursive
imports Main "HOL-Library.Nat_Bijection"
begin
text \<open>This chapter lays the foundation for Chapter~\ref{c:iirf}.
Essentially it develops recursion theory up to the point of certain
fixed-point theorems. This in turn requires standard results such as the
existence of a universal function and the $s$-$m$-$n$ theorem. Besides these,
the chapter contains some results, mostly regarding decidability and the
Kleene normal form, that are not strictly needed later. They are included as
relatively low-hanging fruits to round off the chapter.
The formalization of partial recursive functions is very much inspired by the
Universal Turing Machine AFP entry by Xu
et~al.~\cite{Universal_Turing_Machine-AFP}. It models partial recursive
functions as algorithms whose semantics is given by an evaluation function.
This works well for most of this chapter. For the next chapter, however, it
is beneficial to regard partial recursive functions as ``proper'' partial
functions. To that end, Section~\ref{s:alternative} introduces more
conventional and convenient notation for the common special cases of unary
and binary partial recursive functions.
Especially for the nontrivial proofs I consulted the classical textbook by
Rogers~\cite{Rogers87}, which also partially explains my preferring the
traditional term ``recursive'' to the more modern ``computable''.\<close>
section \<open>Basic definitions\<close>
subsection \<open>Partial recursive functions\<close>
text \<open>To represent partial recursive functions we start from the same
datatype as Xu et~al.~\cite{Universal_Turing_Machine-AFP}, more specifically
from Urban's version of the formalization. In fact the datatype @{text recf}
and the function @{text arity} below have been copied verbatim from it.\<close>
datatype recf =
Z
| S
| Id nat nat
| Cn nat recf "recf list"
| Pr nat recf recf
| Mn nat recf
fun arity :: "recf \<Rightarrow> nat" where
"arity Z = 1"
| "arity S = 1"
| "arity (Id m n) = m"
| "arity (Cn n f gs) = n"
| "arity (Pr n f g) = Suc n"
| "arity (Mn n f) = n"
text \<open>Already we deviate from Xu et~al.\ in that we define a
well-formedness predicate for partial recursive functions. Well-formedness
essentially means arity constraints are respected when combining @{typ
recf}s.\<close>
fun wellf :: "recf \<Rightarrow> bool" where
"wellf Z = True"
| "wellf S = True"
| "wellf (Id m n) = (n < m)"
| "wellf (Cn n f gs) =
(n > 0 \<and> (\<forall>g \<in> set gs. arity g = n \<and> wellf g) \<and> arity f = length gs \<and> wellf f)"
| "wellf (Pr n f g) =
(arity g = Suc (Suc n) \<and> arity f = n \<and> wellf f \<and> wellf g)"
| "wellf (Mn n f) = (n > 0 \<and> arity f = Suc n \<and> wellf f)"
lemma wellf_arity_nonzero: "wellf f \<Longrightarrow> arity f > 0"
by (induction f rule: arity.induct) simp_all
lemma wellf_Pr_arity_greater_1: "wellf (Pr n f g) \<Longrightarrow> arity (Pr n f g) > 1"
using wellf_arity_nonzero by auto
text \<open>For the most part of this chapter this is the meaning of ``$f$
is an $n$-ary partial recursive function'':\<close>
abbreviation recfn :: "nat \<Rightarrow> recf \<Rightarrow> bool" where
"recfn n f \<equiv> wellf f \<and> arity f = n"
text \<open>Some abbreviations for working with @{typ "nat option"}:\<close>
abbreviation divergent :: "nat option \<Rightarrow> bool" ("_ \<up>" [50] 50) where
"x \<up> \<equiv> x = None"
abbreviation convergent :: "nat option \<Rightarrow> bool" ("_ \<down>" [50] 50) where
"x \<down> \<equiv> x \<noteq> None"
abbreviation convergent_eq :: "nat option \<Rightarrow> nat \<Rightarrow> bool" (infix "\<down>=" 50) where
"x \<down>= y \<equiv> x = Some y"
abbreviation convergent_neq :: "nat option \<Rightarrow> nat \<Rightarrow> bool" (infix "\<down>\<noteq>" 50) where
"x \<down>\<noteq> y \<equiv> x \<down> \<and> x \<noteq> Some y"
text \<open>In prose the terms ``halt'', ``terminate'', ``converge'', and
``defined'' will be used interchangeably; likewise for ``not halt'',
``diverge'', and ``undefined''. In names of lemmas, the abbreviations @{text
converg} and @{text diverg} will be used consistently.\<close>
text \<open>Our second major deviation from Xu
et~al.~\cite{Universal_Turing_Machine-AFP} is that we model the semantics of
a @{typ recf} by combining the value and the termination of a function into
one evaluation function with codomain @{typ "nat option"}, rather than
separating both aspects into an evaluation function with codomain @{typ nat}
and a termination predicate.
The value of a well-formed partial recursive function applied to a
correctly-sized list of arguments:\<close>
fun eval_wellf :: "recf \<Rightarrow> nat list \<Rightarrow> nat option" where
"eval_wellf Z xs \<down>= 0"
| "eval_wellf S xs \<down>= Suc (xs ! 0)"
| "eval_wellf (Id m n) xs \<down>= xs ! n"
| "eval_wellf (Cn n f gs) xs =
(if \<forall>g \<in> set gs. eval_wellf g xs \<down>
then eval_wellf f (map (\<lambda>g. the (eval_wellf g xs)) gs)
else None)"
| "eval_wellf (Pr n f g) [] = undefined"
| "eval_wellf (Pr n f g) (0 # xs) = eval_wellf f xs"
| "eval_wellf (Pr n f g) (Suc x # xs) =
Option.bind (eval_wellf (Pr n f g) (x # xs)) (\<lambda>v. eval_wellf g (x # v # xs))"
| "eval_wellf (Mn n f) xs =
(let E = \<lambda>z. eval_wellf f (z # xs)
in if \<exists>z. E z \<down>= 0 \<and> (\<forall>y<z. E y \<down>)
then Some (LEAST z. E z \<down>= 0 \<and> (\<forall>y<z. E y \<down>))
else None)"
text \<open>We define a function value only if the @{typ recf} is well-formed
and its arity matches the number of arguments.\<close>
definition eval :: "recf \<Rightarrow> nat list \<Rightarrow> nat option" where
"recfn (length xs) f \<Longrightarrow> eval f xs \<equiv> eval_wellf f xs"
lemma eval_Z [simp]: "eval Z [x] \<down>= 0"
by (simp add: eval_def)
lemma eval_Z' [simp]: "length xs = 1 \<Longrightarrow> eval Z xs \<down>= 0"
by (simp add: eval_def)
lemma eval_S [simp]: "eval S [x] \<down>= Suc x"
by (simp add: eval_def)
lemma eval_S' [simp]: "length xs = 1 \<Longrightarrow> eval S xs \<down>= Suc (hd xs)"
using eval_def hd_conv_nth[of xs] by fastforce
lemma eval_Id [simp]:
assumes "n < m" and "m = length xs"
shows "eval (Id m n) xs \<down>= xs ! n"
using assms by (simp add: eval_def)
lemma eval_Cn [simp]:
assumes "recfn (length xs) (Cn n f gs)"
shows "eval (Cn n f gs) xs =
(if \<forall>g\<in>set gs. eval g xs \<down>
then eval f (map (\<lambda>g. the (eval g xs)) gs)
else None)"
proof -
have "eval (Cn n f gs) xs = eval_wellf (Cn n f gs) xs"
using assms eval_def by blast
moreover have "\<And>g. g \<in> set gs \<Longrightarrow> eval_wellf g xs = eval g xs"
using assms eval_def by simp
ultimately have "eval (Cn n f gs) xs =
(if \<forall>g\<in>set gs. eval g xs \<down>
then eval_wellf f (map (\<lambda>g. the (eval g xs)) gs)
else None)"
using map_eq_conv[of "\<lambda>g. the (eval_wellf g xs)" gs "\<lambda>g. the (eval g xs)"]
by (auto, metis)
moreover have "\<And>ys. length ys = length gs \<Longrightarrow> eval f ys = eval_wellf f ys"
using assms eval_def by simp
ultimately show ?thesis by auto
qed
lemma eval_Pr_0 [simp]:
assumes "recfn (Suc n) (Pr n f g)" and "n = length xs"
shows "eval (Pr n f g) (0 # xs) = eval f xs"
using assms by (simp add: eval_def)
lemma eval_Pr_diverg_Suc [simp]:
assumes "recfn (Suc n) (Pr n f g)"
and "n = length xs"
and "eval (Pr n f g) (x # xs) \<up>"
shows "eval (Pr n f g) (Suc x # xs) \<up>"
using assms by (simp add: eval_def)
lemma eval_Pr_converg_Suc [simp]:
assumes "recfn (Suc n) (Pr n f g)"
and "n = length xs"
and "eval (Pr n f g) (x # xs) \<down>"
shows "eval (Pr n f g) (Suc x # xs) = eval g (x # the (eval (Pr n f g) (x # xs)) # xs)"
using assms eval_def by auto
lemma eval_Pr_diverg_add:
assumes "recfn (Suc n) (Pr n f g)"
and "n = length xs"
and "eval (Pr n f g) (x # xs) \<up>"
shows "eval (Pr n f g) ((x + y) # xs) \<up>"
using assms by (induction y) auto
lemma eval_Pr_converg_le:
assumes "recfn (Suc n) (Pr n f g)"
and "n = length xs"
and "eval (Pr n f g) (x # xs) \<down>"
and "y \<le> x"
shows "eval (Pr n f g) (y # xs) \<down>"
using assms eval_Pr_diverg_add le_Suc_ex by metis
lemma eval_Pr_Suc_converg:
assumes "recfn (Suc n) (Pr n f g)"
and "n = length xs"
and "eval (Pr n f g) (Suc x # xs) \<down>"
shows "eval g (x # (the (eval (Pr n f g) (x # xs))) # xs) \<down>"
and "eval (Pr n f g) (Suc x # xs) = eval g (x # the (eval (Pr n f g) (x # xs)) # xs)"
using eval_Pr_converg_Suc[of n f g xs x, OF assms(1,2)]
eval_Pr_converg_le[of n f g xs "Suc x" x, OF assms] assms(3)
by simp_all
lemma eval_Mn [simp]:
assumes "recfn (length xs) (Mn n f)"
shows "eval (Mn n f) xs =
(if (\<exists>z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>))
then Some (LEAST z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>))
else None)"
using assms eval_def by auto
text \<open>For $\mu$-recursion, the condition @{term "(\<forall>y<z. eval_wellf f (y # xs) \<down>)"}
inside @{text LEAST} in the definition of @{term eval_wellf} is redundant.\<close>
lemma eval_wellf_Mn:
"eval_wellf (Mn n f) xs =
(if (\<exists>z. eval_wellf f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval_wellf f (y # xs) \<down>))
then Some (LEAST z. eval_wellf f (z # xs) \<down>= 0)
else None)"
proof -
let ?P = "\<lambda>z. eval_wellf f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval_wellf f (y # xs) \<down>)"
{
assume "\<exists>z. ?P z"
moreover define z where "z = Least ?P"
ultimately have "?P z"
using LeastI[of ?P] by blast
have "(LEAST z. eval_wellf f (z # xs) \<down>= 0) = z"
proof (rule Least_equality)
show "eval_wellf f (z # xs) \<down>= 0"
- using `?P z` by simp
+ using \<open>?P z\<close> by simp
show "z \<le> y" if "eval_wellf f (y # xs) \<down>= 0" for y
proof (rule ccontr)
assume "\<not> z \<le> y"
then have "y < z" by simp
moreover from this have "?P y"
- using that `?P z` by simp
+ using that \<open>?P z\<close> by simp
ultimately show False
using that not_less_Least z_def by blast
qed
qed
}
then show ?thesis by simp
qed
lemma eval_Mn':
assumes "recfn (length xs) (Mn n f)"
shows "eval (Mn n f) xs =
(if (\<exists>z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>))
then Some (LEAST z. eval f (z # xs) \<down>= 0)
else None)"
using assms eval_def eval_wellf_Mn by auto
text \<open>Proving that $\mu$-recursion converges is easier if one does not
have to deal with @{text LEAST} directly.\<close>
lemma eval_Mn_convergI:
assumes "recfn (length xs) (Mn n f)"
and "eval f (z # xs) \<down>= 0"
and "\<And>y. y < z \<Longrightarrow> eval f (y # xs) \<down>\<noteq> 0"
shows "eval (Mn n f) xs \<down>= z"
proof -
let ?P = "\<lambda>z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>)"
have "z = Least ?P"
using Least_equality[of ?P z] assms(2,3) not_le_imp_less by blast
moreover have "?P z" using assms(2,3) by simp
ultimately show "eval (Mn n f) xs \<down>= z"
using eval_Mn[OF assms(1)] by meson
qed
text \<open>Similarly, reasoning from a $\mu$-recursive function is
simplified somewhat by the next lemma.\<close>
lemma eval_Mn_convergE:
assumes "recfn (length xs) (Mn n f)" and "eval (Mn n f) xs \<down>= z"
shows "z = (LEAST z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>))"
and "eval f (z # xs) \<down>= 0"
and "\<And>y. y < z \<Longrightarrow> eval f (y # xs) \<down>\<noteq> 0"
proof -
let ?P = "\<lambda>z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>)"
show "z = Least ?P"
using assms eval_Mn[OF assms(1)]
by (metis (no_types, lifting) option.inject option.simps(3))
moreover have "\<exists>z. ?P z"
using assms eval_Mn[OF assms(1)] by (metis option.distinct(1))
ultimately have "?P z"
using LeastI[of ?P] by blast
then have "eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>)"
by simp
then show "eval f (z # xs) \<down>= 0" by simp
show "\<And>y. y < z \<Longrightarrow> eval f (y # xs) \<down>\<noteq> 0"
- using not_less_Least[of _ ?P] `z = Least ?P` `?P z` less_trans by blast
+ using not_less_Least[of _ ?P] \<open>z = Least ?P\<close> \<open>?P z\<close> less_trans by blast
qed
lemma eval_Mn_diverg:
assumes "recfn (length xs) (Mn n f)"
shows "\<not> (\<exists>z. eval f (z # xs) \<down>= 0 \<and> (\<forall>y<z. eval f (y # xs) \<down>)) \<longleftrightarrow> eval (Mn n f) xs \<up>"
using assms eval_Mn[OF assms(1)] by simp
subsection \<open>Extensional equality\<close>
definition exteq :: "recf \<Rightarrow> recf \<Rightarrow> bool" (infix "\<simeq>" 55) where
"f \<simeq> g \<equiv> arity f = arity g \<and> (\<forall>xs. length xs = arity f \<longrightarrow> eval f xs = eval g xs)"
lemma exteq_refl: "f \<simeq> f"
using exteq_def by simp
lemma exteq_sym: "f \<simeq> g \<Longrightarrow> g \<simeq> f"
using exteq_def by simp
lemma exteq_trans: "f \<simeq> g \<Longrightarrow> g \<simeq> h \<Longrightarrow> f \<simeq> h"
using exteq_def by simp
lemma exteqI:
assumes "arity f = arity g" and "\<And>xs. length xs = arity f \<Longrightarrow> eval f xs = eval g xs"
shows "f \<simeq> g"
using assms exteq_def by simp
lemma exteqI1:
assumes "arity f = 1" and "arity g = 1" and "\<And>x. eval f [x] = eval g [x]"
shows "f \<simeq> g"
using assms exteqI by (metis One_nat_def Suc_length_conv length_0_conv)
text \<open>For every partial recursive function @{term f} there are
infinitely many extensionally equal ones, for example, those that wrap @{term
f} arbitrarily often in the identity function.\<close>
fun wrap_Id :: "recf \<Rightarrow> nat \<Rightarrow> recf" where
"wrap_Id f 0 = f"
| "wrap_Id f (Suc n) = Cn (arity f) (Id 1 0) [wrap_Id f n]"
lemma recfn_wrap_Id: "recfn a f \<Longrightarrow> recfn a (wrap_Id f n)"
using wellf_arity_nonzero by (induction n) auto
lemma exteq_wrap_Id: "recfn a f \<Longrightarrow> f \<simeq> wrap_Id f n"
proof (induction n)
case 0
then show ?case by (simp add: exteq_refl)
next
case (Suc n)
have "wrap_Id f n \<simeq> wrap_Id f (Suc n) "
proof (rule exteqI)
show "arity (wrap_Id f n) = arity (wrap_Id f (Suc n))"
using Suc by (simp add: recfn_wrap_Id)
show "eval (wrap_Id f n) xs = eval (wrap_Id f (Suc n)) xs"
if "length xs = arity (wrap_Id f n)" for xs
proof -
have "recfn (length xs) (Cn (arity f) (Id 1 0) [wrap_Id f n])"
using that Suc recfn_wrap_Id by (metis wrap_Id.simps(2))
then show "eval (wrap_Id f n) xs = eval (wrap_Id f (Suc n)) xs"
by auto
qed
qed
then show ?case using Suc exteq_trans by fast
qed
fun depth :: "recf \<Rightarrow> nat" where
"depth Z = 0"
| "depth S = 0"
| "depth (Id m n) = 0"
| "depth (Cn n f gs) = Suc (max (depth f) (Max (set (map depth gs))))"
| "depth (Pr n f g) = Suc (max (depth f) (depth g))"
| "depth (Mn n f) = Suc (depth f)"
lemma depth_wrap_Id: "recfn a f \<Longrightarrow> depth (wrap_Id f n) = depth f + n"
by (induction n) simp_all
lemma wrap_Id_injective:
assumes "recfn a f" and "wrap_Id f n\<^sub>1 = wrap_Id f n\<^sub>2"
shows "n\<^sub>1 = n\<^sub>2"
using assms by (metis add_left_cancel depth_wrap_Id)
lemma exteq_infinite:
assumes "recfn a f"
shows "infinite {g. recfn a g \<and> g \<simeq> f}" (is "infinite ?R")
proof -
have "inj (wrap_Id f)"
- using wrap_Id_injective `recfn a f` by (meson inj_onI)
+ using wrap_Id_injective \<open>recfn a f\<close> by (meson inj_onI)
then have "infinite (range (wrap_Id f))"
using finite_imageD by blast
moreover have "range (wrap_Id f) \<subseteq> ?R"
using assms exteq_sym exteq_wrap_Id recfn_wrap_Id by blast
ultimately show ?thesis by (simp add: infinite_super)
qed
subsection \<open>Primitive recursive and total functions\<close>
fun Mn_free :: "recf \<Rightarrow> bool" where
"Mn_free Z = True"
| "Mn_free S = True"
| "Mn_free (Id m n) = True"
| "Mn_free (Cn n f gs) = ((\<forall>g \<in> set gs. Mn_free g) \<and> Mn_free f)"
| "Mn_free (Pr n f g) = (Mn_free f \<and> Mn_free g)"
| "Mn_free (Mn n f) = False"
text \<open>This is our notion of $n$-ary primitive recursive function:\<close>
abbreviation prim_recfn :: "nat \<Rightarrow> recf \<Rightarrow> bool" where
"prim_recfn n f \<equiv> recfn n f \<and> Mn_free f"
definition total :: "recf \<Rightarrow> bool" where
"total f \<equiv> \<forall>xs. length xs = arity f \<longrightarrow> eval f xs \<down>"
lemma totalI [intro]:
assumes "\<And>xs. length xs = arity f \<Longrightarrow> eval f xs \<down>"
shows "total f"
using assms total_def by simp
lemma totalE [simp]:
assumes "total f" and "recfn n f" and "length xs = n"
shows "eval f xs \<down>"
using assms total_def by simp
lemma totalI1 :
assumes "recfn 1 f" and "\<And>x. eval f [x] \<down>"
shows "total f"
using assms totalI[of f] by (metis One_nat_def length_0_conv length_Suc_conv)
lemma totalI2:
assumes "recfn 2 f" and "\<And>x y. eval f [x, y] \<down>"
shows "total f"
using assms totalI[of f] by (smt length_0_conv length_Suc_conv numeral_2_eq_2)
lemma totalI3:
assumes "recfn 3 f" and "\<And>x y z. eval f [x, y, z] \<down>"
shows "total f"
using assms totalI[of f] by (smt length_0_conv length_Suc_conv numeral_3_eq_3)
lemma totalI4:
assumes "recfn 4 f" and "\<And>w x y z. eval f [w, x, y, z] \<down>"
shows "total f"
proof (rule totalI[of f])
fix xs :: "nat list"
assume "length xs = arity f"
then have "length xs = Suc (Suc (Suc (Suc 0)))"
using assms(1) by simp
then obtain w x y z where "xs = [w, x, y, z]"
by (smt Suc_length_conv length_0_conv)
then show "eval f xs \<down>" using assms(2) by simp
qed
lemma Mn_free_imp_total [intro]:
assumes "wellf f" and "Mn_free f"
shows "total f"
using assms
proof (induction f rule: Mn_free.induct)
case (5 n f g)
have "eval (Pr n f g) (x # xs) \<down>" if "length (x # xs) = arity (Pr n f g)" for x xs
using 5 that by (induction x) auto
then show ?case by (metis arity.simps(5) length_Suc_conv totalI)
qed (auto simp add: total_def eval_def)
lemma prim_recfn_total: "prim_recfn n f \<Longrightarrow> total f"
using Mn_free_imp_total by simp
lemma eval_Pr_prim_Suc:
assumes "h = Pr n f g" and "prim_recfn (Suc n) h" and "length xs = n"
shows "eval h (Suc x # xs) = eval g (x # the (eval h (x # xs)) # xs)"
using assms eval_Pr_converg_Suc prim_recfn_total by simp
lemma Cn_total:
assumes "\<forall>g\<in>set gs. total g" and "total f" and "recfn n (Cn n f gs)"
shows "total (Cn n f gs)"
using assms by (simp add: totalI)
lemma Pr_total:
assumes "total f" and "total g" and "recfn (Suc n) (Pr n f g)"
shows "total (Pr n f g)"
proof -
have "eval (Pr n f g) (x # xs) \<down>" if "length xs = n" for x xs
using that assms by (induction x) auto
then show ?thesis
using assms(3) totalI by (metis Suc_length_conv arity.simps(5))
qed
lemma eval_Mn_total:
assumes "recfn (length xs) (Mn n f)" and "total f"
shows "eval (Mn n f) xs =
(if (\<exists>z. eval f (z # xs) \<down>= 0)
then Some (LEAST z. eval f (z # xs) \<down>= 0)
else None)"
using assms by auto
section \<open>Simple functions\<close>
text \<open>This section, too, bears some similarity to Urban's formalization
in Xu et~al.~\cite{Universal_Turing_Machine-AFP}, but is more minimalistic in
scope.
As a general naming rule, instances of @{typ recf} and functions
returning such instances get names starting with @{text r_}. Typically, for
an @{text r_xyz} there will be a lemma @{text r_xyz_recfn} or @{text
r_xyz_prim} establishing its (primitive) recursiveness, arity, and
well-formedness. Moreover there will be a lemma @{text r_xyz} describing its
semantics, for which we will sometimes introduce an Isabelle function @{text
xyz}.\<close>
subsection \<open>Manipulating parameters\<close>
text \<open>Appending dummy parameters:\<close>
definition r_dummy :: "nat \<Rightarrow> recf \<Rightarrow> recf" where
"r_dummy n f \<equiv> Cn (arity f + n) f (map (\<lambda>i. Id (arity f + n) i) [0..<arity f])"
lemma r_dummy_prim [simp]:
"prim_recfn a f \<Longrightarrow> prim_recfn (a + n) (r_dummy n f)"
using wellf_arity_nonzero by (auto simp add: r_dummy_def)
lemma r_dummy_recfn [simp]:
"recfn a f \<Longrightarrow> recfn (a + n) (r_dummy n f)"
using wellf_arity_nonzero by (auto simp add: r_dummy_def)
lemma r_dummy [simp]:
"r_dummy n f = Cn (arity f + n) f (map (\<lambda>i. Id (arity f + n) i) [0..<arity f])"
unfolding r_dummy_def by simp
lemma r_dummy_append:
assumes "recfn (length xs) f" and "length ys = n"
shows "eval (r_dummy n f) (xs @ ys) = eval f xs"
proof -
let ?r = "r_dummy n f"
let ?gs = "map (\<lambda>i. Id (arity f + n) i) [0..<arity f]"
have "length ?gs = arity f" by simp
moreover have "?gs ! i = (Id (arity f + n) i)" if "i < arity f" for i
by (simp add: that)
moreover have *: "eval_wellf (?gs ! i) (xs @ ys) \<down>= xs ! i" if "i < arity f" for i
using that assms by (simp add: nth_append)
ultimately have "map (\<lambda>g. the (eval_wellf g (xs @ ys))) ?gs = xs"
by (metis (no_types, lifting) assms(1) length_map nth_equalityI nth_map option.sel)
moreover have "\<forall>g \<in> set ?gs. eval_wellf g (xs @ ys) \<down>"
using * by simp
moreover have "recfn (length (xs @ ys)) ?r"
using assms r_dummy_recfn by fastforce
ultimately show ?thesis
by (auto simp add: assms eval_def)
qed
text \<open>Shrinking a binary function to a unary one is useful when we want
to define a unary function via the @{term Pr} operation, which can only
construct @{typ recf}s of arity two or higher.\<close>
definition r_shrink :: "recf \<Rightarrow> recf" where
"r_shrink f \<equiv> Cn 1 f [Id 1 0, Id 1 0]"
lemma r_shrink_prim [simp]: "prim_recfn 2 f \<Longrightarrow> prim_recfn 1 (r_shrink f)"
by (simp add: r_shrink_def)
lemma r_shrink_recfn [simp]: "recfn 2 f \<Longrightarrow> recfn 1 (r_shrink f)"
by (simp add: r_shrink_def)
lemma r_shrink [simp]: "recfn 2 f \<Longrightarrow> eval (r_shrink f) [x] = eval f [x, x]"
by (simp add: r_shrink_def)
definition r_swap :: "recf \<Rightarrow> recf" where
"r_swap f \<equiv> Cn 2 f [Id 2 1, Id 2 0]"
lemma r_swap_recfn [simp]: "recfn 2 f \<Longrightarrow> recfn 2 (r_swap f)"
by (simp add: r_swap_def)
lemma r_swap_prim [simp]: "prim_recfn 2 f \<Longrightarrow> prim_recfn 2 (r_swap f)"
by (simp add: r_swap_def)
lemma r_swap [simp]: "recfn 2 f \<Longrightarrow> eval (r_swap f) [x, y] = eval f [y, x]"
by (simp add: r_swap_def)
text \<open>Prepending one dummy parameter:\<close>
definition r_shift :: "recf \<Rightarrow> recf" where
"r_shift f \<equiv> Cn (Suc (arity f)) f (map (\<lambda>i. Id (Suc (arity f)) (Suc i)) [0..<arity f])"
lemma r_shift_prim [simp]: "prim_recfn a f \<Longrightarrow> prim_recfn (Suc a) (r_shift f)"
by (simp add: r_shift_def)
lemma r_shift_recfn [simp]: "recfn a f \<Longrightarrow> recfn (Suc a) (r_shift f)"
by (simp add: r_shift_def)
lemma r_shift [simp]:
assumes "recfn (length xs) f"
shows "eval (r_shift f) (x # xs) = eval f xs"
proof -
let ?r = "r_shift f"
let ?gs = "map (\<lambda>i. Id (Suc (arity f)) (Suc i)) [0..<arity f]"
have "length ?gs = arity f" by simp
moreover have "?gs ! i = (Id (Suc (arity f)) (Suc i))" if "i < arity f" for i
by (simp add: that)
moreover have *: "eval (?gs ! i) (x # xs) \<down>= xs ! i" if "i < arity f" for i
using assms nth_append that by simp
ultimately have "map (\<lambda>g. the (eval g (x # xs))) ?gs = xs"
by (metis (no_types, lifting) assms length_map nth_equalityI nth_map option.sel)
moreover have "\<forall>g \<in> set ?gs. eval g (x # xs) \<noteq> None"
using * by simp
ultimately show ?thesis using r_shift_def assms by simp
qed
subsection \<open>Arithmetic and logic\<close>
text \<open>The unary constants:\<close>
fun r_const :: "nat \<Rightarrow> recf" where
"r_const 0 = Z"
| "r_const (Suc c) = Cn 1 S [r_const c]"
lemma r_const_prim [simp]: "prim_recfn 1 (r_const c)"
by (induction c) (simp_all)
lemma r_const [simp]: "eval (r_const c) [x] \<down>= c"
by (induction c) simp_all
text \<open>Constants of higher arities:\<close>
definition "r_constn n c \<equiv> if n = 0 then r_const c else r_dummy n (r_const c)"
lemma r_constn_prim [simp]: "prim_recfn (Suc n) (r_constn n c)"
unfolding r_constn_def by simp
lemma r_constn [simp]: "length xs = Suc n \<Longrightarrow> eval (r_constn n c) xs \<down>= c"
unfolding r_constn_def
by simp (metis length_0_conv length_Suc_conv r_const)
text \<open>We introduce addition, subtraction, and multiplication, but
interestingly enough we can make do without division.\<close>
definition "r_add \<equiv> Pr 1 (Id 1 0) (Cn 3 S [Id 3 1])"
lemma r_add_prim [simp]: "prim_recfn 2 r_add"
by (simp add: r_add_def)
lemma r_add [simp]: "eval r_add [a, b] \<down>= a + b"
unfolding r_add_def by (induction a) simp_all
definition "r_mul \<equiv> Pr 1 Z (Cn 3 r_add [Id 3 1, Id 3 2])"
lemma r_mul_prim [simp]: "prim_recfn 2 r_mul"
unfolding r_mul_def by simp
lemma r_mul [simp]: "eval r_mul [a, b] \<down>= a * b"
unfolding r_mul_def by (induction a) simp_all
definition "r_dec \<equiv> Cn 1 (Pr 1 Z (Id 3 0)) [Id 1 0, Id 1 0]"
lemma r_dec_prim [simp]: "prim_recfn 1 r_dec"
by (simp add: r_dec_def)
lemma r_dec [simp]: "eval r_dec [a] \<down>= a - 1"
proof -
have "eval (Pr 1 Z (Id 3 0)) [x, y] \<down>= x - 1" for x y
by (induction x) simp_all
then show ?thesis by (simp add: r_dec_def)
qed
definition "r_sub \<equiv> r_swap (Pr 1 (Id 1 0) (Cn 3 r_dec [Id 3 1]))"
lemma r_sub_prim [simp]: "prim_recfn 2 r_sub"
unfolding r_sub_def by simp
lemma r_sub [simp]: "eval r_sub [a, b] \<down>= a - b"
proof -
have "eval (Pr 1 (Id 1 0) (Cn 3 r_dec [Id 3 1])) [x, y] \<down>= y - x" for x y
by (induction x) simp_all
then show ?thesis unfolding r_sub_def by simp
qed
definition "r_sign \<equiv> r_shrink (Pr 1 Z (r_constn 2 1))"
lemma r_sign_prim [simp]: "prim_recfn 1 r_sign"
unfolding r_sign_def by simp
lemma r_sign [simp]: "eval r_sign [x] \<down>= (if x = 0 then 0 else 1)"
proof -
have "eval (Pr 1 Z (r_constn 2 1)) [x, y] \<down>= (if x = 0 then 0 else 1)" for x y
by (induction x) simp_all
then show ?thesis unfolding r_sign_def by simp
qed
text \<open>In the logical functions, true will be represented by zero, and
false will be represented by non-zero as argument and by one as
result.\<close>
definition "r_not \<equiv> Cn 1 r_sub [r_const 1, r_sign]"
lemma r_not_prim [simp]: "prim_recfn 1 r_not"
unfolding r_not_def by simp
lemma r_not [simp]: "eval r_not [x] \<down>= (if x = 0 then 1 else 0)"
unfolding r_not_def by simp
definition "r_nand \<equiv> Cn 2 r_not [r_add]"
lemma r_nand_prim [simp]: "prim_recfn 2 r_nand"
unfolding r_nand_def by simp
lemma r_nand [simp]: "eval r_nand [x, y] \<down>= (if x = 0 \<and> y = 0 then 1 else 0)"
unfolding r_nand_def by simp
definition "r_and \<equiv> Cn 2 r_not [r_nand]"
lemma r_and_prim [simp]: "prim_recfn 2 r_and"
unfolding r_and_def by simp
lemma r_and [simp]: "eval r_and [x, y] \<down>= (if x = 0 \<and> y = 0 then 0 else 1)"
unfolding r_and_def by simp
definition "r_or \<equiv> Cn 2 r_sign [r_mul]"
lemma r_or_prim [simp]: "prim_recfn 2 r_or"
unfolding r_or_def by simp
lemma r_or [simp]: "eval r_or [x, y] \<down>= (if x = 0 \<or> y = 0 then 0 else 1)"
unfolding r_or_def by simp
subsection \<open>Comparison and conditions\<close>
definition "r_ifz \<equiv>
let ifzero = (Cn 3 r_mul [r_dummy 2 r_not, Id 3 1]);
ifnzero = (Cn 3 r_mul [r_dummy 2 r_sign, Id 3 2])
in Cn 3 r_add [ifzero, ifnzero]"
lemma r_ifz_prim [simp]: "prim_recfn 3 r_ifz"
unfolding r_ifz_def by simp
lemma r_ifz [simp]: "eval r_ifz [cond, val0, val1] \<down>= (if cond = 0 then val0 else val1)"
unfolding r_ifz_def by (simp add: Let_def)
definition "r_eq \<equiv> Cn 2 r_sign [Cn 2 r_add [r_sub, r_swap r_sub]]"
lemma r_eq_prim [simp]: "prim_recfn 2 r_eq"
unfolding r_eq_def by simp
lemma r_eq [simp]: "eval r_eq [x, y] \<down>= (if x = y then 0 else 1)"
unfolding r_eq_def by simp
definition "r_ifeq \<equiv> Cn 4 r_ifz [r_dummy 2 r_eq, Id 4 2, Id 4 3]"
lemma r_ifeq_prim [simp]: "prim_recfn 4 r_ifeq"
unfolding r_ifeq_def by simp
lemma r_ifeq [simp]: "eval r_ifeq [a, b, v\<^sub>0, v\<^sub>1] \<down>= (if a = b then v\<^sub>0 else v\<^sub>1)"
unfolding r_ifeq_def using r_dummy_append[of r_eq "[a, b]" "[v\<^sub>0, v\<^sub>1]" 2]
by simp
definition "r_neq \<equiv> Cn 2 r_not [r_eq]"
lemma r_neq_prim [simp]: "prim_recfn 2 r_neq"
unfolding r_neq_def by simp
lemma r_neq [simp]: "eval r_neq [x, y] \<down>= (if x = y then 1 else 0)"
unfolding r_neq_def by simp
definition "r_ifle \<equiv> Cn 4 r_ifz [r_dummy 2 r_sub, Id 4 2, Id 4 3]"
lemma r_ifle_prim [simp]: "prim_recfn 4 r_ifle"
unfolding r_ifle_def by simp
lemma r_ifle [simp]: "eval r_ifle [a, b, v\<^sub>0, v\<^sub>1] \<down>= (if a \<le> b then v\<^sub>0 else v\<^sub>1)"
unfolding r_ifle_def using r_dummy_append[of r_sub "[a, b]" "[v\<^sub>0, v\<^sub>1]" 2]
by simp
definition "r_ifless \<equiv> Cn 4 r_ifle [Id 4 1, Id 4 0, Id 4 3, Id 4 2]"
lemma r_ifless_prim [simp]: "prim_recfn 4 r_ifless"
unfolding r_ifless_def by simp
lemma r_ifless [simp]: "eval r_ifless [a, b, v\<^sub>0, v\<^sub>1] \<down>= (if a < b then v\<^sub>0 else v\<^sub>1)"
unfolding r_ifless_def by simp
definition "r_less \<equiv> Cn 2 r_ifle [Id 2 1, Id 2 0, r_constn 1 1, r_constn 1 0]"
lemma r_less_prim [simp]: "prim_recfn 2 r_less"
unfolding r_less_def by simp
lemma r_less [simp]: "eval r_less [x, y] \<down>= (if x < y then 0 else 1)"
unfolding r_less_def by simp
definition "r_le \<equiv> Cn 2 r_ifle [Id 2 0, Id 2 1, r_constn 1 0, r_constn 1 1]"
lemma r_le_prim [simp]: "prim_recfn 2 r_le"
unfolding r_le_def by simp
lemma r_le [simp]: "eval r_le [x, y] \<down>= (if x \<le> y then 0 else 1)"
unfolding r_le_def by simp
text \<open>Arguments are evaluated eagerly. Therefore @{term "r_ifz"}, etc.
cannot be combined with a diverging function to implement a conditionally
diverging function in the naive way. The following function implements a
special case needed in the next section. A \hyperlink{p:r_lifz}{general lazy
version} of @{term "r_ifz"} will be introduced later with the help of a
universal function.\<close>
definition "r_ifeq_else_diverg \<equiv>
Cn 3 r_add [Id 3 2, Mn 3 (Cn 4 r_add [Id 4 0, Cn 4 r_eq [Id 4 1, Id 4 2]])]"
lemma r_ifeq_else_diverg_recfn [simp]: "recfn 3 r_ifeq_else_diverg"
unfolding r_ifeq_else_diverg_def by simp
lemma r_ifeq_else_diverg [simp]:
"eval r_ifeq_else_diverg [a, b, v] = (if a = b then Some v else None)"
unfolding r_ifeq_else_diverg_def by simp
section \<open>The halting problem\label{s:halting}\<close>
text \<open>Decidability will be treated more thoroughly in
Section~\ref{s:decidable}. But the halting problem is prominent enough to
deserve an early mention.\<close>
definition decidable :: "nat set \<Rightarrow> bool" where
"decidable X \<equiv> \<exists>f. recfn 1 f \<and> (\<forall>x. eval f [x] \<down>= (if x \<in> X then 1 else 0))"
text \<open>No matter how partial recursive functions are encoded as natural
numbers, the set of all codes of functions halting on their own code is
undecidable.\<close>
theorem halting_problem_undecidable:
fixes code :: "nat \<Rightarrow> recf"
assumes "\<And>f. recfn 1 f \<Longrightarrow> \<exists>i. code i = f"
shows "\<not> decidable {x. eval (code x) [x] \<down>}" (is "\<not> decidable ?K")
proof
assume "decidable ?K"
then obtain f where "recfn 1 f" and f: "\<forall>x. eval f [x] \<down>= (if x \<in> ?K then 1 else 0)"
using decidable_def by auto
define g where "g \<equiv> Cn 1 r_ifeq_else_diverg [f, Z, Z]"
then have "recfn 1 g"
- using `recfn 1 f` r_ifeq_else_diverg_recfn by simp
+ using \<open>recfn 1 f\<close> r_ifeq_else_diverg_recfn by simp
with assms obtain i where i: "code i = g" by auto
from g_def have "eval g [x] = (if x \<notin> ?K then Some 0 else None)" for x
- using r_ifeq_else_diverg_recfn `recfn 1 f` f by simp
+ using r_ifeq_else_diverg_recfn \<open>recfn 1 f\<close> f by simp
then have "eval g [i] \<down> \<longleftrightarrow> i \<notin> ?K" by simp
also have "... \<longleftrightarrow> eval (code i) [i] \<up>" by simp
also have "... \<longleftrightarrow> eval g [i] \<up>"
using i by simp
finally have "eval g [i] \<down> \<longleftrightarrow> eval g [i] \<up>" .
then show False by auto
qed
section \<open>Encoding tuples and lists\<close>
text \<open>This section is based on the Cantor encoding for pairs. Tuples
are encoded by repeated application of the pairing function, lists by pairing
their length with the code for a tuple. Thus tuples have a fixed length that
must be known when decoding, whereas lists are dynamically sized and know
their current length.\<close>
subsection \<open>Pairs and tuples\<close>
subsubsection \<open>The Cantor pairing function\<close>
definition "r_triangle \<equiv> r_shrink (Pr 1 Z (r_dummy 1 (Cn 2 S [r_add])))"
lemma r_triangle_prim: "prim_recfn 1 r_triangle"
unfolding r_triangle_def by simp
lemma r_triangle: "eval r_triangle [n] \<down>= Sum {0..n}"
proof -
let ?r = "r_dummy 1 (Cn 2 S [r_add])"
have "eval ?r [x, y, z] \<down>= Suc (x + y)" for x y z
using r_dummy_append[of "Cn 2 S [r_add]" "[x, y]" "[z]" 1] by simp
then have "eval (Pr 1 Z ?r) [x, y] \<down>= Sum {0..x}" for x y
by (induction x) simp_all
then show ?thesis unfolding r_triangle_def by simp
qed
lemma r_triangle_eq_triangle [simp]: "eval r_triangle [n] \<down>= triangle n"
using r_triangle gauss_sum_nat triangle_def by simp
definition "r_prod_encode \<equiv> Cn 2 r_add [Cn 2 r_triangle [r_add], Id 2 0]"
lemma r_prod_encode_prim [simp]: "prim_recfn 2 r_prod_encode"
unfolding r_prod_encode_def using r_triangle_prim by simp
lemma r_prod_encode [simp]: "eval r_prod_encode [m, n] \<down>= prod_encode (m, n)"
unfolding r_prod_encode_def prod_encode_def using r_triangle_prim by simp
text \<open>These abbreviations are just two more things borrowed from
Xu~et~al.~\cite{Universal_Turing_Machine-AFP}.\<close>
abbreviation "pdec1 z \<equiv> fst (prod_decode z)"
abbreviation "pdec2 z \<equiv> snd (prod_decode z)"
lemma pdec1_le: "pdec1 i \<le> i"
by (metis le_prod_encode_1 prod.collapse prod_decode_inverse)
lemma pdec2_le: "pdec2 i \<le> i"
by (metis le_prod_encode_2 prod.collapse prod_decode_inverse)
lemma pdec_less: "pdec2 i < Suc i"
using pdec2_le by (simp add: le_imp_less_Suc)
lemma pdec1_zero: "pdec1 0 = 0"
using pdec1_le by auto
definition "r_maxletr \<equiv>
Pr 1 Z (Cn 3 r_ifle [r_dummy 2 (Cn 1 r_triangle [S]), Id 3 2, Cn 3 S [Id 3 0], Id 3 1])"
lemma r_maxletr_prim: "prim_recfn 2 r_maxletr"
unfolding r_maxletr_def using r_triangle_prim by simp
lemma not_Suc_Greatest_not_Suc:
assumes "\<not> P (Suc x)" and "\<exists>x. P x"
shows "(GREATEST y. y \<le> x \<and> P y) = (GREATEST y. y \<le> Suc x \<and> P y)"
using assms by (metis le_SucI le_Suc_eq)
lemma r_maxletr: "eval r_maxletr [x\<^sub>0, x\<^sub>1] \<down>= (GREATEST y. y \<le> x\<^sub>0 \<and> triangle y \<le> x\<^sub>1)"
proof -
let ?g = "Cn 3 r_ifle [r_dummy 2 (Cn 1 r_triangle [S]), Id 3 2, Cn 3 S [Id 3 0], Id 3 1]"
have greatest:
"(if triangle (Suc x\<^sub>0) \<le> x\<^sub>1 then Suc x\<^sub>0 else (GREATEST y. y \<le> x\<^sub>0 \<and> triangle y \<le> x\<^sub>1)) =
(GREATEST y. y \<le> Suc x\<^sub>0 \<and> triangle y \<le> x\<^sub>1)"
for x\<^sub>0 x\<^sub>1
proof (cases "triangle (Suc x\<^sub>0) \<le> x\<^sub>1")
case True
then show ?thesis
using Greatest_equality[of "\<lambda>y. y \<le> Suc x\<^sub>0 \<and> triangle y \<le> x\<^sub>1"] by fastforce
next
case False
then show ?thesis
using not_Suc_Greatest_not_Suc[of "\<lambda>y. triangle y \<le> x\<^sub>1" x\<^sub>0] by fastforce
qed
show ?thesis
unfolding r_maxletr_def using r_triangle_prim
proof (induction x\<^sub>0)
case 0
then show ?case
using Greatest_equality[of "\<lambda>y. y \<le> 0 \<and> triangle y \<le> x\<^sub>1" 0] by simp
next
case (Suc x\<^sub>0)
then show ?case using greatest by simp
qed
qed
definition "r_maxlt \<equiv> r_shrink r_maxletr"
lemma r_maxlt_prim: "prim_recfn 1 r_maxlt"
unfolding r_maxlt_def using r_maxletr_prim by simp
lemma r_maxlt: "eval r_maxlt [e] \<down>= (GREATEST y. triangle y \<le> e)"
proof -
have "y \<le> triangle y" for y
by (induction y) auto
then have "triangle y \<le> e \<Longrightarrow> y \<le> e" for y e
using order_trans by blast
then have "(GREATEST y. y \<le> e \<and> triangle y \<le> e) = (GREATEST y. triangle y \<le> e)"
by metis
moreover have "eval r_maxlt [e] \<down>= (GREATEST y. y \<le> e \<and> triangle y \<le> e)"
using r_maxletr r_shrink r_maxlt_def r_maxletr_prim by fastforce
ultimately show ?thesis by simp
qed
definition "pdec1' e \<equiv> e - triangle (GREATEST y. triangle y \<le> e)"
definition "pdec2' e \<equiv> (GREATEST y. triangle y \<le> e) - pdec1' e"
lemma max_triangle_bound: "triangle z \<le> e \<Longrightarrow> z \<le> e"
by (metis Suc_pred add_leD2 less_Suc_eq triangle_Suc zero_le zero_less_Suc)
lemma triangle_greatest_le: "triangle (GREATEST y. triangle y \<le> e) \<le> e"
using max_triangle_bound GreatestI_nat[of "\<lambda>y. triangle y \<le> e" 0 e] by simp
lemma prod_encode_pdec': "prod_encode (pdec1' e, pdec2' e) = e"
proof -
let ?P = "\<lambda>y. triangle y \<le> e"
let ?y = "GREATEST y. ?P y"
have "pdec1' e \<le> ?y"
proof (rule ccontr)
assume "\<not> pdec1' e \<le> ?y"
then have "e - triangle ?y > ?y"
using pdec1'_def by simp
then have "?P (Suc ?y)" by simp
moreover have "\<forall>z. ?P z \<longrightarrow> z \<le> e"
using max_triangle_bound by simp
ultimately have "Suc ?y \<le> ?y"
using Greatest_le_nat[of ?P "Suc ?y" e] by blast
then show False by simp
qed
then have "pdec1' e + pdec2' e = ?y"
using pdec1'_def pdec2'_def by simp
then have "prod_encode (pdec1' e, pdec2' e) = triangle ?y + pdec1' e"
by (simp add: prod_encode_def)
then show ?thesis using pdec1'_def triangle_greatest_le by simp
qed
lemma pdec':
"pdec1' e = pdec1 e"
"pdec2' e = pdec2 e"
using prod_encode_pdec' prod_encode_inverse by (metis fst_conv, metis snd_conv)
definition "r_pdec1 \<equiv> Cn 1 r_sub [Id 1 0, Cn 1 r_triangle [r_maxlt]]"
lemma r_pdec1_prim [simp]: "prim_recfn 1 r_pdec1"
unfolding r_pdec1_def using r_triangle_prim r_maxlt_prim by simp
lemma r_pdec1 [simp]: "eval r_pdec1 [e] \<down>= pdec1 e"
unfolding r_pdec1_def using r_triangle_prim r_maxlt_prim pdec' pdec1'_def
by (simp add: r_maxlt)
definition "r_pdec2 \<equiv> Cn 1 r_sub [r_maxlt, r_pdec1]"
lemma r_pdec2_prim [simp]: "prim_recfn 1 r_pdec2"
unfolding r_pdec2_def using r_maxlt_prim by simp
lemma r_pdec2 [simp]: "eval r_pdec2 [e] \<down>= pdec2 e"
unfolding r_pdec2_def using r_maxlt_prim r_maxlt pdec' pdec2'_def by simp
abbreviation "pdec12 i \<equiv> pdec1 (pdec2 i)"
abbreviation "pdec22 i \<equiv> pdec2 (pdec2 i)"
abbreviation "pdec122 i \<equiv> pdec1 (pdec22 i)"
abbreviation "pdec222 i \<equiv> pdec2 (pdec22 i)"
definition "r_pdec12 \<equiv> Cn 1 r_pdec1 [r_pdec2]"
lemma r_pdec12_prim [simp]: "prim_recfn 1 r_pdec12"
unfolding r_pdec12_def by simp
lemma r_pdec12 [simp]: "eval r_pdec12 [e] \<down>= pdec12 e"
unfolding r_pdec12_def by simp
definition "r_pdec22 \<equiv> Cn 1 r_pdec2 [r_pdec2]"
lemma r_pdec22_prim [simp]: "prim_recfn 1 r_pdec22"
unfolding r_pdec22_def by simp
lemma r_pdec22 [simp]: "eval r_pdec22 [e] \<down>= pdec22 e"
unfolding r_pdec22_def by simp
definition "r_pdec122 \<equiv> Cn 1 r_pdec1 [r_pdec22]"
lemma r_pdec122_prim [simp]: "prim_recfn 1 r_pdec122"
unfolding r_pdec122_def by simp
lemma r_pdec122 [simp]: "eval r_pdec122 [e] \<down>= pdec122 e"
unfolding r_pdec122_def by simp
definition "r_pdec222 \<equiv> Cn 1 r_pdec2 [r_pdec22]"
lemma r_pdec222_prim [simp]: "prim_recfn 1 r_pdec222"
unfolding r_pdec222_def by simp
lemma r_pdec222 [simp]: "eval r_pdec222 [e] \<down>= pdec222 e"
unfolding r_pdec222_def by simp
subsubsection \<open>The Cantor tuple function\<close>
text \<open>The empty tuple gets no code, whereas singletons are encoded by their
only element and other tuples by recursively applying the pairing function.
This yields, for every $n$, the function @{term "tuple_encode n"}, which is a
bijection between the natural numbers and the lists of length $(n + 1)$.\<close>
fun tuple_encode :: "nat \<Rightarrow> nat list \<Rightarrow> nat" where
"tuple_encode n [] = undefined"
| "tuple_encode 0 (x # xs) = x"
| "tuple_encode (Suc n) (x # xs) = prod_encode (x, tuple_encode n xs)"
lemma tuple_encode_prod_encode: "tuple_encode 1 [x, y] = prod_encode (x, y)"
by simp
fun tuple_decode where
"tuple_decode 0 i = [i]"
| "tuple_decode (Suc n) i = pdec1 i # tuple_decode n (pdec2 i)"
lemma tuple_encode_decode [simp]:
"tuple_encode (length xs - 1) (tuple_decode (length xs - 1) i) = i"
proof (induction "length xs - 1" arbitrary: xs i)
case 0
then show ?case by simp
next
case (Suc n)
then have "length xs - 1 > 0" by simp
with Suc have *: "tuple_encode n (tuple_decode n j) = j" for j
by (metis diff_Suc_1 length_tl)
from Suc have "tuple_decode (Suc n) i = pdec1 i # tuple_decode n (pdec2 i)"
using tuple_decode.simps(2) by blast
then have "tuple_encode (Suc n) (tuple_decode (Suc n) i) =
tuple_encode (Suc n) (pdec1 i # tuple_decode n (pdec2 i))"
using Suc by simp
also have "... = prod_encode (pdec1 i, tuple_encode n (tuple_decode n (pdec2 i)))"
by simp
also have "... = prod_encode (pdec1 i, pdec2 i)"
using Suc * by simp
also have "... = i" by simp
finally have "tuple_encode (Suc n) (tuple_decode (Suc n) i) = i" .
then show ?case by (simp add: Suc.hyps(2))
qed
lemma tuple_encode_decode' [simp]: "tuple_encode n (tuple_decode n i) = i"
using tuple_encode_decode by (metis Ex_list_of_length diff_Suc_1 length_Cons)
lemma tuple_decode_encode:
assumes "length xs > 0"
shows "tuple_decode (length xs - 1) (tuple_encode (length xs - 1) xs) = xs"
using assms
proof (induction "length xs - 1" arbitrary: xs)
case 0
moreover from this have "length xs = 1" by linarith
ultimately show ?case
by (metis One_nat_def length_0_conv length_Suc_conv tuple_decode.simps(1)
tuple_encode.simps(2))
next
case (Suc n)
let ?t = "tl xs"
let ?i = "tuple_encode (Suc n) xs"
have "length ?t > 0" and "length ?t - 1 = n"
using Suc by simp_all
then have "tuple_decode n (tuple_encode n ?t) = ?t"
using Suc by blast
moreover have "?i = prod_encode (hd xs, tuple_encode n ?t)"
using Suc by (metis hd_Cons_tl length_greater_0_conv tuple_encode.simps(3))
moreover have "tuple_decode (Suc n) ?i = pdec1 ?i # tuple_decode n (pdec2 ?i)"
using tuple_decode.simps(2) by blast
ultimately have "tuple_decode (Suc n) ?i = xs"
using Suc.prems by simp
then show ?case by (simp add: Suc.hyps(2))
qed
lemma tuple_decode_encode' [simp]:
assumes "length xs = Suc n"
shows "tuple_decode n (tuple_encode n xs) = xs"
using assms tuple_decode_encode by (metis diff_Suc_1 zero_less_Suc)
lemma tuple_decode_length [simp]: "length (tuple_decode n i) = Suc n"
by (induction n arbitrary: i) simp_all
lemma tuple_decode_nonzero:
assumes "n > 0"
shows "tuple_decode n i = pdec1 i # tuple_decode (n - 1) (pdec2 i)"
using assms by (metis One_nat_def Suc_pred tuple_decode.simps(2))
text \<open>The tuple encoding functions are primitive recursive.\<close>
fun r_tuple_encode :: "nat \<Rightarrow> recf" where
"r_tuple_encode 0 = Id 1 0"
| "r_tuple_encode (Suc n) =
Cn (Suc (Suc n)) r_prod_encode [Id (Suc (Suc n)) 0, r_shift (r_tuple_encode n)]"
lemma r_tuple_encode_prim [simp]: "prim_recfn (Suc n) (r_tuple_encode n)"
by (induction n) simp_all
lemma r_tuple_encode:
assumes "length xs = Suc n"
shows "eval (r_tuple_encode n) xs \<down>= tuple_encode n xs"
using assms
proof (induction n arbitrary: xs)
case 0
then show ?case
by (metis One_nat_def eval_Id length_Suc_conv nth_Cons_0
r_tuple_encode.simps(1) tuple_encode.simps(2) zero_less_one)
next
case (Suc n)
then obtain y ys where y_ys: "y # ys = xs"
by (metis length_Suc_conv)
with Suc have "eval (r_tuple_encode n) ys \<down>= tuple_encode n ys"
by auto
with y_ys have "eval (r_shift (r_tuple_encode n)) xs \<down>= tuple_encode n ys"
using Suc.prems r_shift_prim r_tuple_encode_prim by auto
moreover have "eval (Id (Suc (Suc n)) 0) xs \<down>= y"
using y_ys Suc.prems by auto
ultimately have "eval (r_tuple_encode (Suc n)) xs \<down>= prod_encode (y, tuple_encode n ys)"
using Suc.prems by simp
then show ?case using y_ys by auto
qed
subsubsection \<open>Functions on encoded tuples\<close>
text \<open>The function for accessing the $n$-th element of a tuple returns
$0$ for out-of-bounds access.\<close>
definition e_tuple_nth :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_tuple_nth a i n \<equiv> if n \<le> a then (tuple_decode a i) ! n else 0"
lemma e_tuple_nth_le [simp]: "n \<le> a \<Longrightarrow> e_tuple_nth a i n = (tuple_decode a i) ! n"
using e_tuple_nth_def by simp
lemma e_tuple_nth_gr [simp]: "n > a \<Longrightarrow> e_tuple_nth a i n = 0"
using e_tuple_nth_def by simp
lemma tuple_decode_pdec2: "tuple_decode a (pdec2 es) = tl (tuple_decode (Suc a) es)"
by simp
fun iterate :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a)" where
"iterate 0 f = id"
| "iterate (Suc n) f = f \<circ> (iterate n f)"
lemma iterate_additive:
assumes "iterate t\<^sub>1 f x = y" and "iterate t\<^sub>2 f y = z"
shows "iterate (t\<^sub>1 + t\<^sub>2) f x = z"
using assms by (induction t\<^sub>2 arbitrary: z) auto
lemma iterate_additive': "iterate (t\<^sub>1 + t\<^sub>2) f x = iterate t\<^sub>2 f (iterate t\<^sub>1 f x)"
using iterate_additive by metis
lemma e_tuple_nth_elementary:
assumes "k \<le> a"
shows "e_tuple_nth a i k = (if a = k then (iterate k pdec2 i) else (pdec1 (iterate k pdec2 i)))"
proof -
have *: "tuple_decode (a - k) (iterate k pdec2 i) = drop k (tuple_decode a i)"
using assms
by (induction k) (simp, simp add: Suc_diff_Suc tuple_decode_pdec2 drop_Suc tl_drop)
show ?thesis
proof (cases "a = k")
case True
then have "tuple_decode 0 (iterate k pdec2 i) = drop k (tuple_decode a i)"
using assms * by simp
moreover from this have "drop k (tuple_decode a i) = [tuple_decode a i ! k]"
using assms True by (metis nth_via_drop tuple_decode.simps(1))
ultimately show ?thesis using True by simp
next
case False
with assms have "a - k > 0" by simp
with * have "tuple_decode (a - k) (iterate k pdec2 i) = drop k (tuple_decode a i)"
by simp
then have "pdec1 (iterate k pdec2 i) = hd (drop k (tuple_decode a i))"
- using tuple_decode_nonzero `a - k > 0` by (metis list.sel(1))
- with `a - k > 0` have "pdec1 (iterate k pdec2 i) = (tuple_decode a i) ! k"
+ using tuple_decode_nonzero \<open>a - k > 0\<close> by (metis list.sel(1))
+ with \<open>a - k > 0\<close> have "pdec1 (iterate k pdec2 i) = (tuple_decode a i) ! k"
by (simp add: hd_drop_conv_nth)
with False assms show ?thesis by simp
qed
qed
definition "r_nth_inbounds \<equiv>
let r = Pr 1 (Id 1 0) (Cn 3 r_pdec2 [Id 3 1])
in Cn 3 r_ifeq
[Id 3 0,
Id 3 2,
Cn 3 r [Id 3 2, Id 3 1],
Cn 3 r_pdec1 [Cn 3 r [Id 3 2, Id 3 1]]]"
lemma r_nth_inbounds_prim: "prim_recfn 3 r_nth_inbounds"
unfolding r_nth_inbounds_def by (simp add: Let_def)
lemma r_nth_inbounds:
"k \<le> a \<Longrightarrow> eval r_nth_inbounds [a, i, k] \<down>= e_tuple_nth a i k"
"eval r_nth_inbounds [a, i, k] \<down>"
proof -
let ?r = "Pr 1 (Id 1 0) (Cn 3 r_pdec2 [Id 3 1])"
let ?h = "Cn 3 ?r [Id 3 2, Id 3 1]"
have "eval ?r [k, i] \<down>= iterate k pdec2 i" for k i
using r_pdec2_prim by (induction k) (simp_all)
then have "eval ?h [a, i, k] \<down>= iterate k pdec2 i"
using r_pdec2_prim by simp
then have "eval r_nth_inbounds [a, i, k] \<down>=
(if a = k then iterate k pdec2 i else pdec1 (iterate k pdec2 i))"
unfolding r_nth_inbounds_def by (simp add: Let_def)
then show "k \<le> a \<Longrightarrow> eval r_nth_inbounds [a, i, k] \<down>= e_tuple_nth a i k"
and "eval r_nth_inbounds [a, i, k] \<down>"
using e_tuple_nth_elementary by simp_all
qed
definition "r_tuple_nth \<equiv>
Cn 3 r_ifle [Id 3 2, Id 3 0, r_nth_inbounds, r_constn 2 0]"
lemma r_tuple_nth_prim: "prim_recfn 3 r_tuple_nth"
unfolding r_tuple_nth_def using r_nth_inbounds_prim by simp
lemma r_tuple_nth [simp]: "eval r_tuple_nth [a, i, k] \<down>= e_tuple_nth a i k"
unfolding r_tuple_nth_def using r_nth_inbounds_prim r_nth_inbounds by simp
subsection \<open>Lists\<close>
subsubsection \<open>Encoding and decoding\<close>
text \<open>Lists are encoded by pairing the length of the list with the code
for the tuple made up of the list's elements. Then all these codes are
incremented in order to make room for the empty list
(cf.~Rogers~\cite[p.~71]{Rogers87}).\<close>
fun list_encode :: "nat list \<Rightarrow> nat" where
"list_encode [] = 0"
| "list_encode (x # xs) = Suc (prod_encode (length xs, tuple_encode (length xs) (x # xs)))"
lemma list_encode_0 [simp]: "list_encode xs = 0 \<Longrightarrow> xs = []"
using list_encode.elims by blast
lemma list_encode_1: "list_encode [0] = 1"
by (simp add: prod_encode_def)
fun list_decode :: "nat \<Rightarrow> nat list" where
"list_decode 0 = []"
| "list_decode (Suc n) = tuple_decode (pdec1 n) (pdec2 n)"
lemma list_encode_decode [simp]: "list_encode (list_decode n) = n"
proof (cases n)
case 0
then show ?thesis by simp
next
case (Suc k)
then have *: "list_decode n = tuple_decode (pdec1 k) (pdec2 k)" (is "_ = ?t")
by simp
then obtain x xs where xxs: "x # xs = ?t"
by (metis tuple_decode.elims)
then have "list_encode ?t = list_encode (x # xs)" by simp
then have 1: "list_encode ?t = Suc (prod_encode (length xs, tuple_encode (length xs) (x # xs)))"
by simp
have 2: "length xs = length ?t - 1"
using xxs by (metis length_tl list.sel(3))
then have 3: "length xs = pdec1 k"
using * by simp
then have "tuple_encode (length ?t - 1) ?t = pdec2 k"
using 2 tuple_encode_decode by metis
then have "list_encode ?t = Suc (prod_encode (pdec1 k, pdec2 k))"
using 1 2 3 xxs by simp
with * Suc show ?thesis by simp
qed
lemma list_decode_encode [simp]: "list_decode (list_encode xs) = xs"
proof (cases xs)
case Nil
then show ?thesis by simp
next
case (Cons y ys)
then have "list_encode xs =
Suc (prod_encode (length ys, tuple_encode (length ys) xs))"
(is "_ = Suc ?i")
by simp
then have "list_decode (Suc ?i) = tuple_decode (pdec1 ?i) (pdec2 ?i)" by simp
moreover have "pdec1 ?i = length ys" by simp
moreover have "pdec2 ?i = tuple_encode (length ys) xs" by simp
ultimately have "list_decode (Suc ?i) =
tuple_decode (length ys) (tuple_encode (length ys) xs)"
by simp
moreover have "length ys = length xs - 1"
using Cons by simp
ultimately have "list_decode (Suc ?i) =
tuple_decode (length xs - 1) (tuple_encode (length xs - 1) xs)"
by simp
then show ?thesis using Cons by simp
qed
abbreviation singleton_encode :: "nat \<Rightarrow> nat" where
"singleton_encode x \<equiv> list_encode [x]"
lemma list_decode_singleton: "list_decode (singleton_encode x) = [x]"
by simp
definition "r_singleton_encode \<equiv> Cn 1 S [Cn 1 r_prod_encode [Z, Id 1 0]]"
lemma r_singleton_encode_prim [simp]: "prim_recfn 1 r_singleton_encode"
unfolding r_singleton_encode_def by simp
lemma r_singleton_encode [simp]: "eval r_singleton_encode [x] \<down>= singleton_encode x"
unfolding r_singleton_encode_def by simp
definition r_list_encode :: "nat \<Rightarrow> recf" where
"r_list_encode n \<equiv> Cn (Suc n) S [Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]]"
lemma r_list_encode_prim [simp]: "prim_recfn (Suc n) (r_list_encode n)"
unfolding r_list_encode_def by simp
lemma r_list_encode:
assumes "length xs = Suc n"
shows "eval (r_list_encode n) xs \<down>= list_encode xs"
proof -
have "eval (r_tuple_encode n) xs \<down>"
by (simp add: assms r_tuple_encode)
then have "eval (Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]) xs \<down>"
using assms by simp
then have "eval (r_list_encode n) xs =
eval S [the (eval (Cn (Suc n) r_prod_encode [r_constn n n, r_tuple_encode n]) xs)]"
unfolding r_list_encode_def using assms r_tuple_encode by simp
moreover from assms obtain y ys where "xs = y # ys"
by (meson length_Suc_conv)
ultimately show ?thesis
unfolding r_list_encode_def using assms r_tuple_encode by simp
qed
subsubsection \<open>Functions on encoded lists\<close>
text \<open>The functions in this section mimic those on type @{typ "nat
list"}. Their names are prefixed by @{text e_} and the names of the
corresponding @{typ recf}s by @{text r_}.\<close>
abbreviation e_tl :: "nat \<Rightarrow> nat" where
"e_tl e \<equiv> list_encode (tl (list_decode e))"
text \<open>In order to turn @{term e_tl} into a partial recursive function
we first represent it in a more elementary way.\<close>
lemma e_tl_elementary:
"e_tl e =
(if e = 0 then 0
else if pdec1 (e - 1) = 0 then 0
else Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1))))"
proof (cases e)
case 0
then show ?thesis by simp
next
case Suc_d: (Suc d)
then show ?thesis
proof (cases "pdec1 d")
case 0
then show ?thesis using Suc_d by simp
next
case (Suc a)
have *: "list_decode e = tuple_decode (pdec1 d) (pdec2 d)"
using Suc_d by simp
with Suc obtain x xs where xxs: "list_decode e = x # xs" by simp
then have **: "e_tl e = list_encode xs" by simp
have "list_decode (Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1)))) =
tuple_decode (pdec1 (e - 1) - 1) (pdec22 (e - 1))"
(is "?lhs = _")
by simp
also have "... = tuple_decode a (pdec22 (e - 1))"
using Suc Suc_d by simp
also have "... = tl (tuple_decode (Suc a) (pdec2 (e - 1)))"
using tuple_decode_pdec2 Suc by presburger
also have "... = tl (tuple_decode (pdec1 (e - 1)) (pdec2 (e - 1)))"
using Suc Suc_d by auto
also have "... = tl (list_decode e)"
using * Suc_d by simp
also have "... = xs"
using xxs by simp
finally have "?lhs = xs" .
then have "list_encode ?lhs = list_encode xs" by simp
then have "Suc (prod_encode (pdec1 (e - 1) - 1, pdec22 (e - 1))) = list_encode xs"
using list_encode_decode by metis
then show ?thesis using ** Suc_d Suc by simp
qed
qed
definition "r_tl \<equiv>
let r = Cn 1 r_pdec1 [r_dec]
in Cn 1 r_ifz
[Id 1 0,
Z,
Cn 1 r_ifz
[r, Z, Cn 1 S [Cn 1 r_prod_encode [Cn 1 r_dec [r], Cn 1 r_pdec22 [r_dec]]]]]"
lemma r_tl_prim [simp]: "prim_recfn 1 r_tl"
unfolding r_tl_def by (simp add: Let_def)
lemma r_tl [simp]: "eval r_tl [e] \<down>= e_tl e"
unfolding r_tl_def using e_tl_elementary by (simp add: Let_def)
text \<open>We define the head of the empty encoded list to be zero.\<close>
definition e_hd :: "nat \<Rightarrow> nat" where
"e_hd e \<equiv> if e = 0 then 0 else hd (list_decode e)"
lemma e_hd [simp]:
assumes "list_decode e = x # xs"
shows "e_hd e = x"
using e_hd_def assms by auto
lemma e_hd_0 [simp]: "e_hd 0 = 0"
using e_hd_def by simp
lemma e_hd_neq_0 [simp]:
assumes "e \<noteq> 0"
shows "e_hd e = hd (list_decode e)"
using e_hd_def assms by simp
definition "r_hd \<equiv>
Cn 1 r_ifz [Cn 1 r_pdec1 [r_dec], Cn 1 r_pdec2 [r_dec], Cn 1 r_pdec12 [r_dec]]"
lemma r_hd_prim [simp]: "prim_recfn 1 r_hd"
unfolding r_hd_def by simp
lemma r_hd [simp]: "eval r_hd [e] \<down>= e_hd e"
proof -
have "e_hd e = (if pdec1 (e - 1) = 0 then pdec2 (e - 1) else pdec12 (e - 1))"
proof (cases e)
case 0
then show ?thesis using pdec1_zero pdec2_le by auto
next
case (Suc d)
then show ?thesis by (cases "pdec1 d") (simp_all add: pdec1_zero)
qed
then show ?thesis unfolding r_hd_def by simp
qed
abbreviation e_length :: "nat \<Rightarrow> nat" where
"e_length e \<equiv> length (list_decode e)"
lemma e_length_0: "e_length e = 0 \<Longrightarrow> e = 0"
by (metis list_encode.simps(1) length_0_conv list_encode_decode)
definition "r_length \<equiv> Cn 1 r_ifz [Id 1 0, Z, Cn 1 S [Cn 1 r_pdec1 [r_dec]]]"
lemma r_length_prim [simp]: "prim_recfn 1 r_length"
unfolding r_length_def by simp
lemma r_length [simp]: "eval r_length [e] \<down>= e_length e"
unfolding r_length_def by (cases e) simp_all
text \<open>Accessing an encoded list out of bounds yields zero.\<close>
definition e_nth :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_nth e n \<equiv> if e = 0 then 0 else e_tuple_nth (pdec1 (e - 1)) (pdec2 (e - 1)) n"
lemma e_nth [simp]:
"e_nth e n = (if n < e_length e then (list_decode e) ! n else 0)"
by (cases e) (simp_all add: e_nth_def e_tuple_nth_def)
lemma e_hd_nth0: "e_hd e = e_nth e 0"
by (simp add: e_hd_def e_length_0 hd_conv_nth)
definition "r_nth \<equiv>
Cn 2 r_ifz
[Id 2 0,
r_constn 1 0,
Cn 2 r_tuple_nth
[Cn 2 r_pdec1 [r_dummy 1 r_dec], Cn 2 r_pdec2 [r_dummy 1 r_dec], Id 2 1]]"
lemma r_nth_prim [simp]: "prim_recfn 2 r_nth"
unfolding r_nth_def using r_tuple_nth_prim by simp
lemma r_nth [simp]: "eval r_nth [e, n] \<down>= e_nth e n"
unfolding r_nth_def e_nth_def using r_tuple_nth_prim by simp
definition "r_rev_aux \<equiv>
Pr 1 r_hd (Cn 3 r_prod_encode [Cn 3 r_nth [Id 3 2, Cn 3 S [Id 3 0]], Id 3 1])"
lemma r_rev_aux_prim: "prim_recfn 2 r_rev_aux"
unfolding r_rev_aux_def by simp
lemma r_rev_aux:
assumes "list_decode e = xs" and "length xs > 0" and "i < length xs"
shows "eval r_rev_aux [i, e] \<down>= tuple_encode i (rev (take (Suc i) xs))"
using assms(3)
proof (induction i)
case 0
then show ?case
unfolding r_rev_aux_def using assms e_hd_def r_hd by (auto simp add: take_Suc)
next
case (Suc i)
let ?g = "Cn 3 r_prod_encode [Cn 3 r_nth [Id 3 2, Cn 3 S [Id 3 0]], Id 3 1]"
from Suc have "eval r_rev_aux [Suc i, e] = eval ?g [i, the (eval r_rev_aux [i, e]), e]"
unfolding r_rev_aux_def by simp
also have "... \<down>= prod_encode (xs ! (Suc i), tuple_encode i (rev (take (Suc i) xs)))"
using Suc by (simp add: assms(1))
finally show ?case by (simp add: Suc.prems take_Suc_conv_app_nth)
qed
corollary r_rev_aux_full:
assumes "list_decode e = xs" and "length xs > 0"
shows "eval r_rev_aux [length xs - 1, e] \<down>= tuple_encode (length xs - 1) (rev xs)"
using r_rev_aux assms by simp
lemma r_rev_aux_total: "eval r_rev_aux [i, e] \<down>"
using r_rev_aux_prim totalE by fastforce
definition "r_rev \<equiv>
Cn 1 r_ifz
[Id 1 0,
Z,
Cn 1 S
[Cn 1 r_prod_encode
[Cn 1 r_dec [r_length], Cn 1 r_rev_aux [Cn 1 r_dec [r_length], Id 1 0]]]]"
lemma r_rev_prim [simp]: "prim_recfn 1 r_rev"
unfolding r_rev_def using r_rev_aux_prim by simp
lemma r_rev [simp]: "eval r_rev [e] \<down>= list_encode (rev (list_decode e))"
proof -
let ?d = "Cn 1 r_dec [r_length]"
let ?a = "Cn 1 r_rev_aux [?d, Id 1 0]"
let ?p = "Cn 1 r_prod_encode [?d, ?a]"
let ?s = "Cn 1 S [?p]"
have eval_a: "eval ?a [e] = eval r_rev_aux [e_length e - 1, e]"
using r_rev_aux_prim by simp
then have "eval ?s [e] \<down>"
using r_rev_aux_prim by (simp add: r_rev_aux_total)
then have *: "eval r_rev [e] \<down>= (if e = 0 then 0 else the (eval ?s [e]))"
using r_rev_aux_prim by (simp add: r_rev_def)
show ?thesis
proof (cases "e = 0")
case True
then show ?thesis using * by simp
next
case False
then obtain xs where xs: "xs = list_decode e" "length xs > 0"
using e_length_0 by auto
then have len: "length xs = e_length e" by simp
with eval_a have "eval ?a [e] = eval r_rev_aux [length xs - 1, e]"
by simp
then have "eval ?a [e] \<down>= tuple_encode (length xs - 1) (rev xs)"
using xs r_rev_aux_full by simp
then have "eval ?s [e] \<down>=
Suc (prod_encode (length xs - 1, tuple_encode (length xs - 1) (rev xs)))"
using len r_rev_aux_prim by simp
then have "eval ?s [e] \<down>=
Suc (prod_encode
(length (rev xs) - 1, tuple_encode (length (rev xs) - 1) (rev xs)))"
by simp
moreover have "length (rev xs) > 0"
using xs by simp
ultimately have "eval ?s [e] \<down>= list_encode (rev xs)"
by (metis list_encode.elims diff_Suc_1 length_Cons length_greater_0_conv)
then show ?thesis using xs * by simp
qed
qed
abbreviation e_cons :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_cons e es \<equiv> list_encode (e # list_decode es)"
lemma e_cons_elementary:
"e_cons e es =
(if es = 0 then Suc (prod_encode (0, e))
else Suc (prod_encode (e_length es, prod_encode (e, pdec2 (es - 1)))))"
proof (cases "es = 0")
case True
then show ?thesis by simp
next
case False
then have "e_length es = Suc (pdec1 (es - 1))"
by (metis list_decode.elims diff_Suc_1 tuple_decode_length)
moreover have "es = e_tl (list_encode (e # list_decode es))"
by (metis list.sel(3) list_decode_encode list_encode_decode)
ultimately show ?thesis
using False e_tl_elementary
by (metis list_decode.simps(2) diff_Suc_1 list_encode_decode prod.sel(1)
prod_encode_inverse snd_conv tuple_decode.simps(2))
qed
definition "r_cons_else \<equiv>
Cn 2 S
[Cn 2 r_prod_encode
[Cn 2 r_length
[Id 2 1], Cn 2 r_prod_encode [Id 2 0, Cn 2 r_pdec2 [Cn 2 r_dec [Id 2 1]]]]]"
lemma r_cons_else_prim: "prim_recfn 2 r_cons_else"
unfolding r_cons_else_def by simp
lemma r_cons_else:
"eval r_cons_else [e, es] \<down>=
Suc (prod_encode (e_length es, prod_encode (e, pdec2 (es - 1))))"
unfolding r_cons_else_def by simp
definition "r_cons \<equiv>
Cn 2 r_ifz
[Id 2 1, Cn 2 S [Cn 2 r_prod_encode [r_constn 1 0, Id 2 0]], r_cons_else]"
lemma r_cons_prim [simp]: "prim_recfn 2 r_cons"
unfolding r_cons_def using r_cons_else_prim by simp
lemma r_cons [simp]: "eval r_cons [e, es] \<down>= e_cons e es"
unfolding r_cons_def using r_cons_else_prim r_cons_else e_cons_elementary by simp
abbreviation e_snoc :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_snoc es e \<equiv> list_encode (list_decode es @ [e])"
lemma e_nth_snoc_small [simp]:
assumes "n < e_length b"
shows "e_nth (e_snoc b z) n = e_nth b n"
using assms by (simp add: nth_append)
lemma e_hd_snoc [simp]:
assumes "e_length b > 0"
shows "e_hd (e_snoc b x) = e_hd b"
proof -
from assms have "b \<noteq> 0"
using less_imp_neq by force
then have hd: "e_hd b = hd (list_decode b)" by simp
have "e_length (e_snoc b x) > 0" by simp
then have "e_snoc b x \<noteq> 0"
using not_gr_zero by fastforce
then have "e_hd (e_snoc b x) = hd (list_decode (e_snoc b x))" by simp
with assms hd show ?thesis by simp
qed
definition "r_snoc \<equiv> Cn 2 r_rev [Cn 2 r_cons [Id 2 1, Cn 2 r_rev [Id 2 0]]]"
lemma r_snoc_prim [simp]: "prim_recfn 2 r_snoc"
unfolding r_snoc_def by simp
lemma r_snoc [simp]: "eval r_snoc [es, e] \<down>= e_snoc es e"
unfolding r_snoc_def by simp
abbreviation e_butlast :: "nat \<Rightarrow> nat" where
"e_butlast e \<equiv> list_encode (butlast (list_decode e))"
abbreviation e_take :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_take n x \<equiv> list_encode (take n (list_decode x))"
definition "r_take \<equiv>
Cn 2 r_ifle
[Id 2 0, Cn 2 r_length [Id 2 1],
Pr 1 Z (Cn 3 r_snoc [Id 3 1, Cn 3 r_nth [Id 3 2, Id 3 0]]),
Id 2 1]"
lemma r_take_prim [simp]: "prim_recfn 2 r_take"
unfolding r_take_def by simp_all
lemma r_take:
assumes "x = list_encode es"
shows "eval r_take [n, x] \<down>= list_encode (take n es)"
proof -
let ?g = "Cn 3 r_snoc [Id 3 1, Cn 3 r_nth [Id 3 2, Id 3 0]]"
let ?h = "Pr 1 Z ?g"
have "total ?h" using Mn_free_imp_total by simp
have "m \<le> length es \<Longrightarrow> eval ?h [m, x] \<down>= list_encode (take m es)" for m
proof (induction m)
case 0
then show ?case using assms r_take_def by (simp add: r_take_def)
next
case (Suc m)
then have "m < length es" by simp
then have "eval ?h [Suc m, x] = eval ?g [m, the (eval ?h [m, x]), x]"
using Suc r_take_def by simp
also have "... = eval ?g [m, list_encode (take m es), x]"
using Suc by simp
also have "... \<down>= e_snoc (list_encode (take m es)) (es ! m)"
by (simp add: \<open>m < length es\<close> assms)
also have "... \<down>= list_encode ((take m es) @ [es ! m])"
using list_decode_encode by simp
also have "... \<down>= list_encode (take (Suc m) es)"
by (simp add: \<open>m < length es\<close> take_Suc_conv_app_nth)
finally show ?case .
qed
moreover have "eval (Id 2 1) [m, x] \<down>= list_encode (take m es)" if "m > length es" for m
using that assms by simp
moreover have "eval r_take [m, x] \<down>=
(if m \<le> e_length x then the (eval ?h [m, x]) else the (eval (Id 2 1) [m, x]))"
for m
- unfolding r_take_def using `total ?h` by simp
+ unfolding r_take_def using \<open>total ?h\<close> by simp
ultimately show ?thesis unfolding r_take_def by fastforce
qed
corollary r_take' [simp]: "eval r_take [n, x] \<down>= e_take n x"
by (simp add: r_take)
definition "r_last \<equiv> Cn 1 r_hd [r_rev]"
lemma r_last_prim [simp]: "prim_recfn 1 r_last"
unfolding r_last_def by simp
lemma r_last [simp]:
assumes "e = list_encode xs" and "length xs > 0"
shows "eval r_last [e] \<down>= last xs"
proof -
from assms(2) have "length (rev xs) > 0" by simp
then have "list_encode (rev xs) > 0"
by (metis gr0I list.size(3) list_encode_0)
moreover have "eval r_last [e] = eval r_hd [the (eval r_rev [e])]"
unfolding r_last_def by simp
ultimately show ?thesis using assms hd_rev by auto
qed
definition "r_update_aux \<equiv>
let
f = r_constn 2 0;
g = Cn 5 r_snoc
[Id 5 1, Cn 5 r_ifeq [Id 5 0, Id 5 3, Id 5 4, Cn 5 r_nth [Id 5 2, Id 5 0]]]
in Pr 3 f g"
lemma r_update_aux_recfn: "recfn 4 r_update_aux"
unfolding r_update_aux_def by simp
lemma r_update_aux:
assumes "n \<le> e_length b"
shows "eval r_update_aux [n, b, j, v] \<down>= list_encode ((take n (list_decode b))[j:=v])"
using assms
proof (induction n)
case 0
then show ?case unfolding r_update_aux_def by simp
next
case (Suc n)
then have n: "n < e_length b"
by simp
let ?a = "Cn 5 r_nth [Id 5 2, Id 5 0]"
let ?b = "Cn 5 r_ifeq [Id 5 0, Id 5 3, Id 5 4, ?a]"
define g where "g \<equiv> Cn 5 r_snoc [Id 5 1, ?b]"
then have g: "eval g [n, r, b, j, v] \<down>= e_snoc r (if n = j then v else e_nth b n)" for r
by simp
have "Pr 3 (r_constn 2 0) g = r_update_aux"
using r_update_aux_def g_def by simp
then have "eval r_update_aux [Suc n, b, j, v] =
eval g [n, the (eval r_update_aux [n, b, j, v]), b, j, v]"
using r_update_aux_recfn Suc n eval_Pr_converg_Suc
by (metis arity.simps(5) length_Cons list.size(3) nat_less_le
numeral_3_eq_3 option.simps(3))
then have *: "eval r_update_aux [Suc n, b, j, v] \<down>= e_snoc
(list_encode ((take n (list_decode b))[j:=v]))
(if n = j then v else e_nth b n)"
using g Suc by simp
consider (j_eq_n) "j = n" | (j_less_n) "j < n" | (j_gt_n) "j > n"
by linarith
then show ?case
proof (cases)
case j_eq_n
moreover from this have "(take (Suc n) (list_decode b))[j:=v] =
(take n (list_decode b))[j:=v] @ [v]"
using n
by (metis length_list_update nth_list_update_eq take_Suc_conv_app_nth take_update_swap)
ultimately show ?thesis using * by simp
next
case j_less_n
moreover from this have "(take (Suc n) (list_decode b))[j:=v] =
(take n (list_decode b))[j:=v] @ [(list_decode b) ! n]"
using n
by (simp add: le_eq_less_or_eq list_update_append min_absorb2 take_Suc_conv_app_nth)
ultimately show ?thesis using * by auto
next
case j_gt_n
moreover from this have "(take (Suc n) (list_decode b))[j:=v] =
(take n (list_decode b))[j:=v] @ [(list_decode b) ! n]"
using n take_Suc_conv_app_nth by auto
ultimately show ?thesis using * by auto
qed
qed
abbreviation e_update :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_update b j v \<equiv> list_encode ((list_decode b)[j:=v])"
definition "r_update \<equiv>
Cn 3 r_update_aux [Cn 3 r_length [Id 3 0], Id 3 0, Id 3 1, Id 3 2]"
lemma r_update_recfn [simp]: "recfn 3 r_update"
unfolding r_update_def using r_update_aux_recfn by simp
lemma r_update [simp]: "eval r_update [b, j, v] \<down>= e_update b j v"
unfolding r_update_def using r_update_aux r_update_aux_recfn by simp
lemma e_length_update [simp]: "e_length (e_update b k v) = e_length b"
by simp
definition e_append :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_append xs ys \<equiv> list_encode (list_decode xs @ list_decode ys)"
lemma e_length_append: "e_length (e_append xs ys) = e_length xs + e_length ys"
using e_append_def by simp
lemma e_nth_append_small:
assumes "n < e_length xs"
shows "e_nth (e_append xs ys) n = e_nth xs n"
using e_append_def assms by (simp add: nth_append)
lemma e_nth_append_big:
assumes "n \<ge> e_length xs"
shows "e_nth (e_append xs ys) n = e_nth ys (n - e_length xs)"
using e_append_def assms e_nth by (simp add: less_diff_conv2 nth_append)
definition "r_append \<equiv>
let
f = Id 2 0;
g = Cn 4 r_snoc [Id 4 1, Cn 4 r_nth [Id 4 3, Id 4 0]]
in Cn 2 (Pr 2 f g) [Cn 2 r_length [Id 2 1], Id 2 0, Id 2 1]"
lemma r_append_prim [simp]: "prim_recfn 2 r_append"
unfolding r_append_def by simp
lemma r_append [simp]: "eval r_append [a, b] \<down>= e_append a b"
proof -
define g where "g = Cn 4 r_snoc [Id 4 1, Cn 4 r_nth [Id 4 3, Id 4 0]]"
then have g: "eval g [j, r, a, b] \<down>= e_snoc r (e_nth b j)" for j r
by simp
let ?h = "Pr 2 (Id 2 0) g"
have "eval ?h [n, a, b] \<down>= list_encode (list_decode a @ (take n (list_decode b)))"
if "n \<le> e_length b" for n
using that g g_def by (induction n) (simp_all add: take_Suc_conv_app_nth)
then show ?thesis
unfolding r_append_def g_def e_append_def by simp
qed
definition e_append_zeros :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"e_append_zeros b z \<equiv> e_append b (list_encode (replicate z 0))"
lemma e_append_zeros_length: "e_length (e_append_zeros b z) = e_length b + z"
using e_append_def e_append_zeros_def by simp
lemma e_nth_append_zeros: "e_nth (e_append_zeros b z) i = e_nth b i"
using e_append_zeros_def e_nth_append_small e_nth_append_big by auto
lemma e_nth_append_zeros_big:
assumes "i \<ge> e_length b"
shows "e_nth (e_append_zeros b z) i = 0"
unfolding e_append_zeros_def
using e_nth_append_big[of b i "list_encode (replicate z 0)", OF assms(1)]
by simp
definition "r_append_zeros \<equiv>
r_swap (Pr 1 (Id 1 0) (Cn 3 r_snoc [Id 3 1, r_constn 2 0]))"
lemma r_append_zeros_prim [simp]: "prim_recfn 2 r_append_zeros"
unfolding r_append_zeros_def by simp
lemma r_append_zeros: "eval r_append_zeros [b, z] \<down>= e_append_zeros b z"
proof -
let ?r = "Pr 1 (Id 1 0) (Cn 3 r_snoc [Id 3 1, r_constn 2 0])"
have "eval ?r [z, b] \<down>= e_append_zeros b z"
using e_append_zeros_def e_append_def
by (induction z) (simp_all add: replicate_append_same)
then show ?thesis by (simp add: r_append_zeros_def)
qed
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/R1_BC.thy b/thys/Inductive_Inference/R1_BC.thy
--- a/thys/Inductive_Inference/R1_BC.thy
+++ b/thys/Inductive_Inference/R1_BC.thy
@@ -1,534 +1,534 @@
section \<open>@{term "\<R>"} is not in BC\label{s:r1_bc}\<close>
theory R1_BC
imports Lemma_R
CP_FIN_NUM (* for V0 *)
begin
text \<open>We show that @{term "U\<^sub>0 \<union> V\<^sub>0"} is not in BC,
which implies @{term "\<R> \<notin> BC"}.
The proof is by contradiction. Assume there is a strategy $S$ learning @{term
"U\<^sub>0 \<union> V\<^sub>0"} behaviorally correct in the limit with respect to our
standard Gödel numbering $\varphi$. Thanks to Lemma~R for BC we can assume
$S$ to be total. Then we construct a function in @{term "U\<^sub>0 \<union> V\<^sub>0"} for
which $S$ fails.
As usual, there is a computable process building prefixes of functions
$\psi_j$. For every $j$ it starts with the singleton prefix $b = [j]$ and
computes the next prefix from a given prefix $b$ as follows:
\begin{enumerate}
\item Simulate $\varphi_{S(b0^k)}(|b| + k)$ for increasing $k$ for an
increasing number of steps.
\item Once a $k$ with $\varphi_{S(b0^k)}(|b| + k) = 0$ is found, extend the
prefix by $0^k1$.
\end{enumerate}
There is always such a $k$ because by assumption $S$ learns $b0^\infty \in
U_0$ and thus outputs a hypothesis for $b0^\infty$ on almost all of its
prefixes. Therefore for almost all prefixes of the form $b0^k$, we have
$\varphi_{S(b0^k)} = b0^\infty$ and hence $\varphi_{S(b0^k)}(|b| + k) = 0$.
But Step~2 constructs $\psi_j$ such that $\psi_j(|b| + k) = 1$. Therefore $S$
does not hypothesize $\psi_j$ on the prefix $b0^k$ of $\psi_j$. And since the
process runs forever, $S$ outputs infinitely many incorrect hypotheses for
$\psi_j$ and thus does not learn $\psi_j$.
Applying Kleene's fixed-point theorem to @{term "\<psi> \<in> \<R>\<^sup>2"}
yields a $j$ with $\varphi_j = \psi_j$ and thus $\psi_j \in V_0$. But $S$
does not learn any $\psi_j$, contradicting our assumption.
The result @{prop "\<R> \<notin> BC"} can be obtained more directly by
running the process with the empty prefix, thereby constructing only one
function instead of a numbering. This function is in @{term R1}, and $S$
fails to learn it by the same reasoning as above. The stronger statement
about @{term "U\<^sub>0 \<union> V\<^sub>0"} will be exploited in
Section~\ref{s:union}.
In the following locale the assumption that $S$ learns @{term "U\<^sub>0"}
suffices for analyzing the process. However, in order to arrive at the
desired contradiction this assumption is too weak because the functions built
by the process are not in @{term "U\<^sub>0"}.\<close>
locale r1_bc =
fixes s :: partial1
assumes s_in_R1: "s \<in> \<R>" and s_learn_U0: "learn_bc \<phi> U\<^sub>0 s"
begin
lemma s_learn_prenum: "\<And>b. learn_bc \<phi> {prenum b} s"
using s_learn_U0 U0_altdef learn_bc_closed_subseteq by blast
text \<open>A @{typ recf} for the strategy:\<close>
definition r_s :: recf where
"r_s \<equiv> SOME rs. recfn 1 rs \<and> total rs \<and> s = (\<lambda>x. eval rs [x])"
lemma r_s_recfn [simp]: "recfn 1 r_s"
and r_s_total: "\<And>x. eval r_s [x] \<down>"
and eval_r_s: "\<And>x. s x = eval r_s [x]"
using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all
text \<open>We begin with the function that finds the $k$ from Step~1 of the
construction of $\psi$.\<close>
definition "r_find_k \<equiv>
let k = Cn 2 r_pdec1 [Id 2 0];
r = Cn 2 r_result1
[Cn 2 r_pdec2 [Id 2 0],
Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, k]],
Cn 2 r_add [Cn 2 r_length [Id 2 1], k]]
in Cn 1 r_pdec1 [Mn 1 (Cn 2 r_eq [r, r_constn 1 1])]"
lemma r_find_k_recfn [simp]: "recfn 1 r_find_k"
unfolding r_find_k_def by (simp add: Let_def)
text \<open>There is always a suitable $k$, since the strategy learns
$b0^\infty$ for all $b$.\<close>
lemma learn_bc_prenum_eventually_zero:
"\<exists>k. \<phi> (the (s (e_append_zeros b k))) (e_length b + k) \<down>= 0"
proof -
let ?f = "prenum b"
have "\<exists>n\<ge>e_length b. \<phi> (the (s (?f \<triangleright> n))) = ?f"
using learn_bcE s_learn_prenum by (meson le_cases singletonI)
then obtain n where n: "n \<ge> e_length b" "\<phi> (the (s (?f \<triangleright> n))) = ?f"
by auto
define k where "k = Suc n - e_length b"
let ?e = "e_append_zeros b k"
have len: "e_length ?e = Suc n"
using k_def n e_append_zeros_length by simp
have "?f \<triangleright> n = ?e"
proof -
have "e_length ?e > 0"
using len n(1) by simp
moreover have "?f x \<down>= e_nth ?e x" for x
proof (cases "x < e_length b")
case True
then show ?thesis using e_nth_append_zeros by simp
next
case False
then have "?f x \<down>= 0" by simp
moreover from False have "e_nth ?e x = 0"
using e_nth_append_zeros_big by simp
ultimately show ?thesis by simp
qed
ultimately show ?thesis using initI[of "?e"] len by simp
qed
with n(2) have "\<phi> (the (s ?e)) = ?f" by simp
then have "\<phi> (the (s ?e)) (e_length ?e) \<down>= 0"
using len n(1) by auto
then show ?thesis using e_append_zeros_length by auto
qed
lemma if_eq_eq: "(if v = 1 then (0 :: nat) else 1) = 0 \<Longrightarrow> v = 1"
by presburger
lemma r_find_k:
shows "eval r_find_k [b] \<down>"
and "let k = the (eval r_find_k [b])
in \<phi> (the (s (e_append_zeros b k))) (e_length b + k) \<down>= 0"
proof -
let ?k = "Cn 2 r_pdec1 [Id 2 0]"
let ?argt = "Cn 2 r_pdec2 [Id 2 0]"
let ?argi = "Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, ?k]]"
let ?argx = "Cn 2 r_add [Cn 2 r_length [Id 2 1], ?k]"
let ?r = "Cn 2 r_result1 [?argt, ?argi, ?argx]"
define f where "f \<equiv>
let k = Cn 2 r_pdec1 [Id 2 0];
r = Cn 2 r_result1
[Cn 2 r_pdec2 [Id 2 0],
Cn 2 r_s [Cn 2 r_append_zeros [Id 2 1, k]],
Cn 2 r_add [Cn 2 r_length [Id 2 1], k]]
in Cn 2 r_eq [r, r_constn 1 1]"
then have "recfn 2 f" by (simp add: Let_def)
have "total r_s"
by (simp add: r_s_total totalI1)
then have "total f"
unfolding f_def using Cn_total Mn_free_imp_total by (simp add: Let_def)
have "eval ?argi [z, b] = s (e_append_zeros b (pdec1 z))" for z
using r_append_zeros \<open>recfn 2 f\<close> eval_r_s by auto
then have "eval ?argi [z, b] \<down>= the (s (e_append_zeros b (pdec1 z)))" for z
using eval_r_s r_s_total by simp
moreover have "recfn 2 ?r" using \<open>recfn 2 f\<close> by auto
ultimately have r: "eval ?r [z, b] =
eval r_result1 [pdec2 z, the (s (e_append_zeros b (pdec1 z))), e_length b + pdec1 z]"
for z
by simp
then have f: "eval f [z, b] \<down>= (if the (eval ?r [z, b]) = 1 then 0 else 1)" for z
- using f_def `recfn 2 f` prim_recfn_total by (auto simp add: Let_def)
+ using f_def \<open>recfn 2 f\<close> prim_recfn_total by (auto simp add: Let_def)
have "\<exists>k. \<phi> (the (s (e_append_zeros b k))) (e_length b + k) \<down>= 0"
using s_learn_prenum learn_bc_prenum_eventually_zero by auto
then obtain k where "\<phi> (the (s (e_append_zeros b k))) (e_length b + k) \<down>= 0"
by auto
then obtain t where "eval r_result1 [t, the (s (e_append_zeros b k)), e_length b + k] \<down>= Suc 0"
using r_result1_converg_phi(1) by blast
then have t: "eval r_result1 [t, the (s (e_append_zeros b k)), e_length b + k] \<down>= Suc 0"
by simp
let ?z = "prod_encode (k, t)"
have "eval ?r [?z, b] \<down>= Suc 0"
using t r by (metis fst_conv prod_encode_inverse snd_conv)
with f have fzb: "eval f [?z, b] \<down>= 0" by simp
moreover have "eval (Mn 1 f) [b] =
(if (\<exists>z. eval f ([z, b]) \<down>= 0)
then Some (LEAST z. eval f [z, b] \<down>= 0)
else None)"
- using eval_Mn_total[of 1 f "[b]"] `total f` `recfn 2 f` by simp
+ using eval_Mn_total[of 1 f "[b]"] \<open>total f\<close> \<open>recfn 2 f\<close> by simp
ultimately have mn1f: "eval (Mn 1 f) [b] \<down>= (LEAST z. eval f [z, b] \<down>= 0)"
by auto
with fzb have "eval f [the (eval (Mn 1 f) [b]), b] \<down>= 0" (is "eval f [?zz, b] \<down>= 0")
using \<open>total f\<close> \<open>recfn 2 f\<close> LeastI_ex[of "%z. eval f [z, b] \<down>= 0"] by auto
moreover have "eval f [?zz, b] \<down>= (if the (eval ?r [?zz, b]) = 1 then 0 else 1)"
using f by simp
ultimately have "(if the (eval ?r [?zz, b]) = 1 then (0 :: nat) else 1) = 0" by auto
then have "the (eval ?r [?zz, b]) = 1"
using if_eq_eq[of "the (eval ?r [?zz, b])"] by simp
then have
"eval r_result1
[pdec2 ?zz, the (s (e_append_zeros b (pdec1 ?zz))), e_length b + pdec1 ?zz] \<down>=
1"
using r r_result1_total r_result1_prim totalE
by (metis length_Cons list.size(3) numeral_3_eq_3 option.collapse)
then have *: "\<phi> (the (s (e_append_zeros b (pdec1 ?zz)))) (e_length b + pdec1 ?zz) \<down>= 0"
by (simp add: r_result1_some_phi)
define Mn1f where "Mn1f = Mn 1 f"
then have "eval Mn1f [b] \<down>= ?zz"
using mn1f by auto
moreover have "recfn 1 (Cn 1 r_pdec1 [Mn1f])"
- using `recfn 2 f` Mn1f_def by simp
+ using \<open>recfn 2 f\<close> Mn1f_def by simp
ultimately have "eval (Cn 1 r_pdec1 [Mn1f]) [b] = eval r_pdec1 [the (eval (Mn1f) [b])]"
by auto
then have "eval (Cn 1 r_pdec1 [Mn1f]) [b] = eval r_pdec1 [?zz]"
using Mn1f_def by blast
then have 1: "eval (Cn 1 r_pdec1 [Mn1f]) [b] \<down>= pdec1 ?zz"
by simp
moreover have "recfn 1 (Cn 1 S [Cn 1 r_pdec1 [Mn1f]])"
- using `recfn 2 f` Mn1f_def by simp
+ using \<open>recfn 2 f\<close> Mn1f_def by simp
ultimately have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] =
eval S [the (eval (Cn 1 r_pdec1 [Mn1f]) [b])]"
by simp
then have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] = eval S [pdec1 ?zz]"
using 1 by simp
then have "eval (Cn 1 S [Cn 1 r_pdec1 [Mn1f]]) [b] \<down>= Suc (pdec1 ?zz)"
by simp
moreover have "eval r_find_k [b] = eval (Cn 1 r_pdec1 [Mn1f]) [b]"
unfolding r_find_k_def Mn1f_def f_def by metis
ultimately have r_find_ksb: "eval r_find_k [b] \<down>= pdec1 ?zz"
using 1 by simp
then show "eval r_find_k [b] \<down>" by simp_all
from r_find_ksb have "the (eval r_find_k [b]) = pdec1 ?zz"
by simp
moreover have "\<phi> (the (s (e_append_zeros b (pdec1 ?zz)))) (e_length b + pdec1 ?zz) \<down>= 0"
using * by simp
ultimately show "let k = the (eval r_find_k [b])
in \<phi> (the (s (e_append_zeros b k))) (e_length b + k) \<down>= 0"
by simp
qed
lemma r_find_k_total: "total r_find_k"
by (simp add: s_learn_prenum r_find_k(1) totalI1)
text \<open>The following function represents one iteration of the
process.\<close>
abbreviation "r_next \<equiv>
Cn 3 r_snoc [Cn 3 r_append_zeros [Id 3 1, Cn 3 r_find_k [Id 3 1]], r_constn 2 1]"
text \<open>Using @{term r_next} we define the function @{term r_prefixes}
that computes the prefix after every iteration of the process.\<close>
definition r_prefixes :: recf where
"r_prefixes \<equiv> Pr 1 r_singleton_encode r_next"
lemma r_prefixes_recfn: "recfn 2 r_prefixes"
unfolding r_prefixes_def by simp
lemma r_prefixes_total: "total r_prefixes"
proof -
have "recfn 3 r_next" by simp
then have "total r_next"
- using `recfn 3 r_next` r_find_k_total Cn_total Mn_free_imp_total by auto
+ using \<open>recfn 3 r_next\<close> r_find_k_total Cn_total Mn_free_imp_total by auto
then show ?thesis
by (simp add: Mn_free_imp_total Pr_total r_prefixes_def)
qed
lemma r_prefixes_0: "eval r_prefixes [0, j] \<down>= list_encode [j]"
unfolding r_prefixes_def by simp
lemma r_prefixes_Suc:
"eval r_prefixes [Suc n, j] \<down>=
(let b = the (eval r_prefixes [n, j])
in e_snoc (e_append_zeros b (the (eval r_find_k [b]))) 1)"
proof -
have "recfn 3 r_next" by simp
then have "total r_next"
- using `recfn 3 r_next` r_find_k_total Cn_total Mn_free_imp_total by auto
+ using \<open>recfn 3 r_next\<close> r_find_k_total Cn_total Mn_free_imp_total by auto
have eval_next: "eval r_next [t, v, j] \<down>=
e_snoc (e_append_zeros v (the (eval r_find_k [v]))) 1"
for t v j
- using r_find_k_total `recfn 3 r_next` r_append_zeros by simp
+ using r_find_k_total \<open>recfn 3 r_next\<close> r_append_zeros by simp
then have "eval r_prefixes [Suc n, j] = eval r_next [n, the (eval r_prefixes [n, j]), j]"
using r_prefixes_total by (simp add: r_prefixes_def)
then show "eval r_prefixes [Suc n, j] \<down>=
(let b = the (eval r_prefixes [n, j])
in e_snoc (e_append_zeros b (the (eval r_find_k [b]))) 1)"
using eval_next by metis
qed
text \<open>Since @{term r_prefixes} is total, we can get away with
introducing a total function.\<close>
definition prefixes :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"prefixes j t \<equiv> the (eval r_prefixes [t, j])"
lemma prefixes_Suc:
"prefixes j (Suc t) =
e_snoc (e_append_zeros (prefixes j t) (the (eval r_find_k [prefixes j t]))) 1"
unfolding prefixes_def using r_prefixes_Suc by (simp_all add: Let_def)
lemma prefixes_Suc_length:
"e_length (prefixes j (Suc t)) =
Suc (e_length (prefixes j t) + the (eval r_find_k [prefixes j t]))"
using e_append_zeros_length prefixes_Suc by simp
lemma prefixes_length_mono: "e_length (prefixes j t) < e_length (prefixes j (Suc t))"
using prefixes_Suc_length by simp
lemma prefixes_length_mono': "e_length (prefixes j t) \<le> e_length (prefixes j (t + d))"
proof (induction d)
case 0
then show ?case by simp
next
case (Suc d)
then show ?case using prefixes_length_mono le_less_trans by fastforce
qed
lemma prefixes_length_lower_bound: "e_length (prefixes j t) \<ge> Suc t"
proof (induction t)
case 0
then show ?case by (simp add: prefixes_def r_prefixes_0)
next
case (Suc t)
moreover have "Suc (e_length (prefixes j t)) \<le> e_length (prefixes j (Suc t))"
using prefixes_length_mono by (simp add: Suc_leI)
ultimately show ?case by simp
qed
lemma prefixes_Suc_nth:
assumes "x < e_length (prefixes j t)"
shows "e_nth (prefixes j t) x = e_nth (prefixes j (Suc t)) x"
proof -
define k where "k = the (eval r_find_k [prefixes j t])"
let ?u = "e_append_zeros (prefixes j t) k"
have "prefixes j (Suc t) =
e_snoc (e_append_zeros (prefixes j t) (the (eval r_find_k [prefixes j t]))) 1"
using prefixes_Suc by simp
with k_def have "prefixes j (Suc t) = e_snoc ?u 1"
by simp
then have "e_nth (prefixes j (Suc t)) x = e_nth (e_snoc ?u 1) x"
by simp
moreover have "x < e_length ?u"
using assms e_append_zeros_length by auto
ultimately have "e_nth (prefixes j (Suc t)) x = e_nth ?u x"
using e_nth_snoc_small by simp
moreover have "e_nth ?u x = e_nth (prefixes j t) x"
using assms e_nth_append_zeros by simp
ultimately show "e_nth (prefixes j t) x = e_nth (prefixes j (Suc t)) x"
by simp
qed
lemma prefixes_Suc_last: "e_nth (prefixes j (Suc t)) (e_length (prefixes j (Suc t)) - 1) = 1"
using prefixes_Suc by simp
lemma prefixes_le_nth:
assumes "x < e_length (prefixes j t)"
shows "e_nth (prefixes j t) x = e_nth (prefixes j (t + d)) x"
proof (induction d)
case 0
then show ?case by simp
next
case (Suc d)
have "x < e_length (prefixes j (t + d))"
using s_learn_prenum assms prefixes_length_mono'
by (simp add: less_eq_Suc_le order_trans_rules(23))
then have "e_nth (prefixes j (t + d)) x = e_nth (prefixes j (t + Suc d)) x"
using prefixes_Suc_nth by simp
with Suc show ?case by simp
qed
text \<open>The numbering $\psi$ is defined via @{term[names_short] prefixes}.\<close>
definition psi :: partial2 ("\<psi>") where
"\<psi> j x \<equiv> Some (e_nth (prefixes j (Suc x)) x)"
lemma psi_in_R2: "\<psi> \<in> \<R>\<^sup>2"
proof
define r where "r \<equiv> Cn 2 r_nth [Cn 2 r_prefixes [Cn 2 S [Id 2 1], Id 2 0], Id 2 1]"
then have "recfn 2 r"
using r_prefixes_recfn by simp
then have "eval r [j, x] \<down>= e_nth (prefixes j (Suc x)) x" for j x
unfolding r_def prefixes_def using r_prefixes_total r_prefixes_recfn e_nth by simp
then have "eval r [j, x] = \<psi> j x" for j x
unfolding psi_def by simp
then show "\<psi> \<in> \<P>\<^sup>2"
- using `recfn 2 r` by auto
+ using \<open>recfn 2 r\<close> by auto
show "total2 \<psi>"
unfolding psi_def by auto
qed
lemma psi_eq_nth_prefixes:
assumes "x < e_length (prefixes j t)"
shows "\<psi> j x \<down>= e_nth (prefixes j t) x"
proof (cases "Suc x < t")
case True
have "x \<le> e_length (prefixes j x)"
using prefixes_length_lower_bound by (simp add: Suc_leD)
also have "... < e_length (prefixes j (Suc x))"
using prefixes_length_mono s_learn_prenum by simp
finally have "x < e_length (prefixes j (Suc x))" .
with True have "e_nth (prefixes j (Suc x)) x = e_nth (prefixes j t) x"
using prefixes_le_nth[of x j "Suc x" "t - Suc x"] by simp
then show ?thesis using psi_def by simp
next
case False
then have "e_nth (prefixes j (Suc x)) x = e_nth (prefixes j t) x"
using prefixes_le_nth[of x j t "Suc x - t"] assms by simp
then show ?thesis using psi_def by simp
qed
lemma psi_at_0: "\<psi> j 0 \<down>= j"
using psi_eq_nth_prefixes[of 0 j 0] prefixes_length_lower_bound[of 0 j]
by (simp add: prefixes_def r_prefixes_0)
text \<open>The prefixes output by the process @{term[names_short] "prefixes j"} are
indeed prefixes of $\psi_j$.\<close>
lemma prefixes_init_psi: "\<psi> j \<triangleright> (e_length (prefixes j (Suc t)) - 1) = prefixes j (Suc t)"
proof (rule initI[of "prefixes j (Suc t)"])
let ?e = "prefixes j (Suc t)"
show "e_length ?e > 0"
using prefixes_length_lower_bound[of "Suc t" j] by auto
show "\<And>x. x < e_length ?e \<Longrightarrow> \<psi> j x \<down>= e_nth ?e x"
using prefixes_Suc_nth psi_eq_nth_prefixes by simp
qed
text \<open>Every prefix of $\psi_j$ generated by the process
@{term[names_short] "prefixes j"} (except for the initial one) is of the form
$b0^k1$. But $k$ is chosen such that $\varphi_{S(b0^k)}(|b|+k) = 0 \neq 1 =
b0^k1_{|b|+k}$. Therefore the hypothesis $S(b0^k)$ is incorrect for
$\psi_j$.\<close>
lemma hyp_wrong_at_last:
"\<phi> (the (s (e_butlast (prefixes j (Suc t))))) (e_length (prefixes j (Suc t)) - 1) \<noteq>
\<psi> j (e_length (prefixes j (Suc t)) - 1)"
(is "?lhs \<noteq> ?rhs")
proof -
let ?b = "prefixes j t"
let ?k = "the (eval r_find_k [?b])"
let ?x = "e_length (prefixes j (Suc t)) - 1"
have "e_butlast (prefixes j (Suc t)) = e_append_zeros ?b ?k"
using s_learn_prenum prefixes_Suc by simp
then have "?lhs = \<phi> (the (s (e_append_zeros ?b ?k))) ?x"
by simp
moreover have "?x = e_length ?b + ?k"
using prefixes_Suc_length by simp
ultimately have "?lhs = \<phi> (the (s (e_append_zeros ?b ?k))) (e_length ?b + ?k)"
by simp
then have "?lhs \<down>= 0"
using r_find_k(2) r_s_total s_learn_prenum by metis
moreover have "?x < e_length (prefixes j (Suc t))"
using prefixes_length_lower_bound le_less_trans linorder_not_le s_learn_prenum
by fastforce
ultimately have "?rhs \<down>= e_nth (prefixes j (Suc t)) ?x"
using psi_eq_nth_prefixes[of ?x j "Suc t"] by simp
moreover have "e_nth (prefixes j (Suc t)) ?x = 1"
using prefixes_Suc prefixes_Suc_last by simp
ultimately have "?rhs \<down>= 1" by simp
- with `?lhs \<down>= 0` show ?thesis by simp
+ with \<open>?lhs \<down>= 0\<close> show ?thesis by simp
qed
corollary hyp_wrong: "\<phi> (the (s (e_butlast (prefixes j (Suc t))))) \<noteq> \<psi> j"
using hyp_wrong_at_last[of j t] by auto
text \<open>For all $j$, the strategy $S$ outputs infinitely many wrong hypotheses for
$\psi_j$\<close>
lemma infinite_hyp_wrong: "\<exists>m>n. \<phi> (the (s (\<psi> j \<triangleright> m))) \<noteq> \<psi> j"
proof -
let ?b = "prefixes j (Suc (Suc n))"
let ?bb = "e_butlast ?b"
have len_b: "e_length ?b > Suc (Suc n)"
using prefixes_length_lower_bound by (simp add: Suc_le_lessD)
then have len_bb: "e_length ?bb > Suc n" by simp
define m where "m = e_length ?bb - 1"
with len_bb have "m > n" by simp
have "\<psi> j \<triangleright> m = ?bb"
proof -
have "\<psi> j \<triangleright> (e_length ?b - 1) = ?b"
using prefixes_init_psi by simp
then have "\<psi> j \<triangleright> (e_length ?b - 2) = ?bb"
using init_butlast_init psi_in_R2 R2_proj_R1 R1_imp_total1 len_bb length_init
by (metis Suc_1 diff_diff_left length_butlast length_greater_0_conv
list.size(3) list_decode_encode not_less0 plus_1_eq_Suc)
then show ?thesis by (metis diff_Suc_1 length_init m_def)
qed
moreover have "\<phi> (the (s ?bb)) \<noteq> \<psi> j"
using hyp_wrong by simp
ultimately have "\<phi> (the (s (\<psi> j \<triangleright> m))) \<noteq> \<psi> j"
by simp
- with `m > n` show ?thesis by auto
+ with \<open>m > n\<close> show ?thesis by auto
qed
lemma U0_V0_not_learn_bc: "\<not> learn_bc \<phi> (U\<^sub>0 \<union> V\<^sub>0) s"
proof -
obtain j where j: "\<phi> j = \<psi> j"
using R2_imp_P2 kleene_fixed_point psi_in_R2 by blast
moreover have "\<exists>m>n. \<phi> (the (s ((\<psi> j) \<triangleright> m))) \<noteq> \<psi> j" for n
using infinite_hyp_wrong[of _ j] by simp
ultimately have "\<not> learn_bc \<phi> {\<psi> j} s"
using infinite_hyp_wrong_not_BC by simp
moreover have "\<psi> j \<in> V\<^sub>0"
proof -
have "\<psi> j \<in> \<R>" (is "?f \<in> \<R>")
using psi_in_R2 by simp
moreover have "\<phi> (the (?f 0)) = ?f"
using j psi_at_0[of j] by simp
ultimately show ?thesis by (simp add: V0_def)
qed
ultimately show "\<not> learn_bc \<phi> (U\<^sub>0 \<union> V\<^sub>0) s"
using learn_bc_closed_subseteq by auto
qed
end
lemma U0_V0_not_in_BC: "U\<^sub>0 \<union> V\<^sub>0 \<notin> BC"
proof
assume in_BC: "U\<^sub>0 \<union> V\<^sub>0 \<in> BC"
then have "U\<^sub>0 \<union> V\<^sub>0 \<in> BC_wrt \<phi>"
using BC_wrt_phi_eq_BC by simp
then obtain s where "learn_bc \<phi> (U\<^sub>0 \<union> V\<^sub>0) s"
using BC_wrt_def by auto
then obtain s' where s': "s' \<in> \<R>" "learn_bc \<phi> (U\<^sub>0 \<union> V\<^sub>0) s'"
using lemma_R_for_BC_simple by blast
then have learn_U0: "learn_bc \<phi> U\<^sub>0 s'"
using learn_bc_closed_subseteq[of \<phi> "U\<^sub>0 \<union> V\<^sub>0" "s'"] by simp
then interpret r1_bc s'
by (simp add: r1_bc_def s'(1))
have "\<not> learn_bc \<phi> (U\<^sub>0 \<union> V\<^sub>0) s'"
using learn_bc_closed_subseteq U0_V0_not_learn_bc by simp
with s'(2) show False by simp
qed
theorem R1_not_in_BC: "\<R> \<notin> BC"
proof -
have "U\<^sub>0 \<union> V\<^sub>0 \<subseteq> \<R>"
using V0_def U0_in_NUM by auto
then show ?thesis
using U0_V0_not_in_BC BC_closed_subseteq by auto
qed
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/Standard_Results.thy b/thys/Inductive_Inference/Standard_Results.thy
--- a/thys/Inductive_Inference/Standard_Results.thy
+++ b/thys/Inductive_Inference/Standard_Results.thy
@@ -1,1595 +1,1595 @@
theory Standard_Results
imports Universal
begin
section \<open>Kleene normal form and the number of $\mu$-operations\<close>
text \<open>Kleene's original normal form theorem~\cite{Kleene43} states that
every partial recursive $f$ can be expressed as $f(x) = u(\mu y[t(i, x, y) =
0]$ for some $i$, where $u$ and $t$ are specially crafted primitive recursive
functions tied to Kleene's definition of partial recursive functions.
Rogers~\cite[p.~29f.]{Rogers87} relaxes the theorem by allowing $u$ and $t$
to be any primitive recursive functions of arity one and three, respectively.
Both versions require a separate $t$-predicate for every arity. We will show
a unified version for all arities by treating $x$ as an encoded list of
arguments.
Our universal function @{thm[display,names_short] "r_univ_def"} can represent
all partial recursive functions (see theorem @{thm[source] r_univ}). Moreover
@{term "r_result"}, @{term "r_dec"}, and @{term "r_not"} are primitive
recursive. As such @{term r_univ} could almost serve as the right-hand side
$u(\mu y[t(i, x, y) = 0]$. Its only flaw is that the outer function, the
composition of @{term r_dec} and @{term r_result}, is ternary rather than
unary.\<close>
lemma r_univ_almost_kleene_nf:
"r_univ \<simeq>
(let u = Cn 3 r_dec [r_result];
t = Cn 3 r_not [r_result]
in Cn 2 u [Mn 2 t, Id 2 0, Id 2 1])"
unfolding r_univ_def by (rule exteqI) simp_all
text \<open>We can remedy the wrong arity with some encoding and
projecting.\<close>
definition r_nf_t :: recf where
"r_nf_t \<equiv> Cn 3 r_and
[Cn 3 r_eq [Cn 3 r_pdec2 [Id 3 0], Cn 3 r_prod_encode [Id 3 1, Id 3 2]],
Cn 3 r_not
[Cn 3 r_result
[Cn 3 r_pdec1 [Id 3 0],
Cn 3 r_pdec12 [Id 3 0],
Cn 3 r_pdec22 [Id 3 0]]]]"
lemma r_nf_t_prim: "prim_recfn 3 r_nf_t"
unfolding r_nf_t_def by simp
definition r_nf_u :: recf where
"r_nf_u \<equiv> Cn 1 r_dec [Cn 1 r_result [r_pdec1, r_pdec12, r_pdec22]]"
lemma r_nf_u_prim: "prim_recfn 1 r_nf_u"
unfolding r_nf_u_def by simp
lemma r_nf_t_0:
assumes "eval r_result [pdec1 y, pdec12 y, pdec22 y] \<down>\<noteq> 0"
and "pdec2 y = prod_encode (i, x)"
shows "eval r_nf_t [y, i, x] \<down>= 0"
unfolding r_nf_t_def using assms by auto
lemma r_nf_t_1:
assumes "eval r_result [pdec1 y, pdec12 y, pdec22 y] \<down>= 0 \<or> pdec2 y \<noteq> prod_encode (i, x)"
shows "eval r_nf_t [y, i, x] \<down>= 1"
unfolding r_nf_t_def using assms r_result_total by auto
text \<open>The next function is just as universal as @{term r_univ}, but
satisfies the conditions of the Kleene normal form theorem because the
outer funtion @{term r_nf_u} is unary.\<close>
definition "r_normal_form \<equiv> Cn 2 r_nf_u [Mn 2 r_nf_t]"
lemma r_normal_form_recfn: "recfn 2 r_normal_form"
unfolding r_normal_form_def using r_nf_u_prim r_nf_t_prim by simp
lemma r_univ_exteq_r_normal_form: "r_univ \<simeq> r_normal_form"
proof (rule exteqI)
show arity: "arity r_univ = arity r_normal_form"
using r_normal_form_recfn by simp
show "eval r_univ xs = eval r_normal_form xs" if "length xs = arity r_univ" for xs
proof -
have "length xs = 2"
using that by simp
then obtain i x where ix: "[i, x] = xs"
by (smt Suc_length_conv length_0_conv numeral_2_eq_2)
have "eval r_univ [i, x] = eval r_normal_form [i, x]"
proof (cases "\<forall>t. eval r_result [t, i, x] \<down>= 0")
case True
then have "eval r_univ [i, x] \<up>"
unfolding r_univ_def by simp
moreover have "eval r_normal_form [i, x] \<up>"
proof -
have "eval r_nf_t [y, i, x] \<down>= 1" for y
using True r_nf_t_1[of y i x] by fastforce
then show ?thesis
unfolding r_normal_form_def using r_nf_u_prim r_nf_t_prim by simp
qed
ultimately show ?thesis by simp
next
case False
then have "\<exists>t. eval r_result [t, i, x] \<down>\<noteq> 0"
by (simp add: r_result_total)
then obtain t where "eval r_result [t, i, x] \<down>\<noteq> 0"
by auto
then have "eval r_nf_t [triple_encode t i x, i, x] \<down>= 0"
using r_nf_t_0 by simp
then obtain y where y: "eval (Mn 2 r_nf_t) [i, x] \<down>= y"
using r_nf_t_prim Mn_free_imp_total by fastforce
then have "eval r_nf_t [y, i, x] \<down>= 0"
using r_nf_t_prim Mn_free_imp_total eval_Mn_convergE(2)[of 2 r_nf_t "[i, x]" y]
by simp
then have r_result: "eval r_result [pdec1 y, pdec12 y, pdec22 y] \<down>\<noteq> 0"
and pdec2: "pdec2 y = prod_encode (i, x)"
using r_nf_t_0[of y i x] r_nf_t_1[of y i x] r_result_total by auto
then have "eval r_result [pdec1 y, i, x] \<down>\<noteq> 0"
by simp
then obtain v where v:
"eval r_univ [pdec12 y, pdec22 y] \<down>= v"
"eval r_result [pdec1 y, pdec12 y, pdec22 y] \<down>= Suc v"
using r_result r_result_bivalent'[of "pdec12 y" "pdec22 y" _ "pdec1 y"]
r_result_diverg'[of "pdec12 y" "pdec22 y" "pdec1 y"]
by auto
have "eval r_normal_form [i, x] = eval r_nf_u [y]"
unfolding r_normal_form_def using y r_nf_t_prim r_nf_u_prim by simp
also have "... = eval r_dec [the (eval (Cn 1 r_result [r_pdec1, r_pdec12, r_pdec22]) [y])]"
unfolding r_nf_u_def using r_result by simp
also have "... = eval r_dec [Suc v]"
using v by simp
also have "... \<down>= v"
by simp
finally have "eval r_normal_form [i, x] \<down>= v" .
moreover have "eval r_univ [i, x] \<down>= v"
using v(1) pdec2 by simp
ultimately show ?thesis by simp
qed
with ix show ?thesis by simp
qed
qed
theorem normal_form:
assumes "recfn n f"
obtains i where "\<forall>x. e_length x = n \<longrightarrow> eval r_normal_form [i, x] = eval f (list_decode x)"
proof -
have "eval r_normal_form [encode f, x] = eval f (list_decode x)" if "e_length x = n" for x
using r_univ_exteq_r_normal_form assms that exteq_def r_univ' by auto
then show ?thesis using that by auto
qed
text \<open>As a consequence of the normal form theorem every partial
recursive function can be represented with exactly one application of the
$\mu$-operator.\<close>
fun count_Mn :: "recf \<Rightarrow> nat" where
"count_Mn Z = 0"
| "count_Mn S = 0"
| "count_Mn (Id m n) = 0"
| "count_Mn (Cn n f gs) = count_Mn f + sum_list (map count_Mn gs)"
| "count_Mn (Pr n f g) = count_Mn f + count_Mn g"
| "count_Mn (Mn n f) = Suc (count_Mn f)"
lemma count_Mn_zero_iff_prim: "count_Mn f = 0 \<longleftrightarrow> Mn_free f"
by (induction f) auto
text \<open>The normal form has only one $\mu$-recursion.\<close>
lemma count_Mn_normal_form: "count_Mn r_normal_form = 1"
unfolding r_normal_form_def r_nf_u_def r_nf_t_def using count_Mn_zero_iff_prim by simp
lemma one_Mn_suffices:
assumes "recfn n f"
shows "\<exists>g. count_Mn g = 1 \<and> g \<simeq> f"
proof -
have "n > 0"
using assms wellf_arity_nonzero by auto
obtain i where i:
"\<forall>x. e_length x = n \<longrightarrow> eval r_normal_form [i, x] = eval f (list_decode x)"
using normal_form[OF assms(1)] by auto
define g where "g \<equiv> Cn n r_normal_form [r_constn (n - 1) i, r_list_encode (n - 1)]"
then have "recfn n g"
- using r_normal_form_recfn `n > 0` by simp
+ using r_normal_form_recfn \<open>n > 0\<close> by simp
then have "g \<simeq> f"
using g_def r_list_encode i assms by (intro exteqI) simp_all
moreover have "count_Mn g = 1"
unfolding g_def using count_Mn_normal_form count_Mn_zero_iff_prim by simp
ultimately show ?thesis by auto
qed
text \<open>The previous lemma could have been obtained without @{term
"r_normal_form"} directly from @{term "r_univ"}.\<close>
section \<open>The $s$-$m$-$n$ theorem\<close>
text \<open>For all $m, n > 0$ there is an $(m + 1)$-ary primitive recursive
function $s^m_n$ with
\[
\varphi_p^{(m + n)}(c_1, \dots,c_m, x_1, \dots, x_n) =
\varphi_{s^m_n(p, c_1, \dots,c_m)}^{(n)}(x_1, \dots, x_n)
\]
for all $p, c_1, \ldots, c_m, x_1, \ldots, x_n$. Here, $\varphi^{(n)}$ is a
function universal for $n$-ary partial recursive functions, which we will
represent by @{term "r_universal n"}\<close>
text \<open>The $s^m_n$ functions compute codes of functions. We start simple:
computing codes of the unary constant functions.\<close>
fun code_const1 :: "nat \<Rightarrow> nat" where
"code_const1 0 = 0"
| "code_const1 (Suc c) = quad_encode 3 1 1 (singleton_encode (code_const1 c))"
lemma code_const1: "code_const1 c = encode (r_const c)"
by (induction c) simp_all
definition "r_code_const1_aux \<equiv>
Cn 3 r_prod_encode
[r_constn 2 3,
Cn 3 r_prod_encode
[r_constn 2 1,
Cn 3 r_prod_encode
[r_constn 2 1, Cn 3 r_singleton_encode [Id 3 1]]]]"
lemma r_code_const1_aux_prim: "prim_recfn 3 r_code_const1_aux"
by (simp_all add: r_code_const1_aux_def)
lemma r_code_const1_aux:
"eval r_code_const1_aux [i, r, c] \<down>= quad_encode 3 1 1 (singleton_encode r)"
by (simp add: r_code_const1_aux_def)
definition "r_code_const1 \<equiv> r_shrink (Pr 1 Z r_code_const1_aux)"
lemma r_code_const1_prim: "prim_recfn 1 r_code_const1"
by (simp_all add: r_code_const1_def r_code_const1_aux_prim)
lemma r_code_const1: "eval r_code_const1 [c] \<down>= code_const1 c"
proof -
let ?h = "Pr 1 Z r_code_const1_aux"
have "eval ?h [c, x] \<down>= code_const1 c" for x
using r_code_const1_aux r_code_const1_def
by (induction c) (simp_all add: r_code_const1_aux_prim)
then show ?thesis by (simp add: r_code_const1_def r_code_const1_aux_prim)
qed
text \<open>Functions that compute codes of higher-arity constant functions:\<close>
definition code_constn :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"code_constn n c \<equiv>
if n = 1 then code_const1 c
else quad_encode 3 n (code_const1 c) (singleton_encode (triple_encode 2 n 0))"
lemma code_constn: "code_constn (Suc n) c = encode (r_constn n c)"
unfolding code_constn_def using code_const1 r_constn_def
by (cases "n = 0") simp_all
definition r_code_constn :: "nat \<Rightarrow> recf" where
"r_code_constn n \<equiv>
if n = 1 then r_code_const1
else
Cn 1 r_prod_encode
[r_const 3,
Cn 1 r_prod_encode
[r_const n,
Cn 1 r_prod_encode
[r_code_const1,
Cn 1 r_singleton_encode
[Cn 1 r_prod_encode
[r_const 2, Cn 1 r_prod_encode [r_const n, Z]]]]]]"
lemma r_code_constn_prim: "prim_recfn 1 (r_code_constn n)"
by (simp_all add: r_code_constn_def r_code_const1_prim)
lemma r_code_constn: "eval (r_code_constn n) [c] \<down>= code_constn n c"
by (auto simp add: r_code_constn_def r_code_const1 code_constn_def r_code_const1_prim)
text \<open>Computing codes of $m$-ary projections:\<close>
definition code_id :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"code_id m n \<equiv> triple_encode 2 m n"
lemma code_id: "encode (Id m n) = code_id m n"
unfolding code_id_def by simp
text \<open>The functions $s^m_n$ are represented by the following function.
The value $m$ corresponds to the length of @{term "cs"}.\<close>
definition smn :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat" where
"smn n p cs \<equiv> quad_encode
3
n
(encode (r_universal (n + length cs)))
(list_encode (code_constn n p # map (code_constn n) cs @ map (code_id n) [0..<n]))"
lemma smn:
assumes "n > 0"
shows "smn n p cs = encode
(Cn n
(r_universal (n + length cs))
(r_constn (n - 1) p # map (r_constn (n - 1)) cs @ (map (Id n) [0..<n])))"
proof -
let ?p = "r_constn (n - 1) p"
let ?gs1 = "map (r_constn (n - 1)) cs"
let ?gs2 = "map (Id n) [0..<n]"
let ?gs = "?p # ?gs1 @ ?gs2"
have "map encode ?gs1 = map (code_constn n) cs"
by (intro nth_equalityI; auto; metis code_constn assms Suc_pred)
moreover have "map encode ?gs2 = map (code_id n) [0..<n]"
by (rule nth_equalityI) (auto simp add: code_id_def)
moreover have "encode ?p = code_constn n p"
using assms code_constn[of "n - 1" p] by simp
ultimately have "map encode ?gs =
code_constn n p # map (code_constn n) cs @ map (code_id n) [0..<n]"
by simp
then show ?thesis
unfolding smn_def using assms encode.simps(4) by presburger
qed
text \<open>The next function is to help us define @{typ recf}s corresponding
to the $s^m_n$ functions. It maps $m + 1$ arguments $p, c_1, \ldots, c_m$ to
an encoded list of length $m + n + 1$. The list comprises the $m + 1$ codes
of the $n$-ary constants $p, c_1, \ldots, c_m$ and the $n$ codes for all
$n$-ary projections.\<close>
definition r_smn_aux :: "nat \<Rightarrow> nat \<Rightarrow> recf" where
"r_smn_aux n m \<equiv>
Cn (Suc m)
(r_list_encode (m + n))
(map (\<lambda>i. Cn (Suc m) (r_code_constn n) [Id (Suc m) i]) [0..<Suc m] @
map (\<lambda>i. r_constn m (code_id n i)) [0..<n])"
lemma r_smn_aux_prim: "n > 0 \<Longrightarrow> prim_recfn (Suc m) (r_smn_aux n m)"
by (auto simp add: r_smn_aux_def r_code_constn_prim)
lemma r_smn_aux:
assumes "n > 0" and "length cs = m"
shows "eval (r_smn_aux n m) (p # cs) \<down>=
list_encode (map (code_constn n) (p # cs) @ map (code_id n) [0..<n])"
proof -
let ?xs = "map (\<lambda>i. Cn (Suc m) (r_code_constn n) [Id (Suc m) i]) [0..<Suc m]"
let ?ys = "map (\<lambda>i. r_constn m (code_id n i)) [0..<n]"
have len_xs: "length ?xs = Suc m" by simp
have map_xs: "map (\<lambda>g. eval g (p # cs)) ?xs = map Some (map (code_constn n) (p # cs))"
proof (intro nth_equalityI)
show len: "length (map (\<lambda>g. eval g (p # cs)) ?xs) =
length (map Some (map (code_constn n) (p # cs)))"
by (simp add: assms(2))
have "map (\<lambda>g. eval g (p # cs)) ?xs ! i = map Some (map (code_constn n) (p # cs)) ! i"
if "i < Suc m" for i
proof -
have "map (\<lambda>g. eval g (p # cs)) ?xs ! i = (\<lambda>g. eval g (p # cs)) (?xs ! i)"
using len_xs that by (metis nth_map)
also have "... = eval (Cn (Suc m) (r_code_constn n) [Id (Suc m) i]) (p # cs)"
using that len_xs
by (metis (no_types, lifting) add.left_neutral length_map nth_map nth_upt)
also have "... = eval (r_code_constn n) [the (eval (Id (Suc m) i) (p # cs))]"
using r_code_constn_prim assms(2) that by simp
also have "... = eval (r_code_constn n) [(p # cs) ! i]"
using len that by simp
finally have "map (\<lambda>g. eval g (p # cs)) ?xs ! i \<down>= code_constn n ((p # cs) ! i)"
using r_code_constn by simp
then show ?thesis
using len_xs len that by (metis length_map nth_map)
qed
moreover have "length (map (\<lambda>g. eval g (p # cs)) ?xs) = Suc m" by simp
ultimately show "\<And>i. i < length (map (\<lambda>g. eval g (p # cs)) ?xs) \<Longrightarrow>
map (\<lambda>g. eval g (p # cs)) ?xs ! i =
map Some (map (code_constn n) (p # cs)) ! i"
by simp
qed
moreover have "map (\<lambda>g. eval g (p # cs)) ?ys = map Some (map (code_id n) [0..<n])"
using assms(2) by (intro nth_equalityI; auto)
ultimately have "map (\<lambda>g. eval g (p # cs)) (?xs @ ?ys) =
map Some (map (code_constn n) (p # cs) @ map (code_id n) [0..<n])"
by (metis map_append)
moreover have "map (\<lambda>x. the (eval x (p # cs))) (?xs @ ?ys) =
map the (map (\<lambda>x. eval x (p # cs)) (?xs @ ?ys))"
by simp
ultimately have *: "map (\<lambda>g. the (eval g (p # cs))) (?xs @ ?ys) =
(map (code_constn n) (p # cs) @ map (code_id n) [0..<n])"
by simp
have "\<forall>i<length ?xs. eval (?xs ! i) (p # cs) = map (\<lambda>g. eval g (p # cs)) ?xs ! i"
by (metis nth_map)
then have
"\<forall>i<length ?xs. eval (?xs ! i) (p # cs) = map Some (map (code_constn n) (p # cs)) ! i"
using map_xs by simp
then have "\<forall>i<length ?xs. eval (?xs ! i) (p # cs) \<down>"
using assms map_xs by (metis length_map nth_map option.simps(3))
then have xs_converg: "\<forall>z\<in>set ?xs. eval z (p # cs) \<down>"
by (metis in_set_conv_nth)
have "\<forall>i<length ?ys. eval (?ys ! i) (p # cs) = map (\<lambda>x. eval x (p # cs)) ?ys ! i"
by simp
then have
"\<forall>i<length ?ys. eval (?ys ! i) (p # cs) = map Some (map (code_id n) [0..<n]) ! i"
using assms(2) by simp
then have "\<forall>i<length ?ys. eval (?ys ! i) (p # cs) \<down>"
by simp
then have "\<forall>z\<in>set (?xs @ ?ys). eval z (p # cs) \<down>"
using xs_converg by auto
moreover have "recfn (length (p # cs)) (Cn (Suc m) (r_list_encode (m + n)) (?xs @ ?ys))"
using assms r_code_constn_prim by auto
ultimately have "eval (r_smn_aux n m) (p # cs) =
eval (r_list_encode (m + n)) (map (\<lambda>g. the (eval g (p # cs))) (?xs @ ?ys))"
unfolding r_smn_aux_def using assms by simp
then have "eval (r_smn_aux n m) (p # cs) =
eval (r_list_encode (m + n)) (map (code_constn n) (p # cs) @ map (code_id n) [0..<n])"
using * by metis
moreover have "length (?xs @ ?ys) = Suc (m + n)" by simp
ultimately show ?thesis
using r_list_encode * assms(1) by (metis (no_types, lifting) length_map)
qed
text \<open>For all $m, n > 0$, the @{typ recf} corresponding to $s^m_n$ is
given by the next function.\<close>
definition r_smn :: "nat \<Rightarrow> nat \<Rightarrow> recf" where
"r_smn n m \<equiv>
Cn (Suc m) r_prod_encode
[r_constn m 3,
Cn (Suc m) r_prod_encode
[r_constn m n,
Cn (Suc m) r_prod_encode
[r_constn m (encode (r_universal (n + m))), r_smn_aux n m]]]"
lemma r_smn_prim [simp]: "n > 0 \<Longrightarrow> prim_recfn (Suc m) (r_smn n m)"
by (simp_all add: r_smn_def r_smn_aux_prim)
lemma r_smn:
assumes "n > 0" and "length cs = m"
shows "eval (r_smn n m) (p # cs) \<down>= smn n p cs"
using assms r_smn_def r_smn_aux smn_def r_smn_aux_prim by simp
lemma map_eval_Some_the:
assumes "map (\<lambda>g. eval g xs) gs = map Some ys"
shows "map (\<lambda>g. the (eval g xs)) gs = ys"
using assms
by (metis (no_types, lifting) length_map nth_equalityI nth_map option.sel)
text \<open>The essential part of the $s$-$m$-$n$ theorem: For all $m, n > 0$
the function $s^m_n$ satisfies
\[
\varphi_p^{(m + n)}(c_1, \dots,c_m, x_1, \dots, x_n) =
\varphi_{s^m_n(p, c_1, \dots,c_m)}^{(n)}(x_1, \dots, x_n)
\] for all $p, c_i, x_j$.\<close>
lemma smn_lemma:
assumes "n > 0" and len_cs: "length cs = m" and len_xs: "length xs = n"
shows "eval (r_universal (m + n)) (p # cs @ xs) =
eval (r_universal n) ((the (eval (r_smn n m) (p # cs))) # xs)"
proof -
let ?s = "r_smn n m"
let ?f = "Cn n
(r_universal (n + length cs))
(r_constn (n - 1) p # map (r_constn (n - 1)) cs @ (map (Id n) [0..<n]))"
have "eval ?s (p # cs) \<down>= smn n p cs"
using assms r_smn by simp
then have eval_s: "eval ?s (p # cs) \<down>= encode ?f"
by (simp add: assms(1) smn)
have "recfn n ?f"
using len_cs assms by auto
then have *: "eval (r_universal n) ((encode ?f) # xs) = eval ?f xs"
using r_universal[of ?f n, OF _ len_xs] by simp
let ?gs = "r_constn (n - 1) p # map (r_constn (n - 1)) cs @ map (Id n) [0..<n]"
have "\<forall>g\<in>set ?gs. eval g xs \<down>"
using len_cs len_xs assms by auto
then have "eval ?f xs =
eval (r_universal (n + length cs)) (map (\<lambda>g. the (eval g xs)) ?gs)"
- using len_cs len_xs assms `recfn n ?f` by simp
+ using len_cs len_xs assms \<open>recfn n ?f\<close> by simp
then have "eval ?f xs = eval (r_universal (m + n)) (map (\<lambda>g. the (eval g xs)) ?gs)"
by (simp add: len_cs add.commute)
then have "eval (r_universal n) ((the (eval ?s (p # cs))) # xs) =
eval (r_universal (m + n)) (map (\<lambda>g. the (eval g xs)) ?gs)"
using eval_s * by simp
moreover have "map (\<lambda>g. the (eval g xs)) ?gs = p # cs @ xs"
proof (intro nth_equalityI)
show "length (map (\<lambda>g. the (eval g xs)) ?gs) = length (p # cs @ xs)"
by (simp add: len_xs)
have len: "length (map (\<lambda>g. the (eval g xs)) ?gs) = Suc (m + n)"
by (simp add: len_cs)
moreover have "map (\<lambda>g. the (eval g xs)) ?gs ! i = (p # cs @ xs) ! i"
if "i < Suc (m + n)" for i
proof -
from that consider "i = 0" | "i > 0 \<and> i < Suc m" | "Suc m \<le> i \<and> i < Suc (m + n)"
using not_le_imp_less by auto
then show ?thesis
proof (cases)
case 1
then show ?thesis using assms(1) len_xs by simp
next
case 2
then have "?gs ! i = (map (r_constn (n - 1)) cs) ! (i - 1)"
using len_cs
by (metis One_nat_def Suc_less_eq Suc_pred length_map
less_numeral_extra(3) nth_Cons' nth_append)
then have "map (\<lambda>g. the (eval g xs)) ?gs ! i =
(\<lambda>g. the (eval g xs)) ((map (r_constn (n - 1)) cs) ! (i - 1))"
using len by (metis length_map nth_map that)
also have "... = the (eval ((r_constn (n - 1) (cs ! (i - 1)))) xs)"
using 2 len_cs by auto
also have "... = cs ! (i - 1)"
using r_constn len_xs assms(1) by simp
also have "... = (p # cs @ xs) ! i"
using 2 len_cs
by (metis diff_Suc_1 less_Suc_eq_0_disj less_numeral_extra(3) nth_Cons' nth_append)
finally show ?thesis .
next
case 3
then have "?gs ! i = (map (Id n) [0..<n]) ! (i - Suc m)"
using len_cs
by (simp; metis (no_types, lifting) One_nat_def Suc_less_eq add_leE
plus_1_eq_Suc diff_diff_left length_map not_le nth_append
ordered_cancel_comm_monoid_diff_class.add_diff_inverse)
then have "map (\<lambda>g. the (eval g xs)) ?gs ! i =
(\<lambda>g. the (eval g xs)) ((map (Id n) [0..<n]) ! (i - Suc m))"
using len by (metis length_map nth_map that)
also have "... = the (eval ((Id n (i - Suc m))) xs)"
using 3 len_cs by auto
also have "... = xs ! (i - Suc m)"
using len_xs 3 by auto
also have "... = (p # cs @ xs) ! i"
using len_cs len_xs 3
by (metis diff_Suc_1 diff_diff_left less_Suc_eq_0_disj not_le nth_Cons'
nth_append plus_1_eq_Suc)
finally show ?thesis .
qed
qed
ultimately show "map (\<lambda>g. the (eval g xs)) ?gs ! i = (p # cs @ xs) ! i"
if "i < length (map (\<lambda>g. the (eval g xs)) ?gs)" for i
using that by simp
qed
ultimately show ?thesis by simp
qed
theorem smn_theorem:
assumes "n > 0"
shows "\<exists>s. prim_recfn (Suc m) s \<and>
(\<forall>p cs xs. length cs = m \<and> length xs = n \<longrightarrow>
eval (r_universal (m + n)) (p # cs @ xs) =
eval (r_universal n) ((the (eval s (p # cs))) # xs))"
using smn_lemma exI[of _ "r_smn n m"] assms by simp
text \<open>For every numbering, that is, binary partial recursive function,
$\psi$ there is a total recursive function $c$ that translates $\psi$-indices
into $\varphi$-indices.\<close>
lemma numbering_translation:
assumes "recfn 2 psi"
obtains c where
"recfn 1 c"
"total c"
"\<forall>i x. eval psi [i, x] = eval r_phi [the (eval c [i]), x]"
proof -
let ?p = "encode psi"
define c where "c = Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]"
then have "prim_recfn 1 c" by simp
moreover from this have "total c"
by auto
moreover have "eval r_phi [the (eval c [i]), x] = eval psi [i, x]" for i x
proof -
have "eval c [i] = eval (r_smn 1 1) [?p, i]"
using c_def by simp
then have "eval (r_universal 1) [the (eval c [i]), x] =
eval (r_universal 1) [the (eval (r_smn 1 1) [?p, i]), x]"
by simp
also have "... = eval (r_universal (1 + 1)) (?p # [i] @ [x])"
using smn_lemma[of 1 "[i]" 1 "[x]" ?p] by simp
also have "... = eval (r_universal 2) [?p, i, x]"
by (metis append_eq_Cons_conv nat_1_add_1)
also have "... = eval psi [i, x]"
using r_universal[OF assms, of "[i, x]"] by simp
finally have "eval (r_universal 1) [the (eval c [i]), x] = eval psi [i, x]" .
then show ?thesis using r_phi_def by simp
qed
ultimately show ?thesis using that by auto
qed
section \<open>Fixed-point theorems\<close>
text \<open>Fixed-point theorems (also known as recursion theorems) come in
many shapes. We prove the minimum we need for Chapter~\ref{c:iirf}.\<close>
subsection \<open>Rogers's fixed-point theorem\<close>
text \<open>In this section we prove a theorem that Rogers~\cite{Rogers87}
credits to Kleene, but admits that it is a special case and not the original
formulation. We follow Wikipedia~\cite{wiki-krt} and call it the Rogers's
fixed-point theorem.\<close>
lemma s11_inj: "inj (\<lambda>x. smn 1 p [x])"
proof
fix x\<^sub>1 x\<^sub>2 :: nat
assume "smn 1 p [x\<^sub>1] = smn 1 p [x\<^sub>2]"
then have "list_encode [code_constn 1 p, code_constn 1 x\<^sub>1, code_id 1 0] =
list_encode [code_constn 1 p, code_constn 1 x\<^sub>2, code_id 1 0]"
using smn_def by (simp add: prod_encode_eq)
then have "[code_constn 1 p, code_constn 1 x\<^sub>1, code_id 1 0] =
[code_constn 1 p, code_constn 1 x\<^sub>2, code_id 1 0]"
using list_decode_encode by metis
then have "code_constn 1 x\<^sub>1 = code_constn 1 x\<^sub>2" by simp
then show "x\<^sub>1 = x\<^sub>2"
using code_const1 code_constn code_constn_def encode_injective r_constn
by (metis One_nat_def length_Cons list.size(3) option.simps(1))
qed
definition "r_univuniv \<equiv> Cn 2 r_phi [Cn 2 r_phi [Id 2 0, Id 2 0], Id 2 1]"
lemma r_univuniv_recfn: "recfn 2 r_univuniv"
by (simp add: r_univuniv_def)
lemma r_univuniv_converg:
assumes "eval r_phi [x, x] \<down>"
shows "eval r_univuniv [x, y] = eval r_phi [the (eval r_phi [x, x]), y]"
unfolding r_univuniv_def using assms r_univuniv_recfn r_phi_recfn by simp
text \<open>Strictly speaking this is a generalization of Rogers's theorem in
that it shows the existence of infinitely many fixed-points. In conventional
terms it says that for every total recursive $f$ and $k \in \mathbb{N}$ there is
an $n \geq k$ with $\varphi_n = \varphi_{f(n)}$.\<close>
theorem rogers_fixed_point_theorem:
fixes k :: nat
assumes "recfn 1 f" and "total f"
shows "\<exists>n\<ge>k. \<forall>x. eval r_phi [n, x] = eval r_phi [the (eval f [n]), x]"
proof -
let ?p = "encode r_univuniv"
define h where "h = Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]"
then have "prim_recfn 1 h"
by simp
then have "total h"
by blast
have "eval h [x] = eval (Cn 1 (r_smn 1 1) [r_const ?p, Id 1 0]) [x]" for x
unfolding h_def by simp
then have h: "the (eval h [x]) = smn 1 ?p [x]" for x
by (simp add: r_smn)
have "eval r_phi [the (eval h [x]), y] = eval r_univuniv [x, y]" for x y
proof -
have "eval r_phi [the (eval h [x]), y] = eval r_phi [smn 1 ?p [x], y]"
using h by simp
also have "... = eval r_phi [the (eval (r_smn 1 1) [?p, x]), y]"
by (simp add: r_smn)
also have "... = eval (r_universal 2) [?p, x, y]"
using r_phi_def smn_lemma[of 1 "[x]" 1 "[y]" ?p]
by (metis Cons_eq_append_conv One_nat_def Suc_1 length_Cons
less_numeral_extra(1) list.size(3) plus_1_eq_Suc)
finally show "eval r_phi [the (eval h [x]), y] = eval r_univuniv [x, y]"
using r_universal r_univuniv_recfn by simp
qed
then have *: "eval r_phi [the (eval h [x]), y] = eval r_phi [the (eval r_phi [x, x]), y]"
if "eval r_phi [x, x] \<down>" for x y
using r_univuniv_converg that by simp
let ?fh = "Cn 1 f [h]"
have "recfn 1 ?fh"
- using `prim_recfn 1 h` assms by simp
+ using \<open>prim_recfn 1 h\<close> assms by simp
then have "infinite {r. recfn 1 r \<and> r \<simeq> ?fh}"
using exteq_infinite[of ?fh 1] by simp
then have "infinite (encode ` {r. recfn 1 r \<and> r \<simeq> ?fh})" (is "infinite ?E")
using encode_injective by (meson finite_imageD inj_onI)
then have "infinite ((\<lambda>x. smn 1 ?p [x]) ` ?E)"
using s11_inj[of ?p] by (simp add: finite_image_iff inj_on_subset)
moreover have "(\<lambda>x. smn 1 ?p [x]) ` ?E = {smn 1 ?p [encode r] |r. recfn 1 r \<and> r \<simeq> ?fh}"
by auto
ultimately have "infinite {smn 1 ?p [encode r] |r. recfn 1 r \<and> r \<simeq> ?fh}"
by simp
then obtain n where "n \<ge> k" "n \<in> {smn 1 ?p [encode r] |r. recfn 1 r \<and> r \<simeq> ?fh}"
by (meson finite_nat_set_iff_bounded_le le_cases)
then obtain r where r: "recfn 1 r" "n = smn 1 ?p [encode r]" "recfn 1 r \<and> r \<simeq> ?fh"
by auto
then have eval_r: "eval r [encode r] = eval ?fh [encode r]"
by (simp add: exteq_def)
then have eval_r': "eval r [encode r] = eval f [the (eval h [encode r])]"
- using assms `total h` `prim_recfn 1 h` by simp
+ using assms \<open>total h\<close> \<open>prim_recfn 1 h\<close> by simp
then have "eval r [encode r] \<down>"
- using `prim_recfn 1 h` assms(1,2) by simp
+ using \<open>prim_recfn 1 h\<close> assms(1,2) by simp
then have "eval r_phi [encode r, encode r] \<down>"
by (simp add: \<open>recfn 1 r\<close> r_phi)
then have "eval r_phi [the (eval h [encode r]), y] =
eval r_phi [(the (eval r_phi [encode r, encode r])), y]"
for y
using * by simp
then have "eval r_phi [the (eval h [encode r]), y] =
eval r_phi [(the (eval r [encode r])), y]"
for y
by (simp add: \<open>recfn 1 r\<close> r_phi)
moreover have "n = the (eval h [encode r])" by (simp add: h r(2))
ultimately have "eval r_phi [n, y] = eval r_phi [the (eval r [encode r]), y]" for y
by simp
then have "eval r_phi [n, y] = eval r_phi [the (eval ?fh [encode r]), y]" for y
using r by (simp add: eval_r)
moreover have "eval ?fh [encode r] = eval f [n]"
using eval_r eval_r' \<open>n = the (eval h [encode r])\<close> by auto
ultimately have "eval r_phi [n, y] = eval r_phi [the (eval f [n]), y]" for y
by simp
- with `n \<ge> k` show ?thesis by auto
+ with \<open>n \<ge> k\<close> show ?thesis by auto
qed
subsection \<open>Kleene's fixed-point theorem\<close>
text \<open>The next theorem is what Rogers~\cite[p.~214]{Rogers87} calls
Kleene's version of what we call Rogers's fixed-point theorem. More precisely
this would be Kleene's \emph{second} fixed-point theorem, but since we do not
cover the first one, we leave out the number.\<close>
theorem kleene_fixed_point_theorem:
fixes k :: nat
assumes "recfn 2 psi"
shows "\<exists>n\<ge>k. \<forall>x. eval r_phi [n, x] = eval psi [n, x]"
proof -
from numbering_translation[OF assms] obtain c where c:
"recfn 1 c"
"total c"
"\<forall>i x. eval psi [i, x] = eval r_phi [the (eval c [i]), x]"
by auto
then obtain n where "n \<ge> k" and "\<forall>x. eval r_phi [n, x] = eval r_phi [the (eval c [n]), x]"
using rogers_fixed_point_theorem by blast
with c(3) have "\<forall>x. eval r_phi [n, x] = eval psi [n, x]"
by simp
- with `n \<ge> k` show ?thesis by auto
+ with \<open>n \<ge> k\<close> show ?thesis by auto
qed
text \<open>Kleene's fixed-point theorem can be generalized to arbitrary
arities. But we need to generalize it only to binary functions in order to
show Smullyan's double fixed-point theorem in
Section~\ref{s:smullyan}.\<close>
definition "r_univuniv2 \<equiv>
Cn 3 r_phi [Cn 3 (r_universal 2) [Id 3 0, Id 3 0, Id 3 1], Id 3 2]"
lemma r_univuniv2_recfn: "recfn 3 r_univuniv2"
by (simp add: r_univuniv2_def)
lemma r_univuniv2_converg:
assumes "eval (r_universal 2) [u, u, x] \<down>"
shows "eval r_univuniv2 [u, x, y] = eval r_phi [the (eval (r_universal 2) [u, u, x]), y]"
unfolding r_univuniv2_def using assms r_univuniv2_recfn by simp
theorem kleene_fixed_point_theorem_2:
assumes "recfn 2 f" and "total f"
shows "\<exists>n.
recfn 1 n \<and>
total n \<and>
(\<forall>x y. eval r_phi [(the (eval n [x])), y] = eval r_phi [(the (eval f [the (eval n [x]), x])), y])"
proof -
let ?p = "encode r_univuniv2"
let ?s = "r_smn 1 2"
define h where "h = Cn 2 ?s [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]"
then have [simp]: "prim_recfn 2 h" by simp
{
fix u x y
have "eval h [u, x] = eval (Cn 2 ?s [r_dummy 1 (r_const ?p), Id 2 0, Id 2 1]) [u, x]"
using h_def by simp
then have "the (eval h [u, x]) = smn 1 ?p [u, x]"
by (simp add: r_smn)
then have "eval r_phi [the (eval h [u, x]), y] = eval r_phi [smn 1 ?p [u, x], y]"
by simp
also have "... =
eval r_phi
[encode (Cn 1 (r_universal 3) (r_constn 0 ?p # r_constn 0 u # r_constn 0 x # [Id 1 0])),
y]"
using smn[of 1 ?p "[u, x]"] by (simp add: numeral_3_eq_3)
also have "... =
eval r_phi
[encode (Cn 1 (r_universal 3) (r_const ?p # r_const u # r_const x # [Id 1 0])), y]"
(is "_ = eval r_phi [encode ?f, y]")
by (simp add: r_constn_def)
also have "... = eval ?f [y]"
using r_phi'[of ?f] by auto
also have "... = eval (r_universal 3) [?p, u, x, y]"
using r_univuniv2_recfn r_universal r_phi by auto
also have "... = eval r_univuniv2 [u, x, y]"
using r_universal
by (simp add: r_universal r_univuniv2_recfn)
finally have "eval r_phi [the (eval h [u, x]), y] = eval r_univuniv2 [u, x, y]" .
}
then have *: "eval r_phi [the (eval h [u, x]), y] =
eval r_phi [the (eval (r_universal 2) [u, u, x]), y]"
if "eval (r_universal 2) [u, u, x] \<down>" for u x y
using r_univuniv2_converg that by simp
let ?fh = "Cn 2 f [h, Id 2 1]"
let ?e = "encode ?fh"
have "recfn 2 ?fh"
using assms by simp
have "total h"
by auto
then have "total ?fh"
using assms Cn_total totalI2[of ?fh] by fastforce
let ?n = "Cn 1 h [r_const ?e, Id 1 0]"
have "recfn 1 ?n"
using assms by simp
moreover have "total ?n"
- using `total h` totalI1[of ?n] by simp
+ using \<open>total h\<close> totalI1[of ?n] by simp
moreover {
fix x y
have "eval r_phi [(the (eval ?n [x])), y] = eval r_phi [(the (eval h [?e, x])), y]"
by simp
also have "... = eval r_phi [the (eval (r_universal 2) [?e, ?e, x]), y]"
using * r_universal[of _ 2] totalE[of ?fh 2] \<open>total ?fh\<close> \<open>recfn 2 ?fh\<close>
by (metis length_Cons list.size(3) numeral_2_eq_2)
also have "... = eval r_phi [the (eval f [the (eval h [?e, x]), x]), y]"
proof -
have "eval (r_universal 2) [?e, ?e, x] \<down>"
- using totalE[OF `total ?fh`] `recfn 2 ?fh` r_universal
+ using totalE[OF \<open>total ?fh\<close>] \<open>recfn 2 ?fh\<close> r_universal
by (metis length_Cons list.size(3) numeral_2_eq_2)
moreover have "eval (r_universal 2) [?e, ?e, x] = eval ?fh [?e, x]"
by (metis \<open>recfn 2 ?fh\<close> length_Cons list.size(3) numeral_2_eq_2 r_universal)
- then show ?thesis using assms `total h` by simp
+ then show ?thesis using assms \<open>total h\<close> by simp
qed
also have "... = eval r_phi [(the (eval f [the (eval ?n [x]), x])), y]"
by simp
finally have "eval r_phi [(the (eval ?n [x])), y] =
eval r_phi [(the (eval f [the (eval ?n [x]), x])), y]" .
}
ultimately show ?thesis by blast
qed
subsection \<open>Smullyan's double fixed-point theorem\label{s:smullyan}\<close>
theorem smullyan_double_fixed_point_theorem:
assumes "recfn 2 g" and "total g" and "recfn 2 h" and "total h"
shows "\<exists>m n.
(\<forall>x. eval r_phi [m, x] = eval r_phi [the (eval g [m, n]), x]) \<and>
(\<forall>x. eval r_phi [n, x] = eval r_phi [the (eval h [m, n]), x])"
proof -
obtain m where
"recfn 1 m" and
"total m" and
m: "\<forall>x y. eval r_phi [the (eval m [x]), y] =
eval r_phi [the (eval g [the (eval m [x]), x]), y]"
using kleene_fixed_point_theorem_2[of g] assms(1,2) by auto
define k where "k = Cn 1 h [m, Id 1 0]"
then have "recfn 1 k"
- using `recfn 1 m` assms(3) by simp
+ using \<open>recfn 1 m\<close> assms(3) by simp
have "total (Id 1 0)"
by (simp add: Mn_free_imp_total)
then have "total k"
- using `total m` assms(4) Cn_total k_def `recfn 1 k` by simp
+ using \<open>total m\<close> assms(4) Cn_total k_def \<open>recfn 1 k\<close> by simp
obtain n where n: "\<forall>x. eval r_phi [n, x] = eval r_phi [the (eval k [n]), x]"
- using rogers_fixed_point_theorem[of k] `recfn 1 k` `total k` by blast
+ using rogers_fixed_point_theorem[of k] \<open>recfn 1 k\<close> \<open>total k\<close> by blast
obtain mm where mm: "eval m [n] \<down>= mm"
- using `total m` `recfn 1 m` by fastforce
+ using \<open>total m\<close> \<open>recfn 1 m\<close> by fastforce
then have "\<forall>x. eval r_phi [mm, x] = eval r_phi [the (eval g [mm, n]), x]"
by (metis m option.sel)
moreover have "\<forall>x. eval r_phi [n, x] = eval r_phi [the (eval h [mm, n]), x]"
- using k_def assms(3) `total m` `recfn 1 m` mm n by simp
+ using k_def assms(3) \<open>total m\<close> \<open>recfn 1 m\<close> mm n by simp
ultimately show ?thesis by blast
qed
section \<open>Decidable and recursively enumerable sets\label{s:decidable}\<close>
text \<open>We defined @{term decidable} already back in
Section~\ref{s:halting}: @{thm[display] decidable_def}\<close>
text \<open>The next theorem is adapted from @{thm[source]
halting_problem_undecidable}.\<close>
theorem halting_problem_phi_undecidable: "\<not> decidable {x. eval r_phi [x, x] \<down>}"
(is "\<not> decidable ?K")
proof
assume "decidable ?K"
then obtain f where "recfn 1 f" and f: "\<forall>x. eval f [x] \<down>= (if x \<in> ?K then 1 else 0)"
using decidable_def by auto
define g where "g \<equiv> Cn 1 r_ifeq_else_diverg [f, Z, Z]"
then have "recfn 1 g"
- using `recfn 1 f` r_ifeq_else_diverg_recfn by simp
+ using \<open>recfn 1 f\<close> r_ifeq_else_diverg_recfn by simp
then obtain i where i: "eval r_phi [i, x] = eval g [x]" for x
using r_phi' by auto
from g_def have "eval g [x] = (if x \<notin> ?K then Some 0 else None)" for x
- using r_ifeq_else_diverg_recfn `recfn 1 f` f by simp
+ using r_ifeq_else_diverg_recfn \<open>recfn 1 f\<close> f by simp
then have "eval g [i] \<down> \<longleftrightarrow> i \<notin> ?K" by simp
also have "... \<longleftrightarrow> eval r_phi [i, i] \<up>" by simp
also have "... \<longleftrightarrow> eval g [i] \<up>"
using i by simp
finally have "eval g [i] \<down> \<longleftrightarrow> eval g [i] \<up>" .
then show False by auto
qed
lemma decidable_complement: "decidable X \<Longrightarrow> decidable (- X)"
proof -
assume "decidable X"
then obtain f where f: "recfn 1 f" "\<forall>x. eval f [x] \<down>= (if x \<in> X then 1 else 0)"
using decidable_def by auto
define g where "g = Cn 1 r_not [f]"
then have "recfn 1 g"
by (simp add: f(1))
moreover have "eval g [x] \<down>= (if x \<in> X then 0 else 1)" for x
by (simp add: g_def f)
ultimately show ?thesis using decidable_def by auto
qed
text \<open>Finite sets are decidable.\<close>
fun r_contains :: "nat list \<Rightarrow> recf" where
"r_contains [] = Z"
| "r_contains (x # xs) = Cn 1 r_ifeq [Id 1 0, r_const x, r_const 1, r_contains xs]"
lemma r_contains_prim: "prim_recfn 1 (r_contains xs)"
by (induction xs) auto
lemma r_contains: "eval (r_contains xs) [x] \<down>= (if x \<in> set xs then 1 else 0)"
proof (induction xs arbitrary: x)
case Nil
then show ?case by simp
next
case (Cons a xs)
have "eval (r_contains (a # xs)) [x] = eval r_ifeq [x, a, 1, the (eval (r_contains xs) [x])]"
using r_contains_prim prim_recfn_total by simp
also have "... \<down>= (if x = a then 1 else if x \<in> set xs then 1 else 0)"
using Cons.IH by simp
also have "... \<down>= (if x = a \<or> x \<in> set xs then 1 else 0)"
by simp
finally show ?case by simp
qed
lemma finite_set_decidable: "finite X \<Longrightarrow> decidable X"
proof -
fix X :: "nat set"
assume "finite X"
then obtain xs where "X = set xs"
using finite_list by auto
then have "\<forall>x. eval (r_contains xs) [x] \<down>= (if x \<in> X then 1 else 0)"
using r_contains by simp
then show "decidable X"
using decidable_def r_contains_prim by blast
qed
definition semidecidable :: "nat set \<Rightarrow> bool" where
"semidecidable X \<equiv> (\<exists>f. recfn 1 f \<and> (\<forall>x. eval f [x] = (if x \<in> X then Some 1 else None)))"
text \<open>The semidecidable sets are the domains of partial recursive functions.\<close>
lemma semidecidable_iff_domain:
"semidecidable X \<longleftrightarrow> (\<exists>f. recfn 1 f \<and> (\<forall>x. eval f [x] \<down> \<longleftrightarrow> x \<in> X))"
proof
show "semidecidable X \<Longrightarrow> \<exists>f. recfn 1 f \<and> (\<forall>x. (eval f [x] \<down>) = (x \<in> X))"
using semidecidable_def by (metis option.distinct(1))
show "semidecidable X" if "\<exists>f. recfn 1 f \<and> (\<forall>x. (eval f [x] \<down>) = (x \<in> X))" for X
proof -
from that obtain f where f: "recfn 1 f" "\<forall>x. (eval f [x] \<down>) = (x \<in> X)"
by auto
let ?g = "Cn 1 (r_const 1) [f]"
have "recfn 1 ?g"
using f(1) by simp
moreover have "\<forall>x. eval ?g [x] = (if x \<in> X then Some 1 else None)"
using f by simp
ultimately show "semidecidable X"
using semidecidable_def by blast
qed
qed
lemma decidable_imp_semidecidable: "decidable X \<Longrightarrow> semidecidable X"
proof -
assume "decidable X"
then obtain f where f: "recfn 1 f" "\<forall>x. eval f [x] \<down>= (if x \<in> X then 1 else 0)"
using decidable_def by auto
define g where "g = Cn 1 r_ifeq_else_diverg [f, r_const 1, r_const 1]"
then have "recfn 1 g"
by (simp add: f(1))
have "eval g [x] = eval r_ifeq_else_diverg [if x \<in> X then 1 else 0, 1, 1]" for x
by (simp add: g_def f)
then have "\<And>x. x \<in> X \<Longrightarrow> eval g [x] \<down>= 1" and "\<And>x. x \<notin> X \<Longrightarrow> eval g [x] \<up>"
by simp_all
then show ?thesis
- using `recfn 1 g` semidecidable_def by auto
+ using \<open>recfn 1 g\<close> semidecidable_def by auto
qed
text \<open>A set is recursively enumerable if it is empty or the image of a
total recursive function.\<close>
definition recursively_enumerable :: "nat set \<Rightarrow> bool" where
"recursively_enumerable X \<equiv>
X = {} \<or> (\<exists>f. recfn 1 f \<and> total f \<and> X = {the (eval f [x]) |x. x \<in> UNIV})"
theorem recursively_enumerable_iff_semidecidable:
"recursively_enumerable X \<longleftrightarrow> semidecidable X"
proof
show "semidecidable X" if "recursively_enumerable X" for X
proof (cases)
assume "X = {}"
then show ?thesis
using finite_set_decidable decidable_imp_semidecidable
recursively_enumerable_def semidecidable_def
by blast
next
assume "X \<noteq> {}"
with that obtain f where f: "recfn 1 f" "total f" "X = {the (eval f [x]) |x. x \<in> UNIV}"
using recursively_enumerable_def by blast
define h where "h = Cn 2 r_eq [Cn 2 f [Id 2 0], Id 2 1]"
then have "recfn 2 h"
using f(1) by simp
from h_def have h: "eval h [x, y] \<down>= 0 \<longleftrightarrow> the (eval f [x]) = y" for x y
using f(1,2) by simp
- from h_def `recfn 2 h` totalI2 f(2) have "total h" by simp
+ from h_def \<open>recfn 2 h\<close> totalI2 f(2) have "total h" by simp
define g where "g = Mn 1 h"
then have "recfn 1 g"
using h_def f(1) by simp
then have "eval g [y] =
(if (\<exists>x. eval h [x, y] \<down>= 0 \<and> (\<forall>x'<x. eval h [x', y] \<down>))
then Some (LEAST x. eval h [x, y] \<down>= 0)
else None)" for y
- using g_def `total h` f(2) by simp
+ using g_def \<open>total h\<close> f(2) by simp
then have "eval g [y] =
(if \<exists>x. eval h [x, y] \<down>= 0
then Some (LEAST x. eval h [x, y] \<down>= 0)
else None)" for y
- using `total h` `recfn 2 h` by simp
+ using \<open>total h\<close> \<open>recfn 2 h\<close> by simp
then have "eval g [y] \<down> \<longleftrightarrow> (\<exists>x. eval h [x, y] \<down>= 0)" for y
by simp
with h have "eval g [y] \<down> \<longleftrightarrow> (\<exists>x. the (eval f [x]) = y)" for y
by simp
with f(3) have "eval g [y] \<down> \<longleftrightarrow> y \<in> X" for y
by auto
- with `recfn 1 g` semidecidable_iff_domain show ?thesis by auto
+ with \<open>recfn 1 g\<close> semidecidable_iff_domain show ?thesis by auto
qed
show "recursively_enumerable X" if "semidecidable X" for X
proof (cases)
assume "X = {}"
then show ?thesis using recursively_enumerable_def by simp
next
assume "X \<noteq> {}"
then obtain x\<^sub>0 where "x\<^sub>0 \<in> X" by auto
from that semidecidable_iff_domain obtain f where f: "recfn 1 f" "\<forall>x. eval f [x] \<down> \<longleftrightarrow> x \<in> X"
by auto
let ?i = "encode f"
have i: "\<And>x. eval f [x] = eval r_phi [?i, x]"
using r_phi' f(1) by simp
- with `x\<^sub>0 \<in> X` f(2) have "eval r_phi [?i, x\<^sub>0] \<down>" by simp
+ with \<open>x\<^sub>0 \<in> X\<close> f(2) have "eval r_phi [?i, x\<^sub>0] \<down>" by simp
then obtain g where g: "recfn 1 g" "total g" "\<forall>x. eval r_phi [?i, x] \<down> = (\<exists>y. eval g [y] \<down>= x)"
using f(1) nonempty_domain_enumerable by blast
with f(2) i have "\<forall>x. x \<in> X = (\<exists>y. eval g [y] \<down>= x)"
by simp
then have "\<forall>x. x \<in> X = (\<exists>y. the (eval g [y]) = x)"
using totalE[OF g(2) g(1)]
by (metis One_nat_def length_Cons list.size(3) option.collapse option.sel)
then have "X = {the (eval g [y]) |y. y \<in> UNIV}"
by auto
with g(1,2) show ?thesis using recursively_enumerable_def by auto
qed
qed
text \<open>The next goal is to show that a set is decidable iff. it and its
complement are semidecidable. For this we use the concurrent evaluation
function.\<close>
lemma semidecidable_decidable:
assumes "semidecidable X" and "semidecidable (- X)"
shows "decidable X"
proof -
obtain f where f: "recfn 1 f \<and> (\<forall>x. eval f [x] \<down> \<longleftrightarrow> x \<in> X)"
using assms(1) semidecidable_iff_domain by auto
let ?i = "encode f"
obtain g where g: "recfn 1 g \<and> (\<forall>x. eval g [x] \<down> \<longleftrightarrow> x \<in> (- X))"
using assms(2) semidecidable_iff_domain by auto
let ?j = "encode g"
define d where "d = Cn 1 r_pdec1 [Cn 1 r_parallel [r_const ?j, r_const ?i, Id 1 0]]"
then have "recfn 1 d"
by (simp add: d_def)
have *: "\<And>x. eval r_phi [?i, x] = eval f [x]" "\<And>x. eval r_phi [?j, x] = eval g [x]"
using f g r_phi' by simp_all
have "eval d [x] \<down>= 1" if "x \<in> X" for x
proof -
have "eval f [x] \<down>"
using f that by simp
moreover have "eval g [x] \<up>"
using g that by blast
ultimately have "eval r_parallel [?j, ?i, x] \<down>= prod_encode (1, the (eval f [x]))"
using * r_parallel(3) by simp
with d_def show ?thesis by simp
qed
moreover have "eval d [x] \<down>= 0" if "x \<notin> X" for x
proof -
have "eval g [x] \<down>"
using g that by simp
moreover have "eval f [x] \<up>"
using f that by blast
ultimately have "eval r_parallel [?j, ?i, x] \<down>= prod_encode (0, the (eval g [x]))"
using * r_parallel(2) by blast
with d_def show ?thesis by simp
qed
ultimately show ?thesis
- using decidable_def `recfn 1 d` by auto
+ using decidable_def \<open>recfn 1 d\<close> by auto
qed
theorem decidable_iff_semidecidable_complement:
"decidable X \<longleftrightarrow> semidecidable X \<and> semidecidable (- X)"
using semidecidable_decidable decidable_imp_semidecidable decidable_complement
by blast
section \<open>Rice's theorem\<close>
definition index_set :: "nat set \<Rightarrow> bool" where
"index_set I \<equiv> \<forall>i j. i \<in> I \<and> (\<forall>x. eval r_phi [i, x] = eval r_phi [j, x]) \<longrightarrow> j \<in> I"
lemma index_set_closed_in:
assumes "index_set I" and "i \<in> I" and "\<forall>x. eval r_phi [i, x] = eval r_phi [j, x]"
shows "j \<in> I"
using index_set_def assms by simp
lemma index_set_closed_not_in:
assumes "index_set I" and "i \<notin> I" and "\<forall>x. eval r_phi [i, x] = eval r_phi [j, x]"
shows "j \<notin> I"
using index_set_def assms by metis
theorem rice_theorem:
assumes "index_set I" and "I \<noteq> UNIV" and "I \<noteq> {}"
shows "\<not> decidable I"
proof
assume "decidable I"
then obtain d where d: "recfn 1 d" "\<forall>i. eval d [i] \<down>= (if i \<in> I then 1 else 0)"
using decidable_def by auto
obtain j\<^sub>1 j\<^sub>2 where "j\<^sub>1 \<notin> I" and "j\<^sub>2 \<in> I"
using assms(2,3) by auto
let ?if = "Cn 2 r_ifz [Cn 2 d [Id 2 0], r_dummy 1 (r_const j\<^sub>2), r_dummy 1 (r_const j\<^sub>1)]"
define psi where "psi = Cn 2 r_phi [?if, Id 2 1] "
then have "recfn 2 psi"
by (simp add: d)
have "eval ?if [x, y] = Some (if x \<in> I then j\<^sub>1 else j\<^sub>2)" for x y
by (simp add: d)
moreover have "eval psi [x, y] = eval (Cn 2 r_phi [?if, Id 2 1]) [x, y]" for x y
using psi_def by simp
ultimately have psi: "eval psi [x, y] = eval r_phi [if x \<in> I then j\<^sub>1 else j\<^sub>2, y]" for x y
by (simp add: d)
then have in_I: "eval psi [x, y] = eval r_phi [j\<^sub>1, y]" if "x \<in> I" for x y
by (simp add: that)
have not_in_I: "eval psi [x, y] = eval r_phi [j\<^sub>2, y]" if "x \<notin> I" for x y
by (simp add: psi that)
obtain n where n: "\<forall>x. eval r_phi [n, x] = eval psi [n, x]"
- using kleene_fixed_point_theorem[OF `recfn 2 psi`] by auto
+ using kleene_fixed_point_theorem[OF \<open>recfn 2 psi\<close>] by auto
show False
proof cases
assume "n \<in> I"
then have "\<forall>x. eval r_phi [n, x] = eval r_phi [j\<^sub>1, x]"
using n in_I by simp
then have "n \<notin> I"
- using `j\<^sub>1 \<notin> I` index_set_closed_not_in[OF assms(1)] by simp
- with `n \<in> I` show False by simp
+ using \<open>j\<^sub>1 \<notin> I\<close> index_set_closed_not_in[OF assms(1)] by simp
+ with \<open>n \<in> I\<close> show False by simp
next
assume "n \<notin> I"
then have "\<forall>x. eval r_phi [n, x] = eval r_phi [j\<^sub>2, x]"
using n not_in_I by simp
then have "n \<in> I"
- using `j\<^sub>2 \<in> I` index_set_closed_in[OF assms(1)] by simp
- with `n \<notin> I` show False by simp
+ using \<open>j\<^sub>2 \<in> I\<close> index_set_closed_in[OF assms(1)] by simp
+ with \<open>n \<notin> I\<close> show False by simp
qed
qed
section \<open>Partial recursive functions as actual functions\label{s:alternative}\<close>
text \<open>A well-formed @{typ recf} describes an algorithm. Usually,
however, partial recursive functions are considered to be partial functions,
that is, right-unique binary relations. This distinction did not matter much
until now, because we were mostly concerned with the \emph{existence} of
partial recursive functions, which is equivalent to the existence of
algorithms. Whenever it did matter, we could use the extensional equivalence
@{term "exteq"}. In Chapter~\ref{c:iirf}, however, we will deal with sets of
functions and sets of sets of functions.
For illustration consider the singleton set containing only the unary zero
function. It could be expressed by @{term "{Z}"}, but this would not contain
@{term[names_short] "Cn 1 (Id 1 0) [Z]"}, which computes the same function.
The alternative representation as @{term "{f. f \<simeq> Z}"} is not a
singleton set. Another alternative would be to identify partial recursive
functions with the equivalence classes of @{term "exteq"}. This would work
for all arities. But since we will only need unary and binary functions, we
can go for the less general but simpler alternative of regarding partial
recursive functions as certain functions of types @{typ "nat \<Rightarrow>
nat option"} and @{typ "nat \<Rightarrow> nat \<Rightarrow> nat option"}.
With this notation we can represent the aforementioned set by @{term
"{\<lambda>_. Some (0::nat)}"} and express that the function @{term "\<lambda>_.
Some (0::nat)"} is total recursive.
In addition terms get shorter, for instance, @{term "eval r_func [i, x]"}
becomes @{term "func i x"}.\<close>
subsection \<open>The definitions\<close>
type_synonym partial1 = "nat \<Rightarrow> nat option"
type_synonym partial2 = "nat \<Rightarrow> nat \<Rightarrow> nat option"
definition total1 :: "partial1 \<Rightarrow> bool" where
"total1 f \<equiv> \<forall>x. f x \<down>"
definition total2 :: "partial2 \<Rightarrow> bool" where
"total2 f \<equiv> \<forall>x y. f x y \<down>"
lemma total1I [intro]: "(\<And>x. f x \<down>) \<Longrightarrow> total1 f"
using total1_def by simp
lemma total2I [intro]: "(\<And>x y. f x y \<down>) \<Longrightarrow> total2 f"
using total2_def by simp
lemma total1E [dest, simp]: "total1 f \<Longrightarrow> f x \<down>"
using total1_def by simp
lemma total2E [dest, simp]: "total2 f \<Longrightarrow> f x y \<down>"
using total2_def by simp
definition P1 :: "partial1 set" ("\<P>") where
"\<P> \<equiv> {\<lambda>x. eval r [x] |r. recfn 1 r}"
definition P2 :: "partial2 set" ("\<P>\<^sup>2") where
"\<P>\<^sup>2 \<equiv> {\<lambda>x y. eval r [x, y] |r. recfn 2 r}"
definition R1 :: "partial1 set" ("\<R>") where
"\<R> \<equiv> {\<lambda>x. eval r [x] |r. recfn 1 r \<and> total r}"
definition R2 :: "partial2 set" ("\<R>\<^sup>2") where
"\<R>\<^sup>2 \<equiv> {\<lambda>x y. eval r [x, y] |r. recfn 2 r \<and> total r}"
definition Prim1 :: "partial1 set" where
"Prim1 \<equiv> {\<lambda>x. eval r [x] |r. prim_recfn 1 r}"
definition Prim2 :: "partial2 set" where
"Prim2 \<equiv> {\<lambda>x y. eval r [x, y] |r. prim_recfn 2 r}"
lemma R1_imp_P1 [simp, elim]: "f \<in> \<R> \<Longrightarrow> f \<in> \<P>"
using R1_def P1_def by auto
lemma R2_imp_P2 [simp, elim]: "f \<in> \<R>\<^sup>2 \<Longrightarrow> f \<in> \<P>\<^sup>2"
using R2_def P2_def by auto
lemma Prim1_imp_R1 [simp, elim]: "f \<in> Prim1 \<Longrightarrow> f \<in> \<R>"
unfolding Prim1_def R1_def by auto
lemma Prim2_imp_R2 [simp, elim]: "f \<in> Prim2 \<Longrightarrow> f \<in> \<R>\<^sup>2"
unfolding Prim2_def R2_def by auto
lemma P1E [elim]:
assumes "f \<in> \<P>"
obtains r where "recfn 1 r" and "\<forall>x. eval r [x] = f x"
using assms P1_def by force
lemma P2E [elim]:
assumes "f \<in> \<P>\<^sup>2"
obtains r where "recfn 2 r" and "\<forall>x y. eval r [x, y] = f x y"
using assms P2_def by force
lemma P1I [intro]:
assumes "recfn 1 r" and "(\<lambda>x. eval r [x]) = f"
shows "f \<in> \<P>"
using assms P1_def by auto
lemma P2I [intro]:
assumes "recfn 2 r" and "\<And>x y. eval r [x, y] = f x y"
shows "f \<in> \<P>\<^sup>2"
proof -
have "(\<lambda>x y. eval r [x, y]) = f"
using assms(2) by simp
then show ?thesis
using assms(1) P2_def by auto
qed
lemma R1I [intro]:
assumes "recfn 1 r" and "total r" and "\<And>x. eval r [x] = f x"
shows "f \<in> \<R>"
unfolding R1_def
using CollectI[of "\<lambda>f. \<exists>r. f = (\<lambda>x. eval r [x]) \<and> recfn 1 r \<and> total r" f] assms
by metis
lemma R1E [elim]:
assumes "f \<in> \<R>"
obtains r where "recfn 1 r" and "total r" and "f = (\<lambda>x. eval r [x])"
using assms R1_def by auto
lemma R2I [intro]:
assumes "recfn 2 r" and "total r" and "\<And>x y. eval r [x, y] = f x y"
shows "f \<in> \<R>\<^sup>2"
unfolding R2_def
using CollectI[of "\<lambda>f. \<exists>r. f = (\<lambda>x y. eval r [x, y]) \<and> recfn 2 r \<and> total r" f] assms
by metis
lemma R1_SOME:
assumes "f \<in> \<R>"
and "r = (SOME r'. recfn 1 r' \<and> total r' \<and> f = (\<lambda>x. eval r' [x]))"
(is "r = (SOME r'. ?P r')")
shows "recfn 1 r"
and "\<And>x. eval r [x] \<down>"
and "\<And>x. f x = eval r [x]"
and "f = (\<lambda>x. eval r [x])"
proof -
obtain r' where "?P r'"
using R1E[OF assms(1)] by auto
then show "recfn 1 r" "\<And>b. eval r [b] \<down>" "\<And>x. f x = eval r [x]"
using someI[of ?P r'] assms(2) totalE[of r] by (auto, metis)
then show "f = (\<lambda>x. eval r [x])" by auto
qed
lemma R2E [elim]:
assumes "f \<in> \<R>\<^sup>2"
obtains r where "recfn 2 r" and "total r" and "f = (\<lambda>x\<^sub>1 x\<^sub>2. eval r [x\<^sub>1, x\<^sub>2])"
using assms R2_def by auto
lemma R1_imp_total1 [simp]: "f \<in> \<R> \<Longrightarrow> total1 f"
using total1I by fastforce
lemma R2_imp_total2 [simp]: "f \<in> \<R>\<^sup>2 \<Longrightarrow> total2 f"
using totalE by fastforce
lemma Prim1I [intro]:
assumes "prim_recfn 1 r" and "\<And>x. f x = eval r [x]"
shows "f \<in> Prim1"
using assms Prim1_def by blast
lemma Prim2I [intro]:
assumes "prim_recfn 2 r" and "\<And>x y. f x y = eval r [x, y]"
shows "f \<in> Prim2"
using assms Prim2_def by blast
lemma P1_total_imp_R1 [intro]:
assumes "f \<in> \<P>" and "total1 f"
shows "f \<in> \<R>"
using assms totalI1 by force
lemma P2_total_imp_R2 [intro]:
assumes "f \<in> \<P>\<^sup>2 " and "total2 f"
shows "f \<in> \<R>\<^sup>2"
using assms totalI2 by force
subsection \<open>Some simple properties\<close>
text \<open>In order to show that a @{typ partial1} or @{typ partial2}
function is in @{term "\<P>"}, @{term "\<P>\<^sup>2"}, @{term "\<R>"}, @{term
"\<R>\<^sup>2"}, @{term "Prim1"}, or @{term "Prim2"} we will usually have to
find a suitable @{typ recf}. But for some simple or frequent cases this
section provides shortcuts.\<close>
lemma identity_in_R1: "Some \<in> \<R>"
proof -
have "\<forall>x. eval (Id 1 0) [x] \<down>= x" by simp
moreover have "recfn 1 (Id 1 0)" by simp
moreover have "total (Id 1 0)"
by (simp add: totalI1)
ultimately show ?thesis by blast
qed
lemma P2_proj_P1 [simp, elim]:
assumes "\<psi> \<in> \<P>\<^sup>2"
shows "\<psi> i \<in> \<P>"
proof -
from assms obtain u where u: "recfn 2 u" "(\<lambda>x\<^sub>1 x\<^sub>2. eval u [x\<^sub>1, x\<^sub>2]) = \<psi>"
by auto
define v where "v \<equiv> Cn 1 u [r_const i, Id 1 0]"
then have "recfn 1 v" "(\<lambda>x. eval v [x]) = \<psi> i"
using u by auto
then show ?thesis by auto
qed
lemma R2_proj_R1 [simp, elim]:
assumes "\<psi> \<in> \<R>\<^sup>2"
shows "\<psi> i \<in> \<R>"
proof -
from assms have "\<psi> \<in> \<P>\<^sup>2" by simp
then have "\<psi> i \<in> \<P>" by auto
moreover have "total1 (\<psi> i)"
using assms by (simp add: total1I)
ultimately show ?thesis by auto
qed
lemma const_in_Prim1: "(\<lambda>_. Some c) \<in> Prim1"
proof -
define r where "r = r_const c"
then have "\<And>x. eval r [x] = Some c" by simp
moreover have "recfn 1 r" "Mn_free r"
using r_def by simp_all
ultimately show ?thesis by auto
qed
lemma concat_P1_P1:
assumes "f \<in> \<P>" and "g \<in> \<P>"
shows "(\<lambda>x. if g x \<down> \<and> f (the (g x)) \<down> then Some (the (f (the (g x)))) else None) \<in> \<P>"
(is "?h \<in> \<P>")
proof -
obtain rf where rf: "recfn 1 rf" "\<forall>x. eval rf [x] = f x"
using assms(1) by auto
obtain rg where rg: "recfn 1 rg" "\<forall>x. eval rg [x] = g x"
using assms(2) by auto
let ?rh = "Cn 1 rf [rg]"
have "recfn 1 ?rh"
using rf(1) rg(1) by simp
moreover have "eval ?rh [x] = ?h x" for x
using rf rg by simp
ultimately show ?thesis by blast
qed
lemma P1_update_P1:
assumes "f \<in> \<P>"
shows "f(x:=z) \<in> \<P>"
proof (cases z)
case None
define re where "re \<equiv> Mn 1 (r_constn 1 1)"
from assms obtain r where r: "recfn 1 r" "(\<lambda>u. eval r [u]) = f"
by auto
define r' where "r' = Cn 1 (r_lifz re r) [Cn 1 r_eq [Id 1 0, r_const x], Id 1 0]"
have "recfn 1 r'"
using r(1) r'_def re_def by simp
then have "eval r' [u] = eval (r_lifz re r) [if u = x then 0 else 1, u]" for u
using r'_def by simp
with r(1) have "eval r' [u] = (if u = x then None else eval r [u])" for u
using re_def re_def by simp
with r(2) have "eval r' [u] = (f(x:=None)) u" for u
by auto
then have "(\<lambda>u. eval r' [u]) = f(x:=None)"
by auto
- with None `recfn 1 r'` show ?thesis by auto
+ with None \<open>recfn 1 r'\<close> show ?thesis by auto
next
case (Some y)
from assms obtain r where r: "recfn 1 r" "(\<lambda>u. eval r [u]) = f"
by auto
define r' where
"r' \<equiv> Cn 1 (r_lifz (r_const y) r) [Cn 1 r_eq [Id 1 0, r_const x], Id 1 0]"
have "recfn 1 r'"
using r(1) r'_def by simp
then have "eval r' [u] = eval (r_lifz (r_const y) r) [if u = x then 0 else 1, u]" for u
using r'_def by simp
with r(1) have "eval r' [u] = (if u = x then Some y else eval r [u])" for u
by simp
with r(2) have "eval r' [u] = (f(x:=Some y)) u" for u
by auto
then have "(\<lambda>u. eval r' [u]) = f(x:=Some y)"
by auto
- with Some `recfn 1 r'` show ?thesis by auto
+ with Some \<open>recfn 1 r'\<close> show ?thesis by auto
qed
lemma swap_P2:
assumes "f \<in> \<P>\<^sup>2"
shows "(\<lambda>x y. f y x) \<in> \<P>\<^sup>2"
proof -
obtain r where r: "recfn 2 r" "\<And>x y. eval r [x, y] = f x y"
using assms by auto
then have "eval (r_swap r) [x, y] = f y x" for x y
by simp
moreover have "recfn 2 (r_swap r)"
using r_swap_recfn r(1) by simp
ultimately show ?thesis by auto
qed
lemma swap_R2:
assumes "f \<in> \<R>\<^sup>2"
shows "(\<lambda>x y. f y x) \<in> \<R>\<^sup>2"
using swap_P2[of f] assms
by (meson P2_total_imp_R2 R2_imp_P2 R2_imp_total2 total2E total2I)
lemma skip_P1:
assumes "f \<in> \<P>"
shows "(\<lambda>x. f (x + n)) \<in> \<P>"
proof -
obtain r where r: "recfn 1 r" "\<And>x. eval r [x] = f x"
using assms by auto
let ?s = "Cn 1 r [Cn 1 r_add [Id 1 0, r_const n]]"
have "recfn 1 ?s"
using r by simp
have "eval ?s [x] = eval r [x + n]" for x
using r by simp
with r have "eval ?s [x] = f (x + n)" for x
by simp
- with `recfn 1 ?s` show ?thesis by blast
+ with \<open>recfn 1 ?s\<close> show ?thesis by blast
qed
lemma skip_R1:
assumes "f \<in> \<R>"
shows "(\<lambda>x. f (x + n)) \<in> \<R>"
using assms skip_P1 R1_imp_total1 total1_def by auto
subsection \<open>The Gödel numbering @{term \<phi>}\label{s:goedel_numbering}\<close>
text \<open>While the term \emph{Gödel numbering} is often used generically for
mappings between natural numbers and mathematical concepts, the inductive
inference literature uses it in a more specific sense. There it is equivalent
to the notion of acceptable numbering~\cite{Rogers87}: For every numbering
there is a recursive function mapping the numbering's indices to equivalent
ones of a Gödel numbering.\<close>
definition goedel_numbering :: "partial2 \<Rightarrow> bool" where
"goedel_numbering \<psi> \<equiv> \<psi> \<in> \<P>\<^sup>2 \<and> (\<forall>\<chi>\<in>\<P>\<^sup>2. \<exists>c\<in>\<R>. \<forall>i. \<chi> i = \<psi> (the (c i)))"
lemma goedel_numbering_P2:
assumes "goedel_numbering \<psi>"
shows "\<psi> \<in> \<P>\<^sup>2"
using goedel_numbering_def assms by simp
lemma goedel_numberingE:
assumes "goedel_numbering \<psi>" and "\<chi> \<in> \<P>\<^sup>2"
obtains c where "c \<in> \<R>" and "\<forall>i. \<chi> i = \<psi> (the (c i))"
using assms goedel_numbering_def by blast
lemma goedel_numbering_universal:
assumes "goedel_numbering \<psi>" and "f \<in> \<P>"
shows "\<exists>i. \<psi> i = f"
proof -
define \<chi> :: partial2 where "\<chi> = (\<lambda>i. f)"
have "\<chi> \<in> \<P>\<^sup>2"
proof -
obtain rf where rf: "recfn 1 rf" "\<And>x. eval rf [x] = f x"
using assms(2) by auto
define r where "r = Cn 2 rf [Id 2 1]"
then have r: "recfn 2 r" "\<And>i x. eval r [i, x] = eval rf [x]"
using rf(1) by simp_all
with rf(2) have "\<And>i x. eval r [i, x] = f x" by simp
with r(1) show ?thesis using \<chi>_def by auto
qed
then obtain c where "c \<in> \<R>" and "\<forall>i. \<chi> i = \<psi> (the (c i))"
using goedel_numbering_def assms(1) by auto
with \<chi>_def show ?thesis by auto
qed
text \<open>Our standard Gödel numbering is based on @{term r_phi}:\<close>
definition phi :: partial2 ("\<phi>") where
"\<phi> i x \<equiv> eval r_phi [i, x]"
lemma phi_in_P2: "\<phi> \<in> \<P>\<^sup>2"
unfolding phi_def using r_phi_recfn by blast
text \<open>Indices of any numbering can be translated into equivalent indices
of @{term phi}, which thus is a Gödel numbering.\<close>
lemma numbering_translation_for_phi:
assumes "\<psi> \<in> \<P>\<^sup>2"
shows "\<exists>c\<in>\<R>. \<forall>i. \<psi> i = \<phi> (the (c i))"
proof -
obtain psi where psi: "recfn 2 psi" "\<And>i x. eval psi [i, x] = \<psi> i x"
using assms by auto
with numbering_translation obtain b where
"recfn 1 b" "total b" "\<forall>i x. eval psi [i, x] = eval r_phi [the (eval b [i]), x]"
by blast
moreover from this obtain c where c: "c \<in> \<R>" "\<forall>i. c i = eval b [i]"
by fast
ultimately have "\<psi> i x = \<phi> (the (c i)) x" for i x
using phi_def psi(2) by presburger
then have "\<psi> i = \<phi> (the (c i))" for i
by auto
then show ?thesis using c(1) by blast
qed
corollary goedel_numbering_phi: "goedel_numbering \<phi>"
unfolding goedel_numbering_def using numbering_translation_for_phi phi_in_P2 by simp
corollary phi_universal:
assumes "f \<in> \<P>"
obtains i where "\<phi> i = f"
using goedel_numbering_universal[OF goedel_numbering_phi assms] by auto
subsection \<open>Fixed-point theorems\<close>
text \<open>The fixed-point theorems look somewhat cleaner in the new
notation. We will only need the following ones in the next chapter.\<close>
theorem kleene_fixed_point:
fixes k :: nat
assumes "\<psi> \<in> \<P>\<^sup>2"
obtains i where "i \<ge> k" and "\<phi> i = \<psi> i"
proof -
obtain r_psi where r_psi: "recfn 2 r_psi" "\<And>i x. eval r_psi [i, x] = \<psi> i x"
using assms by auto
then obtain i where i: "i \<ge> k" "\<forall>x. eval r_phi [i, x] = eval r_psi [i, x]"
using kleene_fixed_point_theorem by blast
then have "\<forall>x. \<phi> i x = \<psi> i x"
using phi_def r_psi by simp
then show ?thesis using i that by blast
qed
theorem smullyan_double_fixed_point:
assumes "g \<in> \<R>\<^sup>2" and "h \<in> \<R>\<^sup>2"
obtains m n where "\<phi> m = \<phi> (the (g m n))" and "\<phi> n = \<phi> (the (h m n))"
proof -
obtain rg where rg: "recfn 2 rg" "total rg" "g = (\<lambda>x y. eval rg [x, y])"
using R2E[OF assms(1)] by auto
moreover obtain rh where rh: "recfn 2 rh" "total rh" "h = (\<lambda>x y. eval rh [x, y])"
using R2E[OF assms(2)] by auto
ultimately obtain m n where
"\<forall>x. eval r_phi [m, x] = eval r_phi [the (eval rg [m, n]), x]"
"\<forall>x. eval r_phi [n, x] = eval r_phi [the (eval rh [m, n]), x]"
using smullyan_double_fixed_point_theorem[of rg rh] by blast
then have "\<phi> m = \<phi> (the (g m n))" and "\<phi> n = \<phi> (the (h m n))"
using phi_def rg rh by auto
then show ?thesis using that by simp
qed
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/TOTAL_CONS.thy b/thys/Inductive_Inference/TOTAL_CONS.thy
--- a/thys/Inductive_Inference/TOTAL_CONS.thy
+++ b/thys/Inductive_Inference/TOTAL_CONS.thy
@@ -1,1468 +1,1468 @@
section \<open>TOTAL is a proper subset of CONS\label{s:total_cons}\<close>
theory TOTAL_CONS
imports Lemma_R (* for r_auxhyp *)
CP_FIN_NUM (* for r_consistent *)
CONS_LIM (* for rmge2, goedel_at *)
begin
text \<open>We first show that TOTAL is a subset of CONS. Then we present a
separating class.\<close>
subsection \<open>TOTAL is a subset of CONS\<close>
text \<open>A TOTAL strategy hypothesizes only total functions, for which the
consistency with the input prefix is decidable. A CONS strategy can thus run
a TOTAL strategy and check if its hypothesis is consistent. If so, it
outputs this hypothesis, otherwise some arbitrary consistent one. Since the
TOTAL strategy converges to a correct hypothesis, which is consistent, the
CONS strategy will converge to the same hypothesis.\<close>
text \<open>Without loss of generality we can assume that learning takes place
with respect to our Gödel numbering $\varphi$. So we need to decide
consistency only for this numbering.\<close>
abbreviation r_consist_phi where
"r_consist_phi \<equiv> r_consistent r_phi"
lemma r_consist_phi_recfn [simp]: "recfn 2 r_consist_phi"
by simp
lemma r_consist_phi:
assumes "\<forall>k<e_length e. \<phi> i k \<down>"
shows "eval r_consist_phi [i, e] \<down>=
(if \<forall>k<e_length e. \<phi> i k \<down>= e_nth e k then 0 else 1)"
proof -
have "\<forall>k<e_length e. eval r_phi [i, k] \<down>"
using assms phi_def by simp
moreover have "recfn 2 r_phi" by simp
ultimately have "eval (r_consistent r_phi) [i, e] \<down>=
(if \<forall>k<e_length e. eval r_phi [i, k] \<down>= e_nth e k then 0 else 1)"
using r_consistent_converg assms by simp
then show ?thesis using phi_def by simp
qed
lemma r_consist_phi_init:
assumes "f \<in> \<R>" and "\<phi> i \<in> \<R>"
shows "eval r_consist_phi [i, f \<triangleright> n] \<down>= (if \<forall>k\<le>n. \<phi> i k = f k then 0 else 1)"
using assms r_consist_phi R1_imp_total1 total1E by (simp add: r_consist_phi)
lemma TOTAL_subseteq_CONS: "TOTAL \<subseteq> CONS"
proof
fix U assume "U \<in> TOTAL"
then have "U \<in> TOTAL_wrt \<phi>"
using TOTAL_wrt_phi_eq_TOTAL by blast
then obtain t' where t': "learn_total \<phi> U t'"
using TOTAL_wrt_def by auto
then obtain t where t: "recfn 1 t" "\<And>x. eval t [x] = t' x"
using learn_totalE(1) P1E by blast
then have t_converg: "eval t [f \<triangleright> n] \<down>" if "f \<in> U" for f n
using t' learn_totalE(1) that by auto
define s where "s \<equiv> Cn 1 r_ifz [Cn 1 r_consist_phi [t, Id 1 0], t, r_auxhyp]"
then have "recfn 1 s"
using r_consist_phi_recfn r_auxhyp_prim t(1) by simp
have consist: "eval r_consist_phi [the (eval t [f \<triangleright> n]), f \<triangleright> n] \<down>=
(if \<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k then 0 else 1)"
if "f \<in> U" for f n
proof -
have "eval r_consist_phi [the (eval t [f \<triangleright> n]), f \<triangleright> n] =
eval (Cn 1 r_consist_phi [t, Id 1 0]) [f \<triangleright> n]"
using that t_converg t(1) by simp
also have "... \<down>= (if \<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k then 0 else 1)"
proof -
from that have "f \<in> \<R>"
using learn_totalE(1) t' by blast
moreover have "\<phi> (the (eval t [f \<triangleright> n])) \<in> \<R>"
using t' t learn_totalE t_converg that by simp
ultimately show ?thesis
using r_consist_phi_init t_converg t(1) that by simp
qed
finally show ?thesis .
qed
have s_eq_t: "eval s [f \<triangleright> n] = eval t [f \<triangleright> n]"
if "\<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k" and "f \<in> U" for f n
using that consist s_def t r_auxhyp_prim prim_recfn_total
by simp
have s_eq_aux: "eval s [f \<triangleright> n] = eval r_auxhyp [f \<triangleright> n]"
if "\<not> (\<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k)" and "f \<in> U" for f n
proof -
from that have "eval r_consist_phi [the (eval t [f \<triangleright> n]), f \<triangleright> n] \<down>= 1"
using consist by simp
moreover have "t' (f \<triangleright> n) \<down>" using t' learn_totalE(1) that(2) by blast
ultimately show ?thesis
using s_def t r_auxhyp_prim t' learn_totalE by simp
qed
have "learn_cons \<phi> U (\<lambda>e. eval s [e])"
proof (rule learn_consI)
have "eval s [f \<triangleright> n] \<down>" if "f \<in> U" for f n
using that t_converg[OF that, of n] s_eq_t[of n f] prim_recfn_total[of r_auxhyp 1]
r_auxhyp_prim s_eq_aux[OF _ that, of n] totalE
by fastforce
then show "environment \<phi> U (\<lambda>e. eval s [e])"
- using t' `recfn 1 s` learn_totalE(1) by blast
+ using t' \<open>recfn 1 s\<close> learn_totalE(1) by blast
show "\<exists>i. \<phi> i = f \<and> (\<forall>\<^sup>\<infinity>n. eval s [f \<triangleright> n] \<down>= i)" if "f \<in> U" for f
proof -
from that t' t learn_totalE obtain i n\<^sub>0 where
i_n0: "\<phi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. eval t [f \<triangleright> n] \<down>= i)"
by metis
then have "\<And>n. n \<ge> n\<^sub>0 \<Longrightarrow> \<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k"
by simp
with s_eq_t have "\<And>n. n \<ge> n\<^sub>0 \<Longrightarrow> eval s [f \<triangleright> n] = eval t [f \<triangleright> n]"
using that by simp
with i_n0 have "\<And>n. n \<ge> n\<^sub>0 \<Longrightarrow> eval s [f \<triangleright> n] \<down>= i"
by auto
with i_n0 show ?thesis by auto
qed
show "\<forall>k\<le>n. \<phi> (the (eval s [f \<triangleright> n])) k = f k" if "f \<in> U" for f n
proof (cases "\<forall>k\<le>n. \<phi> (the (eval t [f \<triangleright> n])) k = f k")
case True
with that s_eq_t show ?thesis by simp
next
case False
then have "eval s [f \<triangleright> n] = eval r_auxhyp [f \<triangleright> n]"
using that s_eq_aux by simp
moreover have "f \<in> \<R>"
using learn_totalE(1)[OF t'] that by auto
ultimately show ?thesis using r_auxhyp by simp
qed
qed
then show "U \<in> CONS" using CONS_def by auto
qed
subsection \<open>The separating class\<close>
subsubsection \<open>Definition of the class\<close>
text \<open>The class that will be shown to be in @{term "CONS - TOTAL"} is
the union of the following two classes.\<close>
definition V_constotal_1 :: "partial1 set" where
"V_constotal_1 \<equiv> {f. \<exists>j p. f = [j] \<odot> p \<and> j \<ge> 2 \<and> p \<in> \<R>\<^sub>0\<^sub>1 \<and> \<phi> j = f}"
definition V_constotal_2 :: "partial1 set" where
"V_constotal_2 \<equiv>
{f. \<exists>j a k.
f = j # a @ [k] \<odot> 0\<^sup>\<infinity> \<and>
j \<ge> 2 \<and>
(\<forall>i<length a. a ! i \<le> 1) \<and>
k \<ge> 2 \<and>
\<phi> j = j # a \<odot> \<up>\<^sup>\<infinity> \<and>
\<phi> k = f}"
definition V_constotal :: "partial1 set" where
"V_constotal \<equiv> V_constotal_1 \<union> V_constotal_2"
lemma V_constotal_2I:
assumes "f = j # a @ [k] \<odot> 0\<^sup>\<infinity>"
and "j \<ge> 2"
and "\<forall>i<length a. a ! i \<le> 1"
and "k \<ge> 2"
and "\<phi> j = j # a \<odot> \<up>\<^sup>\<infinity>"
and "\<phi> k = f"
shows "f \<in> V_constotal_2"
using assms V_constotal_2_def by blast
lemma V_subseteq_R1: "V_constotal \<subseteq> \<R>"
proof
fix f assume "f \<in> V_constotal"
then have "f \<in> V_constotal_1 \<or> f \<in> V_constotal_2"
using V_constotal_def by auto
then show "f \<in> \<R>"
proof
assume "f \<in> V_constotal_1"
then obtain j p where "f = [j] \<odot> p" "p \<in> \<R>\<^sub>0\<^sub>1"
using V_constotal_1_def by blast
then show ?thesis using prepend_in_R1 RPred1_subseteq_R1 by auto
next
assume "f \<in> V_constotal_2"
then obtain j a k where "f = j # a @ [k] \<odot> 0\<^sup>\<infinity>"
using V_constotal_2_def by blast
then show ?thesis using almost0_in_R1 by auto
qed
qed
subsubsection \<open>The class is in CONS\<close>
text \<open>The class can be learned by the strategy @{term rmge2}, which
outputs the rightmost value greater or equal two in the input $f^n$. If $f$
is from $V_1$ then the strategy is correct right from the start. If $f$ is
from $V_2$ the strategy outputs the consistent hypothesis $j$ until it
encounters the correct hypothesis $k$, to which it converges.\<close>
lemma V_in_CONS: "learn_cons \<phi> V_constotal rmge2"
proof (rule learn_consI)
show "environment \<phi> V_constotal rmge2"
using V_subseteq_R1 rmge2_in_R1 R1_imp_total1 phi_in_P2 by simp
have "(\<exists>i. \<phi> i = f \<and> (\<forall>\<^sup>\<infinity>n. rmge2 (f \<triangleright> n) \<down>= i)) \<and>
(\<forall>n. \<forall>k\<le>n. \<phi> (the (rmge2 (f \<triangleright> n))) k = f k)"
if "f \<in> V_constotal" for f
proof (cases "f \<in> V_constotal_1")
case True
then obtain j p where
f: "f = [j] \<odot> p" and
j: "j \<ge> 2" and
p: "p \<in> \<R>\<^sub>0\<^sub>1" and
phi_j: "\<phi> j = f"
using V_constotal_1_def by blast
then have "f 0 \<down>= j" by (simp add: prepend_at_less)
then have f_at_0: "the (f 0) \<ge> 2" by (simp add: j)
have f_at_gr0: "the (f x) \<le> 1" if "x > 0" for x
using that f p by (simp add: RPred1_altdef Suc_leI prepend_at_ge)
have "total1 f"
using V_subseteq_R1 that R1_imp_total1 total1_def by auto
have "rmge2 (f \<triangleright> n) \<down>= j" for n
proof -
let ?P = "\<lambda>i. i < Suc n \<and> the (f i) \<ge> 2"
have "Greatest ?P = 0"
proof (rule Greatest_equality)
show "0 < Suc n \<and> 2 \<le> the (f 0)"
using f_at_0 by simp
show "\<And>y. y < Suc n \<and> 2 \<le> the (f y) \<Longrightarrow> y \<le> 0"
using f_at_gr0 by fastforce
qed
then have "rmge2 (f \<triangleright> n) = f 0"
- using f_at_0 rmge2_init_total[of f n, OF `total1 f`] by auto
+ using f_at_0 rmge2_init_total[of f n, OF \<open>total1 f\<close>] by auto
then show "rmge2 (f \<triangleright> n) \<down>= j"
by (simp add: \<open>f 0 \<down>= j\<close>)
qed
then show ?thesis using phi_j by auto
next
case False
then have "f \<in> V_constotal_2"
using V_constotal_def that by auto
then obtain j a k where jak:
"f = j # a @ [k] \<odot> 0\<^sup>\<infinity>"
"j \<ge> 2"
"\<forall>i<length a. a ! i \<le> 1"
"k \<ge> 2"
"\<phi> j = j # a \<odot> \<up>\<^sup>\<infinity> "
"\<phi> k = f"
using V_constotal_2_def by blast
then have f_at_0: "f 0 \<down>= j" by simp
have f_eq_a: "f x \<down>= a ! (x - 1)" if "0 < x \<and> x < Suc (length a)" for x
proof -
have "x - 1 < length a"
using that by auto
then show ?thesis
by (simp add: jak(1) less_SucI nth_append that)
qed
then have f_at_a: "the (f x) \<le> 1" if "0 < x \<and> x < Suc (length a)" for x
using jak(3) that by auto
from jak have f_k: "f (Suc (length a)) \<down>= k" by auto
from jak have f_at_big: "f x \<down>= 0" if "x > Suc (length a)" for x
using that by simp
let ?P = "\<lambda>n i. i < Suc n \<and> the (f i) \<ge> 2"
have rmge2: "rmge2 (f \<triangleright> n) = f (Greatest (?P n))" for n
proof -
have "\<not> (\<forall>i<Suc n. the (f i) < 2)" for n
using jak(2) f_at_0 by auto
moreover have "total1 f"
using V_subseteq_R1 R1_imp_total1 that total1_def by auto
ultimately show ?thesis using rmge2_init_total[of f n] by auto
qed
have "Greatest (?P n) = 0" if "n < Suc (length a)" for n
proof (rule Greatest_equality)
show "0 < Suc n \<and> 2 \<le> the (f 0)"
using that by (simp add: jak(2) f_at_0)
show "\<And>y. y < Suc n \<and> 2 \<le> the (f y) \<Longrightarrow> y \<le> 0"
using that f_at_a
by (metis Suc_1 dual_order.strict_trans leI less_Suc_eq not_less_eq_eq)
qed
with rmge2 f_at_0 have rmge2_small:
"rmge2 (f \<triangleright> n) \<down>= j" if "n < Suc (length a)" for n
using that by simp
have "Greatest (?P n) = Suc (length a)" if "n \<ge> Suc (length a)" for n
proof (rule Greatest_equality)
show "Suc (length a) < Suc n \<and> 2 \<le> the (f (Suc (length a)))"
using that f_k by (simp add: jak(4) less_Suc_eq_le)
show "\<And>y. y < Suc n \<and> 2 \<le> the (f y) \<Longrightarrow> y \<le> Suc (length a)"
using that f_at_big by (metis leI le_SucI not_less_eq_eq numeral_2_eq_2 option.sel)
qed
with rmge2 f_at_big f_k have rmge2_big:
"rmge2 (f \<triangleright> n) \<down>= k" if "n \<ge> Suc (length a)" for n
using that by simp
then have "\<exists>i n\<^sub>0. \<phi> i = f \<and> (\<forall>n\<ge>n\<^sub>0. rmge2 (f \<triangleright> n) \<down>= i)"
using jak(6) by auto
moreover have "\<forall>k\<le>n. \<phi> (the (rmge2 (f \<triangleright> n))) k = f k" for n
proof (cases "n < Suc (length a)")
case True
then have "rmge2 (f \<triangleright> n) \<down>= j"
using rmge2_small by simp
then have "\<phi> (the (rmge2 (f \<triangleright> n))) = \<phi> j" by simp
with True show ?thesis
using rmge2_small f_at_0 f_eq_a jak(5) prepend_at_less
by (metis le_less_trans le_zero_eq length_Cons not_le_imp_less nth_Cons_0 nth_Cons_pos)
next
case False
then show ?thesis using rmge2_big jak by simp
qed
ultimately show ?thesis by simp
qed
then show "\<And>f. f \<in> V_constotal \<Longrightarrow> \<exists>i. \<phi> i = f \<and> (\<forall>\<^sup>\<infinity>n. rmge2 (f \<triangleright> n) \<down>= i)"
and "\<And>f n. f \<in> V_constotal \<Longrightarrow> \<forall>k\<le>n. \<phi> (the (rmge2 (f \<triangleright> n))) k = f k"
by simp_all
qed
subsubsection \<open>The class is not in TOTAL\<close>
text \<open>Recall that $V$ is the union of $V_1 = \{jp \mid j\geq2 \land p \in
\mathcal{R}_{01} \land \varphi_j = jp\}$ and $V_2 = \{jak0^\infty \mid j\geq 2 \land a
\in \{0, 1\}^* \land k\geq 2 \land \varphi_j = ja\uparrow^\infty \land\
\varphi_k = jak0^\infty\}$.\<close>
text \<open>The proof is adapted from a proof of a stronger result by
Freivalds, Kinber, and Wiehagen~\cite[Theorem~27]{fkw-iisde-95} concerning an
inference type not defined here.
The proof is by contradiction. If $V$ was in TOTAL, there would be
a strategy $S$ learning $V$ in our standard Gödel numbering $\varphi$.
By Lemma R for TOTAL we can assume $S$ to be total.
In order to construct a function $f\in V$ for which $S$ fails we employ a
computable process iteratively building function prefixes. For every $j$ the
process builds a function $\psi_j$. The initial prefix is the singleton
$[j]$. Given a prefix $b$, the next prefix is determined as follows:
\begin{enumerate}
\item Search for a $y \geq |b|$ with $\varphi_{S(b)}(y) \downarrow= v$ for
some $v$.
\item Set the new prefix $b0^{y - |b|}\bar{v}$, where $\bar v = 1 - v$.
\end{enumerate}
Step~1 can diverge, for example, if $\varphi_{S(b)}$ is the empty function.
In this case $\psi_j$ will only be defined for a finite prefix. If, however,
Step~2 is reached, the prefix $b$ is extended to a $b'$ such that
$\varphi_{S(b)}(y) \neq b'_y$, which implies $S(b)$ is a wrong hypothesis for
every function starting with $b'$, in particular for $\psi_j$. Since $\bar v
\in \{0, 1\}$, Step~2 only appends zeros and ones, which is important for
showing membership in $V$.
This process defines a numbering $\psi \in \mathcal{P}^2$, and by Kleene's
fixed-point theorem there is a $j \geq 2$ with $\varphi_j = \psi_j$. For this
$j$ there are two cases:
\begin{enumerate}
\item[Case 1.] Step~1 always succeeds. Then $\psi_j$ is total and
$\psi_j \in V_1$. But $S$ outputs wrong hypotheses on infinitely many
prefixes of $\psi_j$ (namely every prefix constructed by the process).
\item[Case 2.] Step~1 diverges at some iteration, say when the state is $b = ja$
for some $a \in \{0, 1\}^*$.
Then $\psi_j$ has the form $ja\uparrow^\infty$. The numbering $\chi$ with $\chi_k =
jak0^\infty$ is in $\mathcal{P}^2$, and by Kleene's fixed-point theorem there is a
$k\geq 2$ with $\varphi_k = \chi_k = jak0^\infty$. This $jak0^\infty$ is in
$V_2$ and has the prefix $ja$. But Step~1 diverged on this prefix, which
means there is no $y \geq |ja|$ with $\varphi_{S(ja)}(y)\downarrow$. In
other words $S$ hypothesizes a non-total function.
\end{enumerate}
Thus, in both cases there is a function in $V$ where $S$ does not behave like
a TOTAL strategy. This is the desired contradiction.
The following locale formalizes this proof sketch.\<close>
locale total_cons =
fixes s :: partial1
assumes s_in_R1: "s \<in> \<R>"
begin
definition r_s :: recf where
"r_s \<equiv> SOME r_s. recfn 1 r_s \<and> total r_s \<and> s = (\<lambda>x. eval r_s [x])"
lemma rs_recfn [simp]: "recfn 1 r_s"
and rs_total [simp]: "\<And>x. eval r_s [x] \<down>"
and eval_rs: "\<And>x. s x = eval r_s [x]"
using r_s_def R1_SOME[OF s_in_R1, of r_s] by simp_all
text \<open>Performing Step~1 means enumerating the domain of
$\varphi_{S(b)}$ until a $y \geq |b|$ is found. The next function enumerates
all domain values and checks the condition for them.\<close>
definition "r_search_enum \<equiv>
Cn 2 r_le [Cn 2 r_length [Id 2 1], Cn 2 r_enumdom [Cn 2 r_s [Id 2 1], Id 2 0]]"
lemma r_search_enum_recfn [simp]: "recfn 2 r_search_enum"
by (simp add: r_search_enum_def Let_def)
abbreviation search_enum :: partial2 where
"search_enum x b \<equiv> eval r_search_enum [x, b]"
abbreviation enumdom :: partial2 where
"enumdom i y \<equiv> eval r_enumdom [i, y]"
lemma enumdom_empty_domain:
assumes "\<And>x. \<phi> i x \<up>"
shows "\<And>y. enumdom i y \<up>"
using assms r_enumdom_empty_domain by (simp add: phi_def)
lemma enumdom_nonempty_domain:
assumes "\<phi> i x\<^sub>0 \<down>"
shows "\<And>y. enumdom i y \<down>"
and "\<And>x. \<phi> i x \<down> \<longleftrightarrow> (\<exists>y. enumdom i y \<down>= x)"
using assms r_enumdom_nonempty_domain phi_def by metis+
text \<open>Enumerating the empty domain yields the empty function.\<close>
lemma search_enum_empty:
fixes b :: nat
assumes "s b \<down>= i" and "\<And>x. \<phi> i x \<up>"
shows "\<And>x. search_enum x b \<up>"
using assms r_search_enum_def enumdom_empty_domain eval_rs by simp
text \<open>Enumerating a non-empty domain yields a total function.\<close>
lemma search_enum_nonempty:
fixes b y0 :: nat
assumes "s b \<down>= i" and "\<phi> i y\<^sub>0 \<down>" and "e = the (enumdom i x)"
shows "search_enum x b \<down>= (if e_length b \<le> e then 0 else 1)"
proof -
let ?e = "\<lambda>x. the (enumdom i x)"
let ?y = "Cn 2 r_enumdom [Cn 2 r_s [Id 2 1], Id 2 0]"
have "recfn 2 ?y" using assms(1) by simp
moreover have "\<And>x. eval ?y [x, b] = enumdom i x"
using assms(1,2) eval_rs by auto
moreover from this have "\<And>x. eval ?y [x, b] \<down>"
using enumdom_nonempty_domain(1)[OF assms(2)] by simp
ultimately have "eval (Cn 2 r_le [Cn 2 r_length [Id 2 1], ?y]) [x, b] \<down>=
(if e_length b \<le> ?e x then 0 else 1)"
by simp
then show ?thesis using assms by (simp add: r_search_enum_def)
qed
text \<open>If there is a $y$ as desired, the enumeration will eventually return
zero (representing ``true'').\<close>
lemma search_enum_nonempty_eq0:
fixes b y :: nat
assumes "s b \<down>= i" and "\<phi> i y \<down>" and "y \<ge> e_length b"
shows "\<exists>x. search_enum x b \<down>= 0"
proof -
obtain x where x: "enumdom i x \<down>= y"
using enumdom_nonempty_domain(2)[OF assms(2)] assms(2) by auto
from assms(2) have "\<phi> i y \<down>" by simp
with x have "search_enum x b \<down>= 0"
using search_enum_nonempty[where ?e=y] assms by auto
then show ?thesis by auto
qed
text \<open>If there is no $y$ as desired, the enumeration will never return
zero.\<close>
lemma search_enum_nonempty_neq0:
fixes b y0 :: nat
assumes "s b \<down>= i"
and "\<phi> i y\<^sub>0 \<down>"
and "\<not> (\<exists>y. \<phi> i y \<down> \<and> y \<ge> e_length b)"
shows "\<not> (\<exists>x. search_enum x b \<down>= 0)"
proof
assume "\<exists>x. search_enum x b \<down>= 0"
then obtain x where x: "search_enum x b \<down>= 0"
by auto
obtain y where y: "enumdom i x \<down>= y"
using enumdom_nonempty_domain[OF assms(2)] by blast
then have "search_enum x b \<down>= (if e_length b \<le> y then 0 else 1)"
using assms(1-2) search_enum_nonempty by simp
with x have "e_length b \<le> y"
using option.inject by fastforce
moreover have "\<phi> i y \<down>"
using assms(2) enumdom_nonempty_domain(2) y by blast
ultimately show False using assms(3) by force
qed
text \<open>The next function corresponds to Step~1. Given a prefix $b$ it
computes a $y \geq |b|$ with $\varphi_{S(b)}(y)\downarrow$ if such a $y$
exists; otherwise it diverges.\<close>
definition "r_search \<equiv> Cn 1 r_enumdom [r_s, Mn 1 r_search_enum]"
lemma r_search_recfn [simp]: "recfn 1 r_search"
using r_search_def by simp
abbreviation search :: partial1 where
"search b \<equiv> eval r_search [b]"
text \<open>If $\varphi_{S(b)}$ is the empty function, the search process
diverges because already the enumeration of the domain diverges.\<close>
lemma search_empty:
assumes "s b \<down>= i" and "\<And>x. \<phi> i x \<up>"
shows "search b \<up>"
proof -
have "\<And>x. search_enum x b \<up>"
using search_enum_empty[OF assms] by simp
then have "eval (Mn 1 r_search_enum) [b] \<up>" by simp
then show "search b \<up>" unfolding r_search_def by simp
qed
text \<open>If $\varphi_{S(b)}$ is non-empty, but there is no $y$ with the
desired properties, the search process diverges.\<close>
lemma search_nonempty_neq0:
fixes b y0 :: nat
assumes "s b \<down>= i"
and "\<phi> i y\<^sub>0 \<down>"
and "\<not> (\<exists>y. \<phi> i y \<down> \<and> y \<ge> e_length b)"
shows "search b \<up>"
proof -
have "\<not> (\<exists>x. search_enum x b \<down>= 0)"
using assms search_enum_nonempty_neq0 by simp
moreover have "recfn 1 (Mn 1 r_search_enum)"
by (simp add: assms(1))
ultimately have "eval (Mn 1 r_search_enum) [b] \<up>" by simp
then show ?thesis using r_search_def by auto
qed
text \<open>If there is a $y$ as desired, the search process will return
one such $y$.\<close>
lemma search_nonempty_eq0:
fixes b y :: nat
assumes "s b \<down>= i" and "\<phi> i y \<down>" and "y \<ge> e_length b"
shows "search b \<down>"
and "\<phi> i (the (search b)) \<down>"
and "the (search b) \<ge> e_length b"
proof -
have "\<exists>x. search_enum x b \<down>= 0"
using assms search_enum_nonempty_eq0 by simp
moreover have "\<forall>x. search_enum x b \<down>"
using assms search_enum_nonempty by simp
moreover have "recfn 1 (Mn 1 r_search_enum)"
by simp
ultimately have
1: "search_enum (the (eval (Mn 1 r_search_enum) [b])) b \<down>= 0" and
2: "eval (Mn 1 r_search_enum) [b] \<down>"
using eval_Mn_diverg eval_Mn_convergE[of 1 "r_search_enum" "[b]"]
by (metis (no_types, lifting) One_nat_def length_Cons list.size(3) option.collapse,
metis (no_types, lifting) One_nat_def length_Cons list.size(3))
let ?x = "the (eval (Mn 1 r_search_enum) [b])"
have "search b = eval (Cn 1 r_enumdom [r_s, Mn 1 r_search_enum]) [b]"
unfolding r_search_def by simp
then have 3: "search b = enumdom i ?x"
using assms 2 eval_rs by simp
then have "the (search b) = the (enumdom i ?x)" (is "?y = _")
by simp
then have 4: "search_enum ?x b \<down>= (if e_length b \<le> ?y then 0 else 1)"
using search_enum_nonempty assms by simp
from 3 have "\<phi> i ?y \<down>"
using enumdom_nonempty_domain assms(2) by (metis option.collapse)
then show "\<phi> i ?y \<down>"
using phi_def by simp
then show "?y \<ge> e_length b"
using assms 4 1 option.inject by fastforce
show "search b \<down>"
using 3 assms(2) enumdom_nonempty_domain(1) by auto
qed
text \<open>The converse of the previous lemma states that whenever
the search process returns a value it will be one with the
desired properties.\<close>
lemma search_converg:
assumes "s b \<down>= i" and "search b \<down>" (is "?y \<down>")
shows "\<phi> i (the ?y) \<down>"
and "the ?y \<ge> e_length b"
proof -
have "\<exists>y. \<phi> i y \<down>"
using assms search_empty by meson
then have "\<exists>y. y \<ge> e_length b \<and> \<phi> i y \<down>"
using search_nonempty_neq0 assms by meson
then obtain y where y: "y \<ge> e_length b \<and> \<phi> i y \<down>" by auto
then have "\<phi> i y \<down>"
using phi_def by simp
then show "\<phi> i (the (search b)) \<down>"
and "(the (search b)) \<ge> e_length b"
- using y assms search_nonempty_eq0[OF assms(1) `\<phi> i y \<down>`] by simp_all
+ using y assms search_nonempty_eq0[OF assms(1) \<open>\<phi> i y \<down>\<close>] by simp_all
qed
text \<open>Likewise, if the search diverges, there is no appropriate $y$.\<close>
lemma search_diverg:
assumes "s b \<down>= i" and "search b \<up>"
shows "\<not> (\<exists>y. \<phi> i y \<down> \<and> y \<ge> e_length b)"
proof
assume "\<exists>y. \<phi> i y \<down> \<and> y \<ge> e_length b"
then obtain y where y: "\<phi> i y \<down>" "y \<ge> e_length b"
by auto
from y(1) have "\<phi> i y \<down>"
by (simp add: phi_def)
with y(2) search_nonempty_eq0 have "search b \<down>"
using assms by blast
with assms(2) show False by simp
qed
text \<open>Step~2 extends the prefix by a block of the shape $0^n\bar v$.
The next function constructs such a block for given $n$ and $v$.\<close>
definition "r_badblock \<equiv>
let f = Cn 1 r_singleton_encode [r_not];
g = Cn 3 r_cons [r_constn 2 0, Id 3 1]
in Pr 1 f g"
lemma r_badblock_prim [simp]: "recfn 2 r_badblock"
unfolding r_badblock_def by simp
lemma r_badblock: "eval r_badblock [n, v] \<down>= list_encode (replicate n 0 @ [1 - v])"
proof (induction n)
case 0
let ?f = "Cn 1 r_singleton_encode [r_not]"
have "eval r_badblock [0, v] = eval ?f [v]"
unfolding r_badblock_def by simp
also have "... = eval r_singleton_encode [the (eval r_not [v])]"
by simp
also have "... \<down>= list_encode [1 - v]"
by simp
finally show ?case by simp
next
case (Suc n)
let ?g = "Cn 3 r_cons [r_constn 2 0, Id 3 1]"
have "recfn 3 ?g" by simp
have "eval r_badblock [(Suc n), v] = eval ?g [n, the (eval r_badblock [n , v]), v]"
- using `recfn 3 ?g` Suc by (simp add: r_badblock_def)
+ using \<open>recfn 3 ?g\<close> Suc by (simp add: r_badblock_def)
also have "... = eval ?g [n, list_encode (replicate n 0 @ [1 - v]), v]"
using Suc by simp
also have "... = eval r_cons [0, list_encode (replicate n 0 @ [1 - v])]"
by simp
also have "... \<down>= e_cons 0 (list_encode (replicate n 0 @ [1 - v]))"
by simp
also have "... \<down>= list_encode (0 # (replicate n 0 @ [1 - v]))"
by simp
also have "... \<down>= list_encode (replicate (Suc n) 0 @ [1 - v])"
by simp
finally show ?case by simp
qed
lemma r_badblock_only_01: "e_nth (the (eval r_badblock [n, v])) i \<le> 1"
using r_badblock by (simp add: nth_append)
lemma r_badblock_last: "e_nth (the (eval r_badblock [n, v])) n = 1 - v"
using r_badblock by (simp add: nth_append)
text \<open>The following function computes the next prefix from the current
one. In other words, it performs Steps~1 and~2.\<close>
definition "r_next \<equiv>
Cn 1 r_append
[Id 1 0,
Cn 1 r_badblock
[Cn 1 r_sub [r_search, r_length],
Cn 1 r_phi [r_s, r_search]]]"
lemma r_next_recfn [simp]: "recfn 1 r_next"
unfolding r_next_def by simp
text \<open>The name @{text next} is unavailable, so we go for @{term nxt}.\<close>
abbreviation nxt :: partial1 where
"nxt b \<equiv> eval r_next [b]"
lemma nxt_diverg:
assumes "search b \<up>"
shows "nxt b \<up>"
unfolding r_next_def using assms by (simp add: Let_def)
lemma nxt_converg:
assumes "search b \<down>= y"
shows "nxt b \<down>=
e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\<phi> (the (s b)) y)]))"
unfolding r_next_def using assms r_badblock search_converg phi_def eval_rs
by fastforce
lemma nxt_search_diverg:
assumes "nxt b \<up>"
shows "search b \<up>"
proof (rule ccontr)
assume "search b \<down>"
then obtain y where "search b \<down>= y" by auto
then show False
using nxt_converg assms by simp
qed
text \<open>If Step~1 finds a $y$, the hypothesis $S(b)$ is incorrect for
the new prefix.\<close>
lemma nxt_wrong_hyp:
assumes "nxt b \<down>= b'" and "s b \<down>= i"
shows "\<exists>y<e_length b'. \<phi> i y \<down>\<noteq> e_nth b' y"
proof -
obtain y where y: "search b \<down>= y"
using assms nxt_diverg by fastforce
then have y_len: "y \<ge> e_length b"
using assms search_converg(2) by fastforce
then have b': "b' =
(e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\<phi> i y)])))"
using y assms nxt_converg by simp
then have "e_nth b' y = 1 - the (\<phi> i y)"
using y_len e_nth_append_big r_badblock r_badblock_last by auto
moreover have "\<phi> i y \<down>"
using search_converg y y_len assms(2) by fastforce
ultimately have "\<phi> i y \<down>\<noteq> e_nth b' y"
by (metis gr_zeroI less_numeral_extra(4) less_one option.sel zero_less_diff)
moreover have "e_length b' = Suc y"
using y_len e_length_append b' by auto
ultimately show ?thesis by auto
qed
text \<open>If Step~1 diverges, the hypothesis $S(b)$ refers to a non-total
function.\<close>
lemma nxt_nontotal_hyp:
assumes "nxt b \<up>" and "s b \<down>= i"
shows "\<exists>x. \<phi> i x \<up>"
using nxt_search_diverg[OF assms(1)] search_diverg[OF assms(2)] by auto
text \<open>The process only ever extends the given prefix.\<close>
lemma nxt_stable:
assumes "nxt b \<down>= b'"
shows "\<forall>x<e_length b. e_nth b x = e_nth b' x"
proof -
obtain y where y: "search b \<down>= y"
using assms nxt_diverg by fastforce
then have "y \<ge> e_length b"
using search_converg(2) eval_rs rs_total by fastforce
show ?thesis
proof (rule allI, rule impI)
fix x assume "x < e_length b"
let ?i = "the (s b)"
have b': "b' =
(e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\<phi> ?i y)])))"
using assms nxt_converg[OF y] by auto
then show "e_nth b x = e_nth b' x"
using e_nth_append_small \<open>x < e_length b\<close> by auto
qed
qed
text \<open>The following properties of @{term r_next} will be
used to show that some of the constructed functions are in the class
$V$.\<close>
lemma nxt_append_01:
assumes "nxt b \<down>= b'"
shows "\<forall>x. x \<ge> e_length b \<and> x < e_length b' \<longrightarrow> e_nth b' x = 0 \<or> e_nth b' x = 1"
proof -
obtain y where y: "search b \<down>= y"
using assms nxt_diverg by fastforce
let ?i = "the (s b)"
have b': "b' = (e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\<phi> ?i y)])))"
(is "b' = (e_append b ?z)")
using assms y nxt_converg prod_encode_eq by auto
show ?thesis
proof (rule allI, rule impI)
fix x assume x: "e_length b \<le> x \<and> x < e_length b'"
then have "e_nth b' x = e_nth ?z (x - e_length b)"
using b' e_nth_append_big by blast
then show "e_nth b' x = 0 \<or> e_nth b' x = 1"
by (metis less_one nat_less_le option.sel r_badblock r_badblock_only_01)
qed
qed
lemma nxt_monotone:
assumes "nxt b \<down>= b'"
shows "e_length b < e_length b'"
proof -
obtain y where y: "search b \<down>= y"
using assms nxt_diverg by fastforce
let ?i = "the (s b)"
have b': "b' =
(e_append b (list_encode (replicate (y - e_length b) 0 @ [1 - the (\<phi> ?i y)])))"
using assms y nxt_converg prod_encode_eq by auto
then show ?thesis using e_length_append by auto
qed
text \<open>The next function computes the prefixes after each iteration of
the process @{term r_next} when started with the list $[j]$.\<close>
definition r_prefixes :: recf where
"r_prefixes \<equiv> Pr 1 r_singleton_encode (Cn 3 r_next [Id 3 1])"
lemma r_prefixes_recfn [simp]: "recfn 2 r_prefixes"
unfolding r_prefixes_def by (simp add: Let_def)
abbreviation prefixes :: partial2 where
"prefixes t j \<equiv> eval r_prefixes [t, j]"
lemma prefixes_at_0: "prefixes 0 j \<down>= list_encode [j]"
unfolding r_prefixes_def by simp
lemma prefixes_at_Suc:
assumes "prefixes t j \<down>" (is "?b \<down>")
shows "prefixes (Suc t) j = nxt (the ?b)"
using r_prefixes_def assms by auto
lemma prefixes_at_Suc':
assumes "prefixes t j \<down>= b"
shows "prefixes (Suc t) j = nxt b"
using r_prefixes_def assms by auto
lemma prefixes_prod_encode:
assumes "prefixes t j \<down>"
obtains b where "prefixes t j \<down>= b"
using assms surj_prod_encode by force
lemma prefixes_converg_le:
assumes "prefixes t j \<down>" and "t' \<le> t"
shows "prefixes t' j \<down>"
using r_prefixes_def assms eval_Pr_converg_le[of 1 _ _ "[j]"]
by simp
lemma prefixes_diverg_add:
assumes "prefixes t j \<up>"
shows "prefixes (t + d) j \<up>"
using r_prefixes_def assms eval_Pr_diverg_add[of 1 _ _ "[j]"]
by simp
text \<open>Many properties of @{term r_prefixes} can be derived from similar
properties of @{term r_next}.\<close>
lemma prefixes_length:
assumes "prefixes t j \<down>= b"
shows "e_length b > t"
proof (insert assms, induction t arbitrary: b)
case 0
then show ?case using prefixes_at_0 prod_encode_eq by auto
next
case (Suc t)
then have "prefixes t j \<down>"
using prefixes_converg_le Suc_n_not_le_n nat_le_linear by blast
then obtain b' where b': "prefixes t j \<down>= b'"
using prefixes_prod_encode by blast
with Suc have "e_length b' > t" by simp
have "prefixes (Suc t) j = nxt b'"
using b' prefixes_at_Suc' by simp
with Suc have "nxt b' \<down>= b" by simp
then have "e_length b' < e_length b"
using nxt_monotone by simp
- then show ?case using `e_length b' > t` by simp
+ then show ?case using \<open>e_length b' > t\<close> by simp
qed
lemma prefixes_monotone:
assumes "prefixes t j \<down>= b" and "prefixes (t + d) j \<down>= b'"
shows "e_length b \<le> e_length b'"
proof (insert assms, induction d arbitrary: b')
case 0
then show ?case using prod_encode_eq by simp
next
case (Suc d)
moreover have "t + d \<le> t + Suc d" by simp
ultimately have "prefixes (t + d) j \<down>"
using prefixes_converg_le by blast
then obtain b'' where b'': "prefixes (t + d) j \<down>= b''"
using prefixes_prod_encode by blast
with Suc have "prefixes (t + Suc d) j = nxt b''"
by (simp add: prefixes_at_Suc')
with Suc have "nxt b'' \<down>= b'" by simp
then show ?case using nxt_monotone Suc b'' by fastforce
qed
lemma prefixes_stable:
assumes "prefixes t j \<down>= b" and "prefixes (t + d) j \<down>= b'"
shows "\<forall>x<e_length b. e_nth b x = e_nth b' x"
proof (insert assms, induction d arbitrary: b')
case 0
then show ?case using prod_encode_eq by simp
next
case (Suc d)
moreover have "t + d \<le> t + Suc d" by simp
ultimately have "prefixes (t + d) j \<down>"
using prefixes_converg_le by blast
then obtain b'' where b'': "prefixes (t + d) j \<down>= b''"
using prefixes_prod_encode by blast
with Suc have "prefixes (t + Suc d) j = nxt b''"
by (simp add: prefixes_at_Suc')
with Suc have b': "nxt b'' \<down>= b'" by simp
show "\<forall>x<e_length b. e_nth b x = e_nth b' x"
proof (rule allI, rule impI)
fix x assume x: "x < e_length b"
then have "e_nth b x = e_nth b'' x"
using Suc b'' by simp
moreover have "x \<le> e_length b''"
using x prefixes_monotone b'' Suc by fastforce
ultimately show "e_nth b x = e_nth b' x"
using b'' nxt_stable Suc b' prefixes_monotone x
by (metis leD le_neq_implies_less)
qed
qed
lemma prefixes_tl_only_01:
assumes "prefixes t j \<down>= b"
shows "\<forall>x>0. e_nth b x = 0 \<or> e_nth b x = 1"
proof (insert assms, induction t arbitrary: b)
case 0
then show ?case using prefixes_at_0 prod_encode_eq by auto
next
case (Suc t)
then have "prefixes t j \<down>"
using prefixes_converg_le Suc_n_not_le_n nat_le_linear by blast
then obtain b' where b': "prefixes t j \<down>= b'"
using prefixes_prod_encode by blast
show "\<forall>x>0. e_nth b x = 0 \<or> e_nth b x = 1"
proof (rule allI, rule impI)
fix x :: nat
assume x: "x > 0"
show "e_nth b x = 0 \<or> e_nth b x = 1"
proof (cases "x < e_length b'")
case True
then show ?thesis
using Suc b' prefixes_at_Suc' nxt_stable x by metis
next
case False
then show ?thesis
using Suc.prems b' prefixes_at_Suc' nxt_append_01 by auto
qed
qed
qed
lemma prefixes_hd:
assumes "prefixes t j \<down>= b"
shows "e_nth b 0 = j"
proof -
obtain b' where b': "prefixes 0 j \<down>= b'"
by (simp add: prefixes_at_0)
then have "b' = list_encode [j]"
by (simp add: prod_encode_eq prefixes_at_0)
then have "e_nth b' 0 = j" by simp
then show "e_nth b 0 = j"
using assms prefixes_stable[OF b', of t b] prefixes_length[OF b'] by simp
qed
lemma prefixes_nontotal_hyp:
assumes "prefixes t j \<down>= b"
and "prefixes (Suc t) j \<up>"
and "s b \<down>= i"
shows "\<exists>x. \<phi> i x \<up>"
using nxt_nontotal_hyp[OF _ assms(3)] assms(2) prefixes_at_Suc'[OF assms(1)] by simp
text \<open>We now consider the two cases from the proof sketch.\<close>
abbreviation "case_two j \<equiv> \<exists>t. prefixes t j \<up>"
abbreviation "case_one j \<equiv> \<not> case_two j"
text \<open>In Case~2 there is a maximum convergent iteration because
iteration 0 converges.\<close>
lemma case_two:
assumes "case_two j"
shows "\<exists>t. (\<forall>t'\<le>t. prefixes t' j \<down>) \<and> (\<forall>t'>t. prefixes t' j \<up>)"
proof -
let ?P = "\<lambda>t. prefixes t j \<up>"
define t\<^sub>0 where "t\<^sub>0 = Least ?P"
then have "?P t\<^sub>0"
using assms LeastI_ex[of ?P] by simp
then have diverg: "?P t" if "t \<ge> t\<^sub>0" for t
using prefixes_converg_le that by blast
from t\<^sub>0_def have converg: "\<not> ?P t" if "t < t\<^sub>0" for t
using Least_le[of ?P] that not_less by blast
have "t\<^sub>0 > 0"
proof (rule ccontr)
assume "\<not> 0 < t\<^sub>0"
then have "t\<^sub>0 = 0" by simp
- with `?P t\<^sub>0` prefixes_at_0 show False by simp
+ with \<open>?P t\<^sub>0\<close> prefixes_at_0 show False by simp
qed
let ?t = "t\<^sub>0 - 1"
have "\<forall>t'\<le>?t. prefixes t' j \<down>"
using converg \<open>0 < t\<^sub>0\<close> by auto
moreover have "\<forall>t'>?t. prefixes t' j \<up>"
using diverg by simp
ultimately show ?thesis by auto
qed
text \<open>Having completed the modelling of the process, we can now define
the functions $\psi_j$ it computes. The value $\psi_j(x)$ is computed by
running @{term r_prefixes} until the prefix is longer than $x$ and then
taking the $x$-th element of the prefix.\<close>
definition "r_psi \<equiv>
let f = Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]]
in Cn 2 r_nth [Cn 2 r_prefixes [Mn 2 f, Id 2 0], Id 2 1]"
lemma r_psi_recfn: "recfn 2 r_psi"
unfolding r_psi_def by simp
abbreviation psi :: partial2 ("\<psi>") where
"\<psi> j x \<equiv> eval r_psi [j, x]"
lemma psi_in_P2: "\<psi> \<in> \<P>\<^sup>2"
using r_psi_recfn by auto
text \<open>The values of @{term "\<psi>"} can be read off the prefixes.\<close>
lemma psi_eq_nth_prefix:
assumes "prefixes t j \<down>= b" and "e_length b > x"
shows "\<psi> j x \<down>= e_nth b x"
proof -
let ?f = "Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]]"
let ?P = "\<lambda>t. prefixes t j \<down> \<and> e_length (the (prefixes t j)) > x"
from assms have ex_t: "\<exists>t. ?P t" by auto
define t\<^sub>0 where "t\<^sub>0 = Least ?P"
then have "?P t\<^sub>0"
using LeastI_ex[OF ex_t] by simp
from ex_t have not_P: "\<not> ?P t" if "t < t\<^sub>0" for t
using ex_t that Least_le[of ?P] not_le t\<^sub>0_def by auto
have "?P t" using assms by simp
with not_P have "t\<^sub>0 \<le> t" using leI by blast
then obtain b\<^sub>0 where b0: "prefixes t\<^sub>0 j \<down>= b\<^sub>0"
using assms(1) prefixes_converg_le by blast
have "eval ?f [t\<^sub>0, j, x] \<down>= 0"
proof -
have "eval (Cn 3 r_prefixes [Id 3 0, Id 3 1]) [t\<^sub>0, j, x] \<down>= b\<^sub>0"
using b0 by simp
- then show ?thesis using `?P t\<^sub>0` by simp
+ then show ?thesis using \<open>?P t\<^sub>0\<close> by simp
qed
moreover have "eval ?f [t, j, x] \<down>\<noteq> 0" if "t < t\<^sub>0" for t
proof -
obtain bt where bt: "prefixes t j \<down>= bt"
- using prefixes_converg_le[of t\<^sub>0 j t] b0 `t < t\<^sub>0` by auto
+ using prefixes_converg_le[of t\<^sub>0 j t] b0 \<open>t < t\<^sub>0\<close> by auto
moreover have "\<not> ?P t"
using that not_P by simp
ultimately have "e_length bt \<le> x" by simp
moreover have "eval (Cn 3 r_prefixes [Id 3 0, Id 3 1]) [t, j, x] \<down>= bt"
using bt by simp
ultimately show ?thesis by simp
qed
ultimately have "eval (Mn 2 ?f) [j, x] \<down>= t\<^sub>0"
using eval_Mn_convergI[of 2 ?f "[j, x]" t\<^sub>0] by simp
then have "\<psi> j x \<down>= e_nth b\<^sub>0 x"
unfolding r_psi_def using b0 by simp
then show ?thesis
- using `t\<^sub>0 \<le> t` assms(1) prefixes_stable[of t\<^sub>0 j b\<^sub>0 "t - t\<^sub>0" b] b0 `?P t\<^sub>0`
+ using \<open>t\<^sub>0 \<le> t\<close> assms(1) prefixes_stable[of t\<^sub>0 j b\<^sub>0 "t - t\<^sub>0" b] b0 \<open>?P t\<^sub>0\<close>
by simp
qed
lemma psi_converg_imp_prefix:
assumes "\<psi> j x \<down>"
shows "\<exists>t b. prefixes t j \<down>= b \<and> e_length b > x"
proof -
let ?f = "Cn 3 r_less [Id 3 2, Cn 3 r_length [Cn 3 r_prefixes [Id 3 0, Id 3 1]]]"
have "eval (Mn 2 ?f) [j, x] \<down>"
proof (rule ccontr)
assume "\<not> eval (Mn 2 ?f) [j, x] \<down>"
then have "eval (Mn 2 ?f) [j, x] \<up>" by simp
then have "\<psi> j x \<up>"
unfolding r_psi_def by simp
then show False
using assms by simp
qed
then obtain t where t: "eval (Mn 2 ?f) [j, x] \<down>= t"
by blast
have "recfn 2 (Mn 2 ?f)" by simp
then have f_zero: "eval ?f [t, j, x] \<down>= 0"
using eval_Mn_convergE[OF _ t]
by (metis (no_types, lifting) One_nat_def Suc_1 length_Cons list.size(3))
have "prefixes t j \<down>"
proof (rule ccontr)
assume "\<not> prefixes t j \<down>"
then have "prefixes t j \<up>" by simp
then have "eval ?f [t, j, x] \<up>" by simp
with f_zero show False by simp
qed
then obtain b' where b': "prefixes t j \<down>= b'" by auto
moreover have "e_length b' > x"
proof (rule ccontr)
assume "\<not> e_length b' > x"
then have "eval ?f [t, j, x] \<down>= 1"
using b' by simp
with f_zero show False by simp
qed
ultimately show ?thesis by auto
qed
lemma psi_converg_imp_prefix':
assumes "\<psi> j x \<down>"
shows "\<exists>t b. prefixes t j \<down>= b \<and> e_length b > x \<and> \<psi> j x \<down>= e_nth b x"
using psi_converg_imp_prefix[OF assms] psi_eq_nth_prefix by blast
text \<open>In both Case~1 and~2, $\psi_j$ starts with $j$.\<close>
lemma psi_at_0: "\<psi> j 0 \<down>= j"
using prefixes_hd prefixes_length psi_eq_nth_prefix prefixes_at_0 by fastforce
text \<open>In Case~1, $\psi_j$ is total and made up of $j$ followed by zeros
and ones, just as required by the definition of $V_1$.\<close>
lemma case_one_psi_total:
assumes "case_one j" and "x > 0"
shows "\<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1"
proof -
obtain b where b: "prefixes x j \<down>= b"
using assms(1) by auto
then have "e_length b > x"
using prefixes_length by simp
then have "\<psi> j x \<down>= e_nth b x"
using b psi_eq_nth_prefix by simp
moreover have "e_nth b x = 0 \<or> e_nth b x = 1"
using prefixes_tl_only_01[OF b] assms(2) by simp
ultimately show "\<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1"
by simp
qed
text \<open>In Case~2, $\psi_j$ is defined only for a prefix starting with
$j$ and continuing with zeros and ones. This prefix corresponds to $ja$ from
the definition of $V_2$.\<close>
lemma case_two_psi_only_prefix:
assumes "case_two j"
shows "\<exists>y. (\<forall>x. 0 < x \<and> x < y \<longrightarrow> \<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1) \<and>
(\<forall>x \<ge> y. \<psi> j x \<up>)"
proof -
obtain t where
t_le: "\<forall>t'\<le>t. prefixes t' j \<down>" and
t_gr: "\<forall>t'>t. prefixes t' j \<up>"
using assms case_two by blast
then obtain b where b: "prefixes t j \<down>= b"
by auto
let ?y = "e_length b"
have "\<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1" if "x > 0 \<and> x < ?y" for x
using t_le b that by (metis prefixes_tl_only_01 psi_eq_nth_prefix)
moreover have "\<psi> j x \<up>" if "x \<ge> ?y" for x
proof (rule ccontr)
assume "\<psi> j x \<down>"
then obtain t' b' where t': "prefixes t' j \<down>= b'" and "e_length b' > x"
using psi_converg_imp_prefix by blast
then have "e_length b' > ?y"
using that by simp
with t' have "t' > t"
using prefixes_monotone b by (metis add_diff_inverse_nat leD)
with t' t_gr show False by simp
qed
ultimately show ?thesis by auto
qed
definition longest_prefix :: "nat \<Rightarrow> nat" where
"longest_prefix j \<equiv> THE y. (\<forall>x<y. \<psi> j x \<down>) \<and> (\<forall>x\<ge>y. \<psi> j x \<up>)"
lemma longest_prefix:
assumes "case_two j" and "z = longest_prefix j"
shows "(\<forall>x<z. \<psi> j x \<down>) \<and> (\<forall>x\<ge>z. \<psi> j x \<up>)"
proof -
let ?P = "\<lambda>z. (\<forall>x<z. \<psi> j x \<down>) \<and> (\<forall>x\<ge>z. \<psi> j x \<up>)"
obtain y where y:
"\<forall>x. 0 < x \<and> x < y \<longrightarrow> \<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1"
"\<forall>x\<ge>y. \<psi> j x \<up>"
using case_two_psi_only_prefix[OF assms(1)] by auto
have "?P (THE z. ?P z)"
proof (rule theI[of ?P y])
show "?P y"
proof
show "\<forall>x<y. \<psi> j x \<down>"
proof (rule allI, rule impI)
fix x assume "x < y"
show "\<psi> j x \<down>"
proof (cases "x = 0")
case True
then show ?thesis using psi_at_0 by simp
next
case False
- then show ?thesis using y(1) `x < y` by auto
+ then show ?thesis using y(1) \<open>x < y\<close> by auto
qed
qed
show "\<forall>x\<ge>y. \<psi> j x \<up>" using y(2) by simp
qed
show "z = y" if "?P z" for z
proof (rule ccontr, cases "z < y")
case True
moreover assume "z \<noteq> y"
ultimately show False
- using that `?P y` by auto
+ using that \<open>?P y\<close> by auto
next
case False
moreover assume "z \<noteq> y"
then show False
- using that `?P y` y(2) by (meson linorder_cases order_refl)
+ using that \<open>?P y\<close> y(2) by (meson linorder_cases order_refl)
qed
qed
then have "(\<forall>x<(THE z. ?P z). \<psi> j x \<down>) \<and> (\<forall>x\<ge>(THE z. ?P z). \<psi> j x \<up>)"
by blast
moreover have "longest_prefix j = (THE z. ?P z)"
unfolding longest_prefix_def by simp
ultimately show ?thesis using assms(2) by metis
qed
lemma case_two_psi_longest_prefix:
assumes "case_two j" and "y = longest_prefix j"
shows "(\<forall>x. 0 < x \<and> x < y \<longrightarrow> \<psi> j x \<down>= 0 \<or> \<psi> j x \<down>= 1) \<and>
(\<forall>x \<ge> y. \<psi> j x \<up>)"
using assms longest_prefix case_two_psi_only_prefix
by (metis prefixes_tl_only_01 psi_converg_imp_prefix')
text \<open>The prefix cannot be empty because the process starts with prefix $[j]$.\<close>
lemma longest_prefix_gr_0:
assumes "case_two j"
shows "longest_prefix j > 0"
using assms case_two_psi_longest_prefix psi_at_0 by force
lemma psi_not_divergent_init:
assumes "prefixes t j \<down>= b"
shows "(\<psi> j) \<triangleright> (e_length b - 1) = b"
proof (intro initI)
show "0 < e_length b"
using assms prefixes_length by fastforce
show "\<psi> j x \<down>= e_nth b x" if "x < e_length b" for x
using that assms psi_eq_nth_prefix by simp
qed
text \<open>In Case~2, the strategy $S$ outputs a non-total hypothesis on
some prefix of $\psi_j$.\<close>
lemma case_two_nontotal_hyp:
assumes "case_two j"
shows "\<exists>n<longest_prefix j. \<not> total1 (\<phi> (the (s ((\<psi> j) \<triangleright> n))))"
proof -
obtain t where "\<forall>t'\<le>t. prefixes t' j \<down>" and t_gr: "\<forall>t'>t. prefixes t' j \<up>"
using assms case_two by blast
then obtain b where b: "prefixes t j \<down>= b"
by auto
moreover obtain i where i: "s b \<down>= i"
using eval_rs by fastforce
moreover have div: "prefixes (Suc t) j \<up>"
using t_gr by simp
ultimately have "\<exists>x. \<phi> i x \<up>"
using prefixes_nontotal_hyp by simp
then obtain x where "\<phi> i x \<up>" by auto
moreover have init: "\<psi> j \<triangleright> (e_length b - 1) = b" (is "_ \<triangleright> ?n = b")
using psi_not_divergent_init[OF b] by simp
ultimately have "\<phi> (the (s (\<psi> j \<triangleright> ?n))) x \<up>"
using i by simp
then have "\<not> total1 (\<phi> (the (s (\<psi> j \<triangleright> ?n))))"
by auto
moreover have "?n < longest_prefix j"
using case_two_psi_longest_prefix init b div psi_eq_nth_prefix
by (metis length_init lessI not_le_imp_less option.simps(3))
ultimately show ?thesis by auto
qed
text \<open>Consequently, in Case~2 the strategy does not TOTAL-learn
any function starting with the longest prefix of $\psi_j$.\<close>
lemma case_two_not_learn:
assumes "case_two j"
and "f \<in> \<R>"
and "\<And>x. x < longest_prefix j \<Longrightarrow> f x = \<psi> j x"
shows "\<not> learn_total \<phi> {f} s"
proof -
obtain n where n:
"n < longest_prefix j"
"\<not> total1 (\<phi> (the (s (\<psi> j \<triangleright> n))))"
using case_two_nontotal_hyp[OF assms(1)] by auto
have "f \<triangleright> n = \<psi> j \<triangleright> n"
using assms(3) n(1) by (intro init_eqI) auto
with n(2) show ?thesis by (metis R1_imp_total1 learn_totalE(3) singletonI)
qed
text \<open>In Case~1 the strategy outputs a wrong hypothesis
on infinitely many prefixes of $\psi_j$ and thus does not
learn $\psi_j$ in the limit, much less in the sense of TOTAL.\<close>
lemma case_one_wrong_hyp:
assumes "case_one j"
shows "\<exists>n>k. \<phi> (the (s ((\<psi> j) \<triangleright> n))) \<noteq> \<psi> j"
proof -
have all_t: "\<forall>t. prefixes t j \<down>"
using assms by simp
then obtain b where b: "prefixes (Suc k) j \<down>= b"
by auto
then have length: "e_length b > Suc k"
using prefixes_length by simp
then have init: "\<psi> j \<triangleright> (e_length b - 1) = b"
using psi_not_divergent_init b by simp
obtain i where i: "s b \<down>= i"
using eval_rs by fastforce
from all_t obtain b' where b': "prefixes (Suc (Suc k)) j \<down>= b'"
by auto
then have "\<psi> j \<triangleright> (e_length b' - 1) = b'"
using psi_not_divergent_init by simp
moreover have "\<exists>y<e_length b'. \<phi> i y \<down>\<noteq> e_nth b' y"
using nxt_wrong_hyp b b' i prefixes_at_Suc by auto
ultimately have "\<exists>y<e_length b'. \<phi> i y \<noteq> \<psi> j y"
using b' psi_eq_nth_prefix by auto
then have "\<phi> i \<noteq> \<psi> j" by auto
then show ?thesis
using init length i by (metis Suc_less_eq length_init option.sel)
qed
lemma case_one_not_learn:
assumes "case_one j"
shows "\<not> learn_lim \<phi> {\<psi> j} s"
proof (rule infinite_hyp_wrong_not_Lim[of "\<psi> j"])
show "\<psi> j \<in> {\<psi> j}" by simp
show "\<forall>n. \<exists>m>n. \<phi> (the (s (\<psi> j \<triangleright> m))) \<noteq> \<psi> j"
using case_one_wrong_hyp[OF assms] by simp
qed
lemma case_one_not_learn_V:
assumes "case_one j" and "j \<ge> 2" and "\<phi> j = \<psi> j"
shows "\<not> learn_lim \<phi> V_constotal s"
proof -
have "\<psi> j \<in> V_constotal_1"
proof -
define p where "p = (\<lambda>x. (\<psi> j) (x + 1))"
have "p \<in> \<R>\<^sub>0\<^sub>1"
proof -
from p_def have "p \<in> \<P>"
using skip_P1[of "\<psi> j" 1] psi_in_P2 P2_proj_P1 by blast
moreover have "p x \<down>= 0 \<or> p x \<down>= 1" for x
using p_def assms(1) case_one_psi_total by auto
moreover from this have "total1 p" by fast
ultimately show ?thesis using RPred1_def by auto
qed
moreover have "\<psi> j = [j] \<odot> p"
by (intro prepend_eqI, simp add: psi_at_0, simp add: p_def)
ultimately show ?thesis using assms(2,3) V_constotal_1_def by blast
qed
then have "\<psi> j \<in> V_constotal" using V_constotal_def by auto
moreover have "\<not> learn_lim \<phi> {\<psi> j} s"
using case_one_not_learn assms(1) by simp
ultimately show ?thesis using learn_lim_closed_subseteq by auto
qed
text \<open>The next lemma embodies the construction of $\chi$ followed by
the application of Kleene's fixed-point theorem as described in the
proof sketch.\<close>
lemma goedel_after_prefixes:
fixes vs :: "nat list" and m :: nat
shows "\<exists>n\<ge>m. \<phi> n = vs @ [n] \<odot> 0\<^sup>\<infinity>"
proof -
define f :: partial1 where "f \<equiv> vs \<odot> 0\<^sup>\<infinity>"
then have "f \<in> \<R>"
using almost0_in_R1 by auto
then obtain n where n:
"n \<ge> m"
"\<phi> n = (\<lambda>x. if x = length vs then Some n else f x)"
using goedel_at[of f m "length vs"] by auto
moreover have "\<phi> n x = (vs @ [n] \<odot> 0\<^sup>\<infinity>) x" for x
proof -
consider "x < length vs" | "x = length vs" | "x > length vs"
by linarith
then show ?thesis
using n f_def by (cases) (auto simp add: prepend_associative)
qed
ultimately show ?thesis by blast
qed
text \<open>If Case~2 holds for a $j\geq 2$ with $\varphi_j = \psi_j$, that
is, if $\psi_j\in V_1$, then there is a function in $V$, namely $\psi_j$, on
which $S$ fails. Therefore $S$ does not learn $V$.\<close>
lemma case_two_not_learn_V:
assumes "case_two j" and "j \<ge> 2" and "\<phi> j = \<psi> j"
shows "\<not> learn_total \<phi> V_constotal s"
proof -
define z where "z = longest_prefix j"
then have "z > 0"
using longest_prefix_gr_0[OF assms(1)] by simp
define vs where "vs = prefix (\<psi> j) (z - 1)"
then have "vs ! 0 = j"
- using psi_at_0 `z > 0` by simp
+ using psi_at_0 \<open>z > 0\<close> by simp
define a where "a = tl vs"
then have vs: "vs = j # a"
- using vs_def `vs ! 0 = j`
+ using vs_def \<open>vs ! 0 = j\<close>
by (metis length_Suc_conv length_prefix list.sel(3) nth_Cons_0)
obtain k where k: "k \<ge> 2" and phi_k: "\<phi> k = j # a @ [k] \<odot> 0\<^sup>\<infinity>"
using goedel_after_prefixes[of 2 "j # a"] by auto
have phi_j: "\<phi> j = j # a \<odot> \<up>\<^sup>\<infinity> "
proof (rule prepend_eqI)
show "\<And>x. x < length (j # a) \<Longrightarrow> \<phi> j x \<down>= (j # a) ! x"
using assms(1,3) vs vs_def \<open>0 < z\<close>
length_prefix[of "\<psi> j" "z - 1"]
prefix_nth[of _ _ "\<psi> j"]
psi_at_0[of j]
case_two_psi_longest_prefix[OF _ z_def]
longest_prefix[OF _ z_def]
by (metis One_nat_def Suc_pred option.collapse)
show "\<And>x. \<phi> j (length (j # a) + x) \<up>"
using assms(3) vs_def
by (simp add: vs assms(1) case_two_psi_longest_prefix z_def)
qed
moreover have "\<phi> k \<in> V_constotal_2"
proof (intro V_constotal_2I[of _ j a k])
show "\<phi> k = j # a @ [k] \<odot> 0\<^sup>\<infinity>"
using phi_k .
show "2 \<le> j"
- using `2 \<le> j` .
+ using \<open>2 \<le> j\<close> .
show "2 \<le> k"
- using `2 \<le> k` .
+ using \<open>2 \<le> k\<close> .
show "\<forall>i<length a. a ! i \<le> 1"
proof (rule allI, rule impI)
fix i assume i: "i < length a"
then have "Suc i < z"
using z_def vs_def length_prefix \<open>0 < z\<close> vs
by (metis One_nat_def Suc_mono Suc_pred length_Cons)
have "a ! i = vs ! (Suc i)"
using vs by simp
also have "... = the (\<psi> j (Suc i))"
using vs_def vs i length_Cons length_prefix prefix_nth
by (metis Suc_mono)
finally show "a ! i \<le> 1"
- using case_two_psi_longest_prefix `Suc i < z` z_def
+ using case_two_psi_longest_prefix \<open>Suc i < z\<close> z_def
by (metis assms(1) less_or_eq_imp_le not_le_imp_less not_one_less_zero
option.sel zero_less_Suc)
qed
qed (auto simp add: phi_j)
then have "\<phi> k \<in> V_constotal"
using V_constotal_def by auto
moreover have "\<not> learn_total \<phi> {\<phi> k} s"
proof -
have "\<phi> k \<in> \<R>"
by (simp add: phi_k almost0_in_R1)
moreover have "\<And>x. x < longest_prefix j \<Longrightarrow> \<phi> k x = \<psi> j x"
using phi_k vs_def z_def length_prefix phi_j prepend_associative prepend_at_less
by (metis One_nat_def Suc_pred \<open>0 < z\<close> \<open>vs = j # a\<close> append_Cons assms(3))
ultimately show ?thesis
using case_two_not_learn[OF assms(1)] by simp
qed
ultimately show "\<not> learn_total \<phi> V_constotal s"
using learn_total_closed_subseteq by auto
qed
text \<open>The strategy $S$ does not learn $V$ in either case.\<close>
lemma not_learn_total_V: "\<not> learn_total \<phi> V_constotal s"
proof -
obtain j where "j \<ge> 2" "\<phi> j = \<psi> j"
using kleene_fixed_point psi_in_P2 by auto
then show ?thesis
using case_one_not_learn_V learn_total_def case_two_not_learn_V
by (cases "case_two j") auto
qed
end
lemma V_not_in_TOTAL: "V_constotal \<notin> TOTAL"
proof (rule ccontr)
assume "\<not> V_constotal \<notin> TOTAL"
then have "V_constotal \<in> TOTAL" by simp
then have "V_constotal \<in> TOTAL_wrt \<phi>"
by (simp add: TOTAL_wrt_phi_eq_TOTAL)
then obtain s where "learn_total \<phi> V_constotal s"
using TOTAL_wrt_def by auto
then obtain s' where s': "s' \<in> \<R>" "learn_total \<phi> V_constotal s'"
using lemma_R_for_TOTAL_simple by blast
then interpret total_cons s'
by (simp add: total_cons_def)
have "\<not> learn_total \<phi> V_constotal s'"
by (simp add: not_learn_total_V)
with s'(2) show False by simp
qed
lemma TOTAL_neq_CONS: "TOTAL \<noteq> CONS"
using V_not_in_TOTAL V_in_CONS CONS_def by auto
text \<open>The main result of this section:\<close>
theorem TOTAL_subset_CONS: "TOTAL \<subset> CONS"
using TOTAL_subseteq_CONS TOTAL_neq_CONS by simp
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/Union.thy b/thys/Inductive_Inference/Union.thy
--- a/thys/Inductive_Inference/Union.thy
+++ b/thys/Inductive_Inference/Union.thy
@@ -1,283 +1,283 @@
section \<open>The union of classes\label{s:union}\<close>
theory Union
imports R1_BC TOTAL_CONS
begin
text \<open>None of the inference types introduced in this chapter are closed
under union of classes. For all inference types except FIN this follows from
@{thm[source] "U0_V0_not_in_BC"}.\<close>
lemma not_closed_under_union:
"\<forall>\<I>\<in>{CP, TOTAL, CONS, LIM, BC}. U\<^sub>0 \<in> \<I> \<and> V\<^sub>0 \<in> \<I> \<and> U\<^sub>0 \<union> V\<^sub>0 \<notin> \<I>"
using U0_in_CP U0_in_NUM V0_in_FIN
FIN_subseteq_CP
NUM_subseteq_TOTAL
CP_subseteq_TOTAL
TOTAL_subseteq_CONS
CONS_subseteq_Lim
Lim_subseteq_BC
U0_V0_not_in_BC
by blast
text \<open>In order to show the analogous result for FIN consider the
classes $\{0^\infty\}$ and $\{0^n10^\infty \mid n \in \mathbb{N}\}$. The
former can be learned finitely by a strategy that hypothesizes $0^\infty$ for
every input. The latter can be learned finitely by a strategy that waits for
the 1 and hypothesizes the only function in the class with a 1 at that
position. However, the union of both classes is not in FIN. This is because
any FIN strategy has to hypothesize $0^\infty$ on some prefix of the form
$0^n$. But the strategy then fails for the function $0^n10^\infty$.\<close>
lemma singleton_in_FIN: "f \<in> \<R> \<Longrightarrow> {f} \<in> FIN"
proof -
assume "f \<in> \<R>"
then obtain i where i: "\<phi> i = f"
using phi_universal by blast
define s :: partial1 where "s = (\<lambda>_. Some (Suc i))"
then have "s \<in> \<R>"
using const_in_Prim1[of "Suc i"] by simp
have "learn_fin \<phi> {f} s"
proof (intro learn_finI)
show "environment \<phi> {f} s"
- using `s \<in> \<R>` `f \<in> \<R>` by (simp add: phi_in_P2)
+ using \<open>s \<in> \<R>\<close> \<open>f \<in> \<R>\<close> by (simp add: phi_in_P2)
show "\<exists>i n\<^sub>0. \<phi> i = g \<and> (\<forall>n<n\<^sub>0. s (g \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (g \<triangleright> n) \<down>= Suc i)"
if "g \<in> {f}" for g
proof -
from that have "g = f" by simp
then have "\<phi> i = g"
using i by simp
moreover have "\<forall>n<0. s (g \<triangleright> n) \<down>= 0" by simp
moreover have "\<forall>n\<ge>0. s (g \<triangleright> n) \<down>= Suc i"
using s_def by simp
ultimately show ?thesis by auto
qed
qed
then show "{f} \<in> FIN" using FIN_def by auto
qed
definition U_single :: "partial1 set" where
"U_single \<equiv> {(\<lambda>x. if x = n then Some 1 else Some 0)| n. n \<in> UNIV}"
lemma U_single_in_FIN: "U_single \<in> FIN"
proof -
define psi :: partial2 where "psi \<equiv> \<lambda>n x. if x = n then Some 1 else Some 0"
have "psi \<in> \<R>\<^sup>2"
using psi_def by (intro R2I[of "Cn 2 r_not [r_eq]"]) auto
define s :: partial1 where
"s \<equiv> \<lambda>b. if findr b \<down>= e_length b then Some 0 else Some (Suc (the (findr b)))"
have "s \<in> \<R>"
proof (rule R1I)
let ?r = "Cn 1 r_ifeq [r_findr, r_length, Z, Cn 1 S [r_findr]]"
show "recfn 1 ?r" by simp
show "total ?r" by auto
show "eval ?r [b] = s b" for b
proof -
let ?b = "the (findr b)"
have "eval ?r [b] = (if ?b = e_length b then Some 0 else Some (Suc (?b)))"
using findr_total by simp
then show "eval ?r [b] = s b"
by (metis findr_total option.collapse option.inject s_def)
qed
qed
have "U_single \<subseteq> \<R>"
proof
fix f
assume "f \<in> U_single"
then obtain n where "f = (\<lambda>x. if x = n then Some 1 else Some 0)"
using U_single_def by auto
then have "f = psi n"
using psi_def by simp
then show "f \<in> \<R>"
- using `psi \<in> \<R>\<^sup>2` by simp
+ using \<open>psi \<in> \<R>\<^sup>2\<close> by simp
qed
have "learn_fin psi U_single s"
proof (rule learn_finI)
show "environment psi U_single s"
- using `psi \<in> \<R>\<^sup>2` `s \<in> \<R>` `U_single \<subseteq> \<R>` by simp
+ using \<open>psi \<in> \<R>\<^sup>2\<close> \<open>s \<in> \<R>\<close> \<open>U_single \<subseteq> \<R>\<close> by simp
show "\<exists>i n\<^sub>0. psi i = f \<and> (\<forall>n<n\<^sub>0. s (f \<triangleright> n) \<down>= 0) \<and> (\<forall>n\<ge>n\<^sub>0. s (f \<triangleright> n) \<down>= Suc i)"
if "f \<in> U_single" for f
proof -
from that obtain i where i: "f = (\<lambda>x. if x = i then Some 1 else Some 0)"
using U_single_def by auto
then have "psi i = f"
using psi_def by simp
moreover have "\<forall>n<i. s (f \<triangleright> n) \<down>= 0"
using i s_def findr_def by simp
moreover have "\<forall>n\<ge>i. s (f \<triangleright> n) \<down>= Suc i"
proof (rule allI, rule impI)
fix n
assume "n \<ge> i"
let ?e = "init f n"
have "\<exists>i<e_length ?e. e_nth ?e i \<noteq> 0"
- using `n \<ge> i` i by simp
+ using \<open>n \<ge> i\<close> i by simp
then have less: "the (findr ?e) < e_length ?e"
and nth_e: "e_nth ?e (the (findr ?e)) \<noteq> 0"
using findr_ex by blast+
then have "s ?e \<down>= Suc (the (findr ?e))"
using s_def by auto
moreover have "the (findr ?e) = i"
using nth_e less i by (metis length_init nth_init option.sel)
ultimately show "s ?e \<down>= Suc i" by simp
qed
ultimately show ?thesis by auto
qed
qed
then show "U_single \<in> FIN" using FIN_def by blast
qed
lemma zero_U_single_not_in_FIN: "{0\<^sup>\<infinity>} \<union> U_single \<notin> FIN"
proof
assume "{0\<^sup>\<infinity>} \<union> U_single \<in> FIN"
then obtain psi s where learn: "learn_fin psi ({0\<^sup>\<infinity>} \<union> U_single) s"
using FIN_def by blast
then have "learn_fin psi {0\<^sup>\<infinity>} s"
using learn_fin_closed_subseteq by auto
then obtain i n\<^sub>0 where i:
"psi i = 0\<^sup>\<infinity>"
"\<forall>n<n\<^sub>0. s (0\<^sup>\<infinity> \<triangleright> n) \<down>= 0"
"\<forall>n\<ge>n\<^sub>0. s (0\<^sup>\<infinity> \<triangleright> n) \<down>= Suc i"
using learn_finE(2) by blast
let ?f = "\<lambda>x. if x = Suc n\<^sub>0 then Some 1 else Some 0"
have "?f \<noteq> 0\<^sup>\<infinity>" by (metis option.inject zero_neq_one)
have "?f \<in> U_single"
using U_single_def by auto
then have "learn_fin psi {?f} s"
using learn learn_fin_closed_subseteq by simp
then obtain j m\<^sub>0 where j:
"psi j = ?f"
"\<forall>n<m\<^sub>0. s (?f \<triangleright> n) \<down>= 0"
"\<forall>n\<ge>m\<^sub>0. s (?f \<triangleright> n) \<down>= Suc j"
using learn_finE(2) by blast
consider
(less) "m\<^sub>0 < n\<^sub>0" | (eq) "m\<^sub>0 = n\<^sub>0" | (gr) "m\<^sub>0 > n\<^sub>0"
by linarith
then show False
proof (cases)
case less
then have "s (0\<^sup>\<infinity>\<triangleright> m\<^sub>0) \<down>= 0"
using i by simp
moreover have "0\<^sup>\<infinity> \<triangleright> m\<^sub>0 = ?f \<triangleright> m\<^sub>0"
using less init_eqI[of m\<^sub>0 ?f "0\<^sup>\<infinity>"] by simp
ultimately have "s (?f \<triangleright> m\<^sub>0) \<down>= 0" by simp
then show False using j by simp
next
case eq
then have "0\<^sup>\<infinity> \<triangleright> m\<^sub>0 = ?f \<triangleright> m\<^sub>0"
using init_eqI[of m\<^sub>0 ?f "0\<^sup>\<infinity>"] by simp
then have "s (0\<^sup>\<infinity> \<triangleright> m\<^sub>0) = s (?f \<triangleright> m\<^sub>0)" by simp
then have "i = j"
using i j eq by simp
then have "psi i = psi j" by simp
- then show False using `?f \<noteq> 0\<^sup>\<infinity>` i j by simp
+ then show False using \<open>?f \<noteq> 0\<^sup>\<infinity>\<close> i j by simp
next
case gr
have "0\<^sup>\<infinity> \<triangleright> n\<^sub>0 = ?f \<triangleright> n\<^sub>0"
using init_eqI[of n\<^sub>0 ?f "0\<^sup>\<infinity>"] by simp
moreover have "s (0\<^sup>\<infinity> \<triangleright> n\<^sub>0) \<down>= Suc i"
using i by simp
moreover have "s (?f \<triangleright> n\<^sub>0) \<down>= 0"
using j gr by simp
ultimately show False by simp
qed
qed
lemma FIN_not_closed_under_union: "\<exists>U V. U \<in> FIN \<and> V \<in> FIN \<and> U \<union> V \<notin> FIN"
proof -
have "{0\<^sup>\<infinity>} \<in> FIN"
using singleton_in_FIN const_in_Prim1 by simp
moreover have "U_single \<in> FIN"
using U_single_in_FIN by simp
ultimately show ?thesis
using zero_U_single_not_in_FIN by blast
qed
text \<open>In contrast to the inference types, NUM is closed under the union
of classes. The total numberings that exist for each NUM class can be
interleaved to produce a total numbering encompassing the union of the
classes. To define the interleaving, modulo and division by two will be
helpful.\<close>
definition "r_div2 \<equiv>
r_shrink
(Pr 1 Z
(Cn 3 r_ifle
[Cn 3 r_mul [r_constn 2 2, Cn 3 S [Id 3 0]], Id 3 2, Cn 3 S [Id 3 1], Id 3 1]))"
lemma r_div2_prim [simp]: "prim_recfn 1 r_div2"
unfolding r_div2_def by simp
lemma r_div2 [simp]: "eval r_div2 [n] \<down>= n div 2"
proof -
let ?p = "Pr 1 Z
(Cn 3 r_ifle
[Cn 3 r_mul [r_constn 2 2, Cn 3 S [Id 3 0]], Id 3 2, Cn 3 S [Id 3 1], Id 3 1])"
have "eval ?p [i, n] \<down>= min (n div 2) i" for i
by (induction i) auto
then have "eval ?p [n, n] \<down>= n div 2" by simp
then show ?thesis unfolding r_div2_def by simp
qed
definition "r_mod2 \<equiv> Cn 1 r_sub [Id 1 0, Cn 1 r_mul [r_const 2, r_div2]]"
lemma r_mod2_prim [simp]: "prim_recfn 1 r_mod2"
unfolding r_mod2_def by simp
lemma r_mod2 [simp]: "eval r_mod2 [n] \<down>= n mod 2"
unfolding r_mod2_def using Rings.semiring_modulo_class.minus_mult_div_eq_mod
by auto
lemma NUM_closed_under_union:
assumes "U \<in> NUM" and "V \<in> NUM"
shows "U \<union> V \<in> NUM"
proof -
from assms obtain psi_u psi_v where
psi_u: "psi_u \<in> \<R>\<^sup>2" "\<And>f. f \<in> U \<Longrightarrow> \<exists>i. psi_u i = f" and
psi_v: "psi_v \<in> \<R>\<^sup>2" "\<And>f. f \<in> V \<Longrightarrow> \<exists>i. psi_v i = f"
by fastforce
define psi where "psi \<equiv> \<lambda>i. if i mod 2 = 0 then psi_u (i div 2) else psi_v (i div 2)"
from psi_u(1) obtain u where u: "recfn 2 u" "total u" "\<And>x y. eval u [x, y] = psi_u x y"
by auto
from psi_v(1) obtain v where v: "recfn 2 v" "total v" "\<And>x y. eval v [x, y] = psi_v x y"
by auto
let ?r_psi = "Cn 2 r_ifz
[Cn 2 r_mod2 [Id 2 0],
Cn 2 u [Cn 2 r_div2 [Id 2 0], Id 2 1],
Cn 2 v [Cn 2 r_div2 [Id 2 0], Id 2 1]]"
show ?thesis
proof (rule NUM_I[of psi])
show "psi \<in> \<R>\<^sup>2"
proof (rule R2I)
show "recfn 2 ?r_psi"
using u(1) v(1) by simp
show "eval ?r_psi [x, y] = psi x y" for x y
using u v psi_def prim_recfn_total R2_imp_total2[OF psi_u(1)]
R2_imp_total2[OF psi_v(1)]
by simp
moreover have "psi x y \<down>" for x y
using psi_def psi_u(1) psi_v(1) by simp
ultimately show "total ?r_psi"
- using `recfn 2 ?r_psi` totalI2 by simp
+ using \<open>recfn 2 ?r_psi\<close> totalI2 by simp
qed
show "\<exists>i. psi i = f" if "f \<in> U \<union> V" for f
proof (cases "f \<in> U")
case True
then obtain j where "psi_u j = f"
using psi_u(2) by auto
then have "psi (2 * j) = f"
using psi_def by simp
then show ?thesis by auto
next
case False
then have "f \<in> V"
using that by simp
then obtain j where "psi_v j = f"
using psi_v(2) by auto
then have "psi (Suc (2 * j)) = f"
using psi_def by simp
then show ?thesis by auto
qed
qed
qed
end
\ No newline at end of file
diff --git a/thys/Inductive_Inference/Universal.thy b/thys/Inductive_Inference/Universal.thy
--- a/thys/Inductive_Inference/Universal.thy
+++ b/thys/Inductive_Inference/Universal.thy
@@ -1,2538 +1,2538 @@
section \<open>A universal partial recursive function\<close>
theory Universal
imports Partial_Recursive
begin
text \<open>The main product of this section is a universal partial recursive
function, which given a code $i$ of an $n$-ary partial recursive function $f$
and an encoded list @{term xs} of $n$ arguments, computes @{term "eval f
xs"}. From this we can derive fixed-arity universal functions satisfying the
usual results such as the $s$-$m$-$n$ theorem. To represent the code $i$, we
need a way to encode @{typ recf}s as natural numbers (Section~\ref{s:recf_enc}). To
construct the universal function, we devise a ternary function taking $i$,
$xs$, and a step bound $t$ and simulating the execution of $f$ on input $xs$ for
$t$ steps. This function is useful in its own right, enabling techniques like
dovetailing or ``concurrent'' evaluation of partial recursive functions.
The notion of a ``step'' is not part of the definition of (the evaluation of)
partial recursive functions, but one can simulate the evaluation on an
abstract machine (Section~\ref{s:step}). This machine's configurations can be
encoded as natural numbers, and this leads us to a step function @{typ "nat
\<Rightarrow> nat"} on encoded configurations (Section~\ref{s:step_enc}).
This function in turn can be computed by a primitive recursive function, from
which we develop the aforementioned ternary function of $i$, @{term xs}, and
$t$ (Section~\ref{s:step_recf}). From this we can finally derive
a universal function (Section~\ref{s:the_universal}).\<close>
subsection \<open>A step function\label{s:step}\<close>
text \<open>We simulate the stepwise execution of a partial recursive
function in a fairly straightforward way reminiscent of the execution of
function calls in an imperative programming language. A configuration of the
abstract machine is a pair consisting of:
\begin{enumerate}
\item A stack of frames. A frame represents the execution of a function and is
a triple @{term "(f, xs, locals)"} of
\begin{enumerate}
\item a @{typ recf} @{term f} being executed,
\item a @{typ "nat list"} of arguments of @{term f},
\item a @{typ "nat list"} of local variables, which holds intermediate
values when @{term f} is of the form @{term Cn}, @{term Pr}, or @{term Mn}.
\end{enumerate}
\item A register of type @{typ "nat option"} representing the return value of
the last function call: @{term None} signals that in the previous step the
stack was not popped and hence no value was returned, whereas @{term "Some
v"} means that in the previous step a function returned @{term v}.
\end{enumerate}
For computing @{term h} on input @{term xs}, the initial configuration is
@{term "([(h, xs, [])], None)"}. When the computation for a frame ends, it is
popped off the stack, and its return value is put in the register. The entire
computation ends when the stack is empty. In such a final configuration the
register contains the value of @{term h} at @{term xs}. If no final
configuration is ever reached, @{term h} diverges at @{term xs}.
The execution of one step depends on the topmost (that is, active) frame. In
the step when a frame @{term "(h, xs, locals)"} is pushed onto the stack, the
local variables are @{term "locals = []"}. The following happens until the
frame is popped off the stack again (if it ever is):
\begin{itemize}
\item For the base functions @{term "h = Z"}, @{term "h = S"},
@{term[names_short] "h = Id m n"}, the frame is popped off the stack right away,
and the return value is placed in the register.
\item For @{term "h = Cn n f gs"}, for each function $g$ in @{term gs}:
\begin{enumerate}
\item A new frame of the form @{term "(g, xs, [])"} is pushed onto the stack.
\item When (and if) this frame
is eventually popped, the value in the register is @{term "eval g xs"}. This value
is appended to the list @{term locals} of local variables.
\end{enumerate}
When all $g$ in $gs$ have been evaluated in this manner, $f$ is evaluated on the local variables
by pushing @{term "(f, locals, [])"}. The resulting register value is kept
and the active frame for $h$ is popped off the stack.
\item For @{text "h = Pr n f g"}, let @{term "xs = y # ys"}. First @{term "(f,
ys, [])"} is pushed and the return value stored in the @{term
locals}. Then @{term "(g, x # v # ys, [])"} is pushed,
where $x$ is the length of @{term locals} and $v$ the most recently
appended value. The return value is appended to @{term locals}. This is
repeated until the length of @{term locals} reaches @{term y}. Then the most
recently appended local is placed in the register, and the stack is popped.
\item For @{text "h = Mn n f"}, frames @{term "(f, x # xs, [])"} are pushed
for $x = 0, 1, 2, \ldots$ until one of them returns $0$. Then this
$x$ is placed in the register and the stack is popped. Until then $x$ is
stored in @{term locals}. If none of these evaluations return $0$, the
stack never shrinks, and thus the machine never reaches a final state.
\end{itemize}\<close>
type_synonym frame = "recf \<times> nat list \<times> nat list"
type_synonym configuration = "frame list \<times> nat option"
subsubsection \<open>Definition of the step function\<close>
fun step :: "configuration \<Rightarrow> configuration" where
"step ([], rv) = ([], rv)"
| "step (((Z, _, _) # fs), rv) = (fs, Some 0)"
| "step (((S, xs, _) # fs), rv) = (fs, Some (Suc (hd xs)))"
| "step (((Id m n, xs, _) # fs), rv) = (fs, Some (xs ! n))"
| "step (((Cn n f gs, xs, ls) # fs), rv) =
(if length ls = length gs
then if rv = None
then ((f, ls, []) # (Cn n f gs, xs, ls) # fs, None)
else (fs, rv)
else if rv = None
then if length ls < length gs
then ((gs ! (length ls), xs, []) # (Cn n f gs, xs, ls) # fs, None)
else (fs, rv) \<comment>\<open>cannot occur, so don't-care term\<close>
else ((Cn n f gs, xs, ls @ [the rv]) # fs, None))"
| "step (((Pr n f g, xs, ls) # fs), rv) =
(if ls = []
then if rv = None
then ((f, tl xs, []) # (Pr n f g, xs, ls) # fs, None)
else ((Pr n f g, xs, [the rv]) # fs, None)
else if length ls = Suc (hd xs)
then (fs, Some (hd ls))
else if rv = None
then ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None)
else ((Pr n f g, xs, (the rv) # ls) # fs, None))"
| "step (((Mn n f, xs, ls) # fs), rv) =
(if ls = []
then ((f, 0 # xs, []) # (Mn n f, xs, [0]) # fs, None)
else if rv = Some 0
then (fs, Some (hd ls))
else ((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None))"
definition reachable :: "configuration \<Rightarrow> configuration \<Rightarrow> bool" where
"reachable x y \<equiv> \<exists>t. iterate t step x = y"
lemma step_reachable [intro]:
assumes "step x = y"
shows "reachable x y"
unfolding reachable_def using assms by (metis iterate.simps(1,2) comp_id)
lemma reachable_transitive [trans]:
assumes "reachable x y" and "reachable y z"
shows "reachable x z"
using assms iterate_additive[where ?f=step] reachable_def by metis
lemma reachable_refl: "reachable x x"
unfolding reachable_def by (metis iterate.simps(1) eq_id_iff)
text \<open>From a final configuration, that is, when the stack is empty,
only final configurations are reachable.\<close>
lemma step_empty_stack:
assumes "fst x = []"
shows "fst (step x) = []"
using assms by (metis prod.collapse step.simps(1))
lemma reachable_empty_stack:
assumes "fst x = []" and "reachable x y"
shows "fst y = []"
proof -
have "fst (iterate t step x) = []" for t
using assms step_empty_stack by (induction t) simp_all
then show ?thesis
using reachable_def assms(2) by auto
qed
abbreviation nonterminating :: "configuration \<Rightarrow> bool" where
"nonterminating x \<equiv> \<forall>t. fst (iterate t step x) \<noteq> []"
lemma reachable_nonterminating:
assumes "reachable x y" and "nonterminating y"
shows "nonterminating x"
proof -
from assms(1) obtain t\<^sub>1 where t1: "iterate t\<^sub>1 step x = y"
using reachable_def by auto
have "fst (iterate t step x) \<noteq> []" for t
proof (cases "t \<le> t\<^sub>1")
case True
then show ?thesis
using t1 assms(2) reachable_def reachable_empty_stack iterate_additive'
by (metis le_Suc_ex)
next
case False
then have "iterate t step x = iterate (t\<^sub>1 + (t - t\<^sub>1)) step x"
by simp
then have "iterate t step x = iterate (t - t\<^sub>1) step (iterate t\<^sub>1 step x)"
by (simp add: iterate_additive')
then have "iterate t step x = iterate (t - t\<^sub>1) step y"
using t1 by simp
then show "fst (iterate t step x) \<noteq> []"
using assms(2) by simp
qed
then show ?thesis ..
qed
text \<open>The function @{term step} is underdefined, for example, when the
top frame contains a non-well-formed @{typ recf} or too few arguments. All is
well, though, if every frame contains a well-formed @{typ recf} whose arity
matches the number of arguments. Such stacks will be called
\emph{valid}.\<close>
definition valid :: "frame list \<Rightarrow> bool" where
"valid stack \<equiv> \<forall>s\<in>set stack. recfn (length (fst (snd s))) (fst s)"
lemma valid_frame: "valid (s # ss) \<Longrightarrow> valid ss \<and> recfn (length (fst (snd s))) (fst s)"
using valid_def by simp
lemma valid_ConsE: "valid ((f, xs, locs) # rest) \<Longrightarrow> valid rest \<and> recfn (length xs) f"
using valid_def by simp
lemma valid_ConsI: "valid rest \<Longrightarrow> recfn (length xs) f \<Longrightarrow> valid ((f, xs, locs) # rest)"
using valid_def by simp
text \<open>Stacks in initial configurations are valid, and performing a step
maintains the validity of the stack.\<close>
lemma step_valid: "valid stack \<Longrightarrow> valid (fst (step (stack, rv)))"
proof (cases stack)
case Nil
then show ?thesis using valid_def by simp
next
case (Cons s ss)
assume valid: "valid stack"
then have *: "valid ss \<and> recfn (length (fst (snd s))) (fst s)"
using valid_frame Cons by simp
show ?thesis
proof (cases "fst s")
case Z
then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(2))
next
case S
then show ?thesis using Cons valid * by (metis fst_conv prod.collapse step.simps(3))
next
case Id
then show ?thesis using Cons valid * by (metis fstI prod.collapse step.simps(4))
next
case (Cn n f gs)
then obtain xs ls where "s = (Cn n f gs, xs, ls)"
using Cons by (metis prod.collapse)
moreover consider
"length ls = length gs \<and> rv \<up>"
| "length ls = length gs \<and> rv \<down>"
| "length ls < length gs \<and> rv \<up>"
| "length ls \<noteq> length gs \<and> rv \<down>"
| "length ls > length gs \<and> rv \<up>"
by linarith
ultimately show ?thesis using valid Cons valid_def by (cases) auto
next
case (Pr n f g)
then obtain xs ls where s: "s = (Pr n f g, xs, ls)"
using Cons by (metis prod.collapse)
consider
"length ls = 0 \<and> rv \<up>"
| "length ls = 0 \<and> rv \<down>"
| "length ls \<noteq> 0 \<and> length ls = Suc (hd xs)"
| "length ls \<noteq> 0 \<and> length ls \<noteq> Suc (hd xs) \<and> rv \<up>"
| "length ls \<noteq> 0 \<and> length ls \<noteq> Suc (hd xs) \<and> rv \<down>"
by linarith
then show ?thesis using Cons * valid_def s by (cases) auto
next
case (Mn n f)
then obtain xs ls where s: "s = (Mn n f, xs, ls)"
using Cons by (metis prod.collapse)
consider
"length ls = 0"
| "length ls \<noteq> 0 \<and> rv \<up>"
| "length ls \<noteq> 0 \<and> rv \<down>"
by linarith
then show ?thesis using Cons * valid_def s by (cases) auto
qed
qed
corollary iterate_step_valid:
assumes "valid stack"
shows "valid (fst (iterate t step (stack, rv)))"
using assms
proof (induction t)
case 0
then show ?case by simp
next
case (Suc t)
moreover have "iterate (Suc t) step (stack, rv) = step (iterate t step (stack, rv))"
by simp
ultimately show ?case using step_valid valid_def by (metis prod.collapse)
qed
subsubsection \<open>Correctness of the step function\<close>
text \<open>The function @{term step} works correctly for a @{typ recf} $f$
on arguments @{term xs} in some configuration if (1) in case $f$ converges, @{term
step} reaches a configuration with the topmost frame popped and @{term "eval
f xs"} in the register, and (2) in case $f$ diverges, @{term step} does not
reach a final configuration.\<close>
fun correct :: "configuration \<Rightarrow> bool" where
"correct ([], r) = True"
| "correct ((f, xs, ls) # rest, r) =
(if eval f xs \<down> then reachable ((f, xs, ls) # rest, r) (rest, eval f xs)
else nonterminating ((f, xs, ls) # rest, None))"
lemma correct_convergI:
assumes "eval f xs \<down>" and "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)"
shows "correct ((f, xs, ls) # rest, None)"
using assms by auto
lemma correct_convergE:
assumes "correct ((f, xs, ls) # rest, None)" and "eval f xs \<down>"
shows "reachable ((f, xs, ls) # rest, None) (rest, eval f xs)"
using assms by simp
text \<open>The correctness proof for @{term step} is by structural induction
on the @{typ recf} in the top frame. The base cases @{term Z}, @{term S},
and @{term[names_short] Id} are simple. For @{text "X = Cn, Pr, Mn"}, the
lemmas named @{text reachable_X} show which configurations are reachable for
@{typ recf}s of shape @{text X}. Building on those, the lemmas named @{text
step_X_correct} show @{term step}'s correctness for @{text X}.\<close>
lemma reachable_Cn:
assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
and "\<And>g xs rest.
g \<in> set gs \<Longrightarrow> valid ((g, xs, []) # rest) \<Longrightarrow> correct ((g, xs, []) # rest, None)"
and "\<forall>i<k. eval (gs ! i) xs \<down>"
and "k \<le> length gs"
shows "reachable
(?stack, None)
((Cn n f gs, xs, take k (map (\<lambda>g. the (eval g xs)) gs)) # rest, None)"
using assms(4,5)
proof (induction k)
case 0
then show ?case using reachable_refl by simp
next
case (Suc k)
let ?ys = "map (\<lambda>g. the (eval g xs)) gs"
from Suc have "k < length gs" by simp
have valid: "recfn (length xs) (Cn n f gs)" "valid rest"
using assms(1) valid_ConsE[of "(Cn n f gs)"] by simp_all
from Suc have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)"
(is "_ (?stack1, None)")
by simp
also have "reachable ... ((gs ! k, xs, []) # ?stack1, None)"
- using step_reachable `k < length gs`
+ using step_reachable \<open>k < length gs\<close>
by (auto simp: min_absorb2)
also have "reachable ... (?stack1, eval (gs ! k) xs)"
(is "_ (_, ?rv)")
using Suc.prems(1) \<open>k < length gs\<close> assms(3) valid valid_ConsI by auto
also have "reachable ... ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)"
(is "_ (?stack2, None)")
proof -
have "step (?stack1, ?rv) = ((Cn n f gs, xs, (take k ?ys) @ [the ?rv]) # rest, None)"
using Suc by auto
also have "... = ((Cn n f gs, xs, (take (Suc k) ?ys)) # rest, None)"
by (simp add: \<open>k < length gs\<close> take_Suc_conv_app_nth)
finally show ?thesis
using step_reachable by auto
qed
finally show "reachable (?stack, None) (?stack2, None)" .
qed
lemma step_Cn_correct:
assumes "valid (((Cn n f gs), xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
and "\<And>g xs rest.
g \<in> set gs \<Longrightarrow> valid ((g, xs, []) # rest) \<Longrightarrow> correct ((g, xs, []) # rest, None)"
shows "correct (?stack, None)"
proof -
have valid: "recfn (length xs) (Cn n f gs)" "valid rest"
using valid_ConsE[OF assms(1)] by auto
let ?ys = "map (\<lambda>g. the (eval g xs)) gs"
consider
(diverg_f) "\<forall>g\<in>set gs. eval g xs \<down>" and "eval f ?ys \<up>"
| (diverg_gs) "\<exists>g\<in>set gs. eval g xs \<up>"
| (converg) "eval (Cn n f gs) xs \<down>"
using valid_ConsE[OF assms(1)] by fastforce
then show ?thesis
proof (cases)
case diverg_f
then have "\<forall>i<length gs. eval (gs ! i) xs \<down>" by simp
then have "reachable (?stack, None) ((Cn n f gs, xs, ?ys) # rest, None)"
(is "_ (?stack1, None)")
using reachable_Cn[OF assms, where ?k="length gs"] by simp
also have "reachable ... ((f, ?ys, []) # ?stack1, None)" (is "_ (?stack2, None)")
by (simp add: step_reachable)
finally have "reachable (?stack, None) (?stack2, None)" .
moreover have "nonterminating (?stack2, None)"
using diverg_f(2) assms(2)[of ?ys ?stack1] valid_ConsE[OF assms(1)] valid_ConsI
by auto
ultimately have "nonterminating (?stack, None)"
using reachable_nonterminating by simp
moreover have "eval (Cn n f gs) xs \<up>"
using diverg_f(2) assms(1) eval_Cn valid_ConsE by presburger
ultimately show ?thesis by simp
next
case diverg_gs
then have ex_i: "\<exists>i<length gs. eval (gs ! i) xs \<up>"
using in_set_conv_nth[of _ gs] by auto
define k where "k = (LEAST i. i < length gs \<and> eval (gs ! i) xs \<up>)" (is "_ = Least ?P")
then have gs_k: "eval (gs ! k) xs \<up>"
using LeastI_ex[OF ex_i] by simp
have "\<forall>i<k. eval (gs ! i) xs \<down>"
using k_def not_less_Least[of _ ?P] LeastI_ex[OF ex_i] by simp
moreover from this have "k < length gs"
using ex_i less_le_trans not_le by blast
ultimately have "reachable (?stack, None) ((Cn n f gs, xs, take k ?ys) # rest, None)"
using reachable_Cn[OF assms] by simp
also have "reachable ...
((gs ! (length (take k ?ys)), xs, []) # (Cn n f gs, xs, take k ?ys) # rest, None)"
(is "_ (?stack1, None)")
proof -
have "length (take k ?ys) < length gs"
by (simp add: \<open>k < length gs\<close> less_imp_le_nat min_less_iff_disj)
then show ?thesis using step_reachable by auto
qed
finally have "reachable (?stack, None) (?stack1, None)" .
moreover have "nonterminating (?stack1, None)"
proof -
have "recfn (length xs) (gs ! k)"
using \<open>k < length gs\<close> valid(1) by simp
then have "correct (?stack1, None)"
using \<open>k < length gs\<close> nth_mem valid valid_ConsI
assms(3)[of "gs ! (length (take k ?ys))" xs]
by auto
moreover have "length (take k ?ys) = k"
by (simp add: \<open>k < length gs\<close> less_imp_le_nat min_absorb2)
ultimately show ?thesis using gs_k by simp
qed
ultimately have "nonterminating (?stack, None)"
using reachable_nonterminating by simp
moreover have "eval (Cn n f gs) xs \<up>"
using diverg_gs valid by fastforce
ultimately show ?thesis by simp
next
case converg
then have f: "eval f ?ys \<down>" and g: "\<And>g. g \<in> set gs \<Longrightarrow> eval g xs \<down>"
using valid(1) by (metis eval_Cn)+
then have "\<forall>i<length gs. eval (gs ! i) xs \<down>"
by simp
then have "reachable (?stack, None) ((Cn n f gs, xs, take (length gs) ?ys) # rest, None)"
using reachable_Cn assms by blast
also have "reachable ... ((Cn n f gs, xs, ?ys) # rest, None)" (is "_ (?stack1, None)")
by (simp add: reachable_refl)
also have "reachable ... ((f, ?ys, []) # ?stack1, None)"
using step_reachable by auto
also have "reachable ... (?stack1, eval f ?ys)"
using assms(2)[of "?ys"] correct_convergE valid f valid_ConsI by auto
also have "reachable (?stack1, eval f ?ys) (rest, eval f ?ys)"
using f by auto
finally have "reachable (?stack, None) (rest, eval f ?ys)" .
moreover have "eval (Cn n f gs) xs = eval f ?ys"
using g valid(1) by auto
ultimately show ?thesis
using converg correct_convergI by auto
qed
qed
text \<open>During the execution of a frame with a partial recursive function
of shape @{term "Pr n f g"} and arguments @{term "x # xs"}, the list of local
variables collects all the function values up to @{term x} in reversed
order. We call such a list a @{term trace} for short.\<close>
definition trace :: "nat \<Rightarrow> recf \<Rightarrow> recf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> nat list" where
"trace n f g xs x \<equiv> map (\<lambda>y. the (eval (Pr n f g) (y # xs))) (rev [0..<Suc x])"
lemma trace_length: "length (trace n f g xs x) = Suc x"
using trace_def by simp
lemma trace_hd: "hd (trace n f g xs x) = the (eval (Pr n f g) (x # xs))"
using trace_def by simp
lemma trace_Suc:
"trace n f g xs (Suc x) = (the (eval (Pr n f g) (Suc x # xs))) # (trace n f g xs x)"
using trace_def by simp
lemma reachable_Pr:
assumes "valid (((Pr n f g), x # xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
and "\<And>xs rest. valid ((g, xs, []) # rest) \<Longrightarrow> correct ((g, xs, []) # rest, None)"
and "y \<le> x"
and "eval (Pr n f g) (y # xs) \<down>"
shows "reachable (?stack, None) ((Pr n f g, x # xs, trace n f g xs y) # rest, None)"
using assms(4,5)
proof (induction y)
case 0
have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest"
using valid_ConsE[OF assms(1)] by simp_all
then have f: "eval f xs \<down>" using 0 by simp
let ?as = "x # xs"
have "reachable (?stack, None) ((f, xs, []) # ((Pr n f g), ?as, []) # rest, None)"
using step_reachable by auto
also have "reachable ... (?stack, eval f xs)"
using assms(2)[of xs "((Pr n f g), ?as, []) # rest"]
correct_convergE[OF _ f] f valid valid_ConsI
by simp
also have "reachable ... ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)"
using step_reachable valid(1) f by auto
finally have "reachable (?stack, None) ((Pr n f g, ?as, [the (eval f xs)]) # rest, None)" .
then show ?case using trace_def valid(1) by simp
next
case (Suc y)
have valid: "recfn (length (x # xs)) (Pr n f g)" "valid rest"
using valid_ConsE[OF assms(1)] by simp_all
let ?ls = "trace n f g xs y"
have lenls: "length ?ls = Suc y"
using trace_length by auto
moreover have hdls: "hd ?ls = the (eval (Pr n f g) (y # xs))"
using Suc trace_hd by auto
ultimately have g:
"eval g (y # hd ?ls # xs) \<down>"
"eval (Pr n f g) (Suc y # xs) = eval g (y # hd ?ls # xs)"
using eval_Pr_Suc_converg hdls valid(1) Suc by simp_all
then have "reachable (?stack, None) ((Pr n f g, x # xs, ?ls) # rest, None)"
(is "_ (?stack1, None)")
using Suc valid(1) by fastforce
also have "reachable ... ((g, y # hd ?ls # xs, []) # (Pr n f g, x # xs, ?ls) # rest, None)"
using Suc.prems lenls by fastforce
also have "reachable ... (?stack1, eval g (y # hd ?ls # xs))"
(is "_ (_, ?rv)")
using assms(3) g(1) valid valid_ConsI by auto
also have "reachable ... ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)"
using Suc.prems(1) g(1) lenls by auto
finally have "reachable (?stack, None) ((Pr n f g, x # xs, (the ?rv) # ?ls) # rest, None)" .
moreover have "trace n f g xs (Suc y) = (the ?rv) # ?ls"
using g(2) trace_Suc by simp
ultimately show ?case by simp
qed
lemma step_Pr_correct:
assumes "valid (((Pr n f g), xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
and "\<And>xs rest. valid ((g, xs, []) # rest) \<Longrightarrow> correct ((g, xs, []) # rest, None)"
shows "correct (?stack, None)"
proof -
have valid: "valid rest" "recfn (length xs) (Pr n f g)"
using valid_ConsE[OF assms(1)] by simp_all
then have "length xs > 0"
by auto
then obtain y ys where y_ys: "xs = y # ys"
using list.exhaust_sel by auto
let ?t = "trace n f g ys"
consider
(converg) "eval (Pr n f g) xs \<down>"
| (diverg_f) "eval (Pr n f g) xs \<up>" and "eval f ys \<up>"
| (diverg) "eval (Pr n f g) xs \<up>" and "eval f ys \<down>"
by auto
then show ?thesis
proof (cases)
case converg
then have "\<And>z. z \<le> y \<Longrightarrow> reachable (?stack, None) (((Pr n f g), xs, ?t z) # rest, None)"
using assms valid by (simp add: eval_Pr_converg_le reachable_Pr y_ys)
then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)"
by simp
moreover have "reachable (((Pr n f g), xs, ?t y) # rest, None) (rest, Some (hd (?t y)))"
using trace_length step_reachable y_ys by fastforce
ultimately have "reachable (?stack, None) (rest, Some (hd (?t y)))"
using reachable_transitive by blast
then show ?thesis
using assms(1) trace_hd converg y_ys by simp
next
case diverg_f
have *: "step (?stack, None) = ((f, ys, []) # ((Pr n f g), xs, []) # tl ?stack, None)"
(is "_ = (?stack1, None)")
using assms(1,2) y_ys by simp
then have "reachable (?stack, None) (?stack1, None)"
using step_reachable by force
moreover have "nonterminating (?stack1, None)"
using assms diverg_f valid valid_ConsI * by auto
ultimately have "nonterminating (?stack, None)"
using reachable_nonterminating by blast
then show ?thesis using diverg_f(1) assms(1) by simp
next
case diverg
let ?h = "\<lambda>z. the (eval (Pr n f g) (z # ys))"
let ?Q = "\<lambda>z. z < y \<and> eval (Pr n f g) (z # ys) \<down>"
have "?Q 0"
using assms diverg neq0_conv y_ys valid by fastforce
define zmax where "zmax = Greatest ?Q"
then have "?Q zmax"
- using `?Q 0` GreatestI_nat[of ?Q 0 y] by simp
+ using \<open>?Q 0\<close> GreatestI_nat[of ?Q 0 y] by simp
have le_zmax: "\<And>z. ?Q z \<Longrightarrow> z \<le> zmax"
using Greatest_le_nat[of ?Q _ y] zmax_def by simp
have len: "length (?t zmax) < Suc y"
by (simp add: \<open>?Q zmax\<close> trace_length)
have "eval (Pr n f g) (y # ys) \<down>" if "y \<le> zmax" for y
- using that zmax_def `?Q zmax` assms eval_Pr_converg_le[of n f g ys zmax y] valid y_ys
+ using that zmax_def \<open>?Q zmax\<close> assms eval_Pr_converg_le[of n f g ys zmax y] valid y_ys
by simp
then have "reachable (?stack, None) (((Pr n f g), xs, ?t y) # rest, None)"
if "y \<le> zmax" for y
- using that `?Q zmax` diverg y_ys assms reachable_Pr by simp
+ using that \<open>?Q zmax\<close> diverg y_ys assms reachable_Pr by simp
then have "reachable (?stack, None) (((Pr n f g), xs, ?t zmax) # rest, None)"
(is "reachable _ (?stack1, None)")
by simp
also have "reachable ...
((g, zmax # ?h zmax # tl xs, []) # (Pr n f g, xs, ?t zmax) # rest, None)"
(is "_ (?stack2, None)")
proof (rule step_reachable)
have "length (?t zmax) \<noteq> Suc (hd xs)"
using len y_ys by simp
moreover have "hd (?t zmax) = ?h zmax"
using trace_hd by auto
moreover have "length (?t zmax) = Suc zmax"
using trace_length by simp
ultimately show "step (?stack1, None) = (?stack2, None)"
by auto
qed
finally have "reachable (?stack, None) (?stack2, None)" .
moreover have "nonterminating (?stack2, None)"
proof -
have "correct (?stack2, None)"
using y_ys assms valid_ConsI valid by simp
moreover have "eval g (zmax # ?h zmax # ys) \<up>"
using \<open>?Q zmax\<close> diverg le_zmax len less_Suc_eq trace_length y_ys valid
by fastforce
ultimately show ?thesis using y_ys by simp
qed
ultimately have "nonterminating (?stack, None)"
using reachable_nonterminating by simp
then show ?thesis using diverg assms(1) by simp
qed
qed
lemma reachable_Mn:
assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
and "\<forall>y<z. eval f (y # xs) \<notin> {None, Some 0}"
shows "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)"
using assms(3)
proof (induction z)
case 0
then have "step (?stack, None) = ((f, 0 # xs, []) # (Mn n f, xs, [0]) # rest, None)"
using assms by simp
then show ?case
using step_reachable assms(1) by force
next
case (Suc z)
have valid: "valid rest" "recfn (length xs) (Mn n f)"
using valid_ConsE[OF assms(1)] by auto
have f: "eval f (z # xs) \<notin> {None, Some 0}"
using Suc by simp
have "reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)"
using Suc by simp
also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))"
using f assms(2)[of "z # xs"] valid correct_convergE valid_ConsI by auto
also have "reachable ... ((f, (Suc z) # xs, []) # (Mn n f, xs, [Suc z]) # rest, None)"
(is "_ (?stack1, None)")
using step_reachable f by force
finally have "reachable (?stack, None) (?stack1, None)" .
then show ?case by simp
qed
lemma iterate_step_empty_stack: "iterate t step ([], rv) = ([], rv)"
using step_empty_stack by (induction t) simp_all
lemma reachable_iterate_step_empty_stack:
assumes "reachable cfg ([], rv)"
shows "\<exists>t. iterate t step cfg = ([], rv) \<and> (\<forall>t'<t. fst (iterate t' step cfg) \<noteq> [])"
proof -
let ?P = "\<lambda>t. iterate t step cfg = ([], rv)"
from assms have "\<exists>t. ?P t"
by (simp add: reachable_def)
moreover define tmin where "tmin = Least ?P"
ultimately have "?P tmin"
using LeastI_ex[of ?P] by simp
have "fst (iterate t' step cfg) \<noteq> []" if "t' < tmin" for t'
proof
assume "fst (iterate t' step cfg) = []"
then obtain v where v: "iterate t' step cfg = ([], v)"
by (metis prod.exhaust_sel)
then have "iterate t'' step ([], v) = ([], v)" for t''
using iterate_step_empty_stack by simp
then have "iterate (t' + t'') step cfg = ([], v)" for t''
using v iterate_additive by fast
moreover obtain t'' where "t' + t'' = tmin"
using \<open>t' < tmin\<close> less_imp_add_positive by auto
ultimately have "iterate tmin step cfg = ([], v)"
by auto
then have "v = rv"
- using `?P tmin` by simp
+ using \<open>?P tmin\<close> by simp
then have "iterate t' step cfg = ([], rv)"
using v by simp
moreover have "\<forall>t'<tmin. \<not> ?P t'"
unfolding tmin_def using not_less_Least[of _ ?P] by simp
ultimately show False
using that by simp
qed
- then show ?thesis using `?P tmin` by auto
+ then show ?thesis using \<open>?P tmin\<close> by auto
qed
lemma step_Mn_correct:
assumes "valid ((Mn n f, xs, []) # rest)" (is "valid ?stack")
and "\<And>xs rest. valid ((f, xs, []) # rest) \<Longrightarrow> correct ((f, xs, []) # rest, None)"
shows "correct (?stack, None)"
proof -
have valid: "valid rest" "recfn (length xs) (Mn n f)"
using valid_ConsE[OF assms(1)] by auto
consider
(diverg) "eval (Mn n f) xs \<up>" and "\<forall>z. eval f (z # xs) \<down>"
| (diverg_f) "eval (Mn n f) xs \<up>" and "\<exists>z. eval f (z # xs) \<up>"
| (converg) "eval (Mn n f) xs \<down>"
by fast
then show ?thesis
proof (cases)
case diverg
then have "\<forall>z. eval f (z # xs) \<noteq> Some 0"
using eval_Mn_diverg[OF valid(2)] by simp
then have "\<forall>y<z. eval f (y # xs) \<notin> {None, Some 0}" for z
using diverg by simp
then have reach_z:
"\<And>z. reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)"
using reachable_Mn[OF assms] diverg by simp
define h :: "nat \<Rightarrow> configuration" where
"h z \<equiv> ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)" for z
then have h_inj: "\<And>x y. x \<noteq> y \<Longrightarrow> h x \<noteq> h y" and z_neq_Nil: "\<And>z. fst (h z) \<noteq> []"
by simp_all
have z: "\<exists>z\<^sub>0. \<forall>z>z\<^sub>0. \<not> (\<exists>t'\<le>t. iterate t' step (?stack, None) = h z)" for t
proof (induction t)
case 0
then show ?case by (metis h_inj le_zero_eq less_not_refl3)
next
case (Suc t)
then show ?case
using h_inj by (metis (no_types, hide_lams) le_Suc_eq less_not_refl3 less_trans)
qed
have "nonterminating (?stack, None)"
proof (rule ccontr)
assume "\<not> nonterminating (?stack, None)"
then obtain t where t: "fst (iterate t step (?stack, None)) = []"
by auto
then obtain z\<^sub>0 where "\<forall>z>z\<^sub>0. \<not> (\<exists>t'\<le>t. iterate t' step (?stack, None) = h z)"
using z by auto
then have not_h: "\<forall>t'\<le>t. iterate t' step (?stack, None) \<noteq> h (Suc z\<^sub>0)"
by simp
have "\<forall>t'\<ge>t. fst (iterate t' step (?stack, None)) = []"
using t iterate_step_empty_stack iterate_additive'[of t]
by (metis le_Suc_ex prod.exhaust_sel)
then have "\<forall>t'\<ge>t. iterate t' step (?stack, None) \<noteq> h (Suc z\<^sub>0)"
using z_neq_Nil by auto
then have "\<forall>t'. iterate t' step (?stack, None) \<noteq> h (Suc z\<^sub>0)"
using not_h nat_le_linear by auto
then have "\<not> reachable (?stack, None) (h (Suc z\<^sub>0))"
using reachable_def by simp
then show False
using reach_z[of "Suc z\<^sub>0"] h_def by simp
qed
then show ?thesis using diverg by simp
next
case diverg_f
let ?P = "\<lambda>z. eval f (z # xs) \<up>"
define zmin where "zmin \<equiv> Least ?P"
then have "\<forall>y<zmin. eval f (y # xs) \<notin> {None, Some 0}"
using diverg_f eval_Mn_diverg[OF valid(2)] less_trans not_less_Least[of _ ?P]
by blast
moreover have f_zmin: "eval f (zmin # xs) \<up>"
using diverg_f LeastI_ex[of ?P] zmin_def by simp
ultimately have
"reachable (?stack, None) ((f, zmin # xs, []) # (Mn n f, xs, [zmin]) # rest, None)"
(is "reachable _ (?stack1, None)")
using reachable_Mn[OF assms] by simp
moreover have "nonterminating (?stack1, None)"
using f_zmin assms valid diverg_f valid_ConsI by auto
ultimately have "nonterminating (?stack, None)"
using reachable_nonterminating by simp
then show ?thesis using diverg_f by simp
next
case converg
then obtain z where z: "eval (Mn n f) xs \<down>= z" by auto
have f_z: "eval f (z # xs) \<down>= 0"
and f_less_z: "\<And>y. y < z \<Longrightarrow> eval f (y # xs) \<down>\<noteq> 0"
using eval_Mn_convergE(2,3)[OF valid(2) z] by simp_all
then have
"reachable (?stack, None) ((f, z # xs, []) # (Mn n f, xs, [z]) # rest, None)"
using reachable_Mn[OF assms] by simp
also have "reachable ... ((Mn n f, xs, [z]) # rest, eval f (z # xs))"
using assms(2)[of "z # xs"] valid f_z valid_ConsI correct_convergE
by auto
also have "reachable ... (rest, Some z)"
using f_z f_less_z step_reachable by auto
finally have "reachable (?stack, None) (rest, Some z)" .
then show ?thesis using z by simp
qed
qed
theorem step_correct:
assumes "valid ((f, xs, []) # rest)"
shows "correct ((f, xs, []) # rest, None)"
using assms
proof (induction f arbitrary: xs rest)
case Z
then show ?case using valid_ConsE[of Z] step_reachable by auto
next
case S
then show ?case using valid_ConsE[of S] step_reachable by auto
next
case (Id m n)
then show ?case using valid_ConsE[of "Id m n"] by auto
next
case Cn
then show ?case using step_Cn_correct by presburger
next
case Pr
then show ?case using step_Pr_correct by simp
next
case Mn
then show ?case using step_Mn_correct by presburger
qed
subsection \<open>Encoding partial recursive functions\label{s:recf_enc}\<close>
text \<open>In this section we define an injective, but not surjective,
mapping from @{typ recf}s to natural numbers.\<close>
abbreviation triple_encode :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
"triple_encode x y z \<equiv> prod_encode (x, prod_encode (y, z))"
abbreviation quad_encode :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
"quad_encode w x y z \<equiv> prod_encode (w, prod_encode (x, prod_encode (y, z)))"
fun encode :: "recf \<Rightarrow> nat" where
"encode Z = 0"
| "encode S = 1"
| "encode (Id m n) = triple_encode 2 m n"
| "encode (Cn n f gs) = quad_encode 3 n (encode f) (list_encode (map encode gs))"
| "encode (Pr n f g) = quad_encode 4 n (encode f) (encode g)"
| "encode (Mn n f) = triple_encode 5 n (encode f)"
lemma prod_encode_gr1: "a > 1 \<Longrightarrow> prod_encode (a, x) > 1"
using le_prod_encode_1 less_le_trans by blast
lemma encode_not_Z_or_S: "encode f = prod_encode (a, b) \<Longrightarrow> a > 1 \<Longrightarrow> f \<noteq> Z \<and> f \<noteq> S"
by (metis encode.simps(1) encode.simps(2) less_numeral_extra(4) not_one_less_zero
prod_encode_gr1)
lemma encode_injective: "encode f = encode g \<Longrightarrow> f = g"
proof (induction g arbitrary: f)
case Z
have "\<And>a x. a > 1 \<Longrightarrow> prod_encode (a, x) > 0"
using prod_encode_gr1 by (meson less_one less_trans)
then have "f \<noteq> Z \<Longrightarrow> encode f > 0"
by (cases f) auto
then have "encode f = 0 \<Longrightarrow> f = Z" by fastforce
then show ?case using Z by simp
next
case S
have "\<And>a x. a > 1 \<Longrightarrow> prod_encode (a, x) \<noteq> Suc 0"
using prod_encode_gr1 by (metis One_nat_def less_numeral_extra(4))
then have "encode f = 1 \<Longrightarrow> f = S"
by (cases f) auto
then show ?case using S by simp
next
case Id
then obtain z where *: "encode f = prod_encode (2, z)" by simp
show ?case
using Id by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq)
next
case Cn
then obtain z where *: "encode f = prod_encode (3, z)" by simp
show ?case
proof (cases f)
case Z
then show ?thesis using * encode_not_Z_or_S by simp
next
case S
then show ?thesis using * encode_not_Z_or_S by simp
next
case Id
then show ?thesis using * by (simp add: prod_encode_eq)
next
case Cn
then show ?thesis
using * Cn.IH Cn.prems list_decode_encode
by (smt encode.simps(4) fst_conv list.inj_map_strong prod_encode_eq snd_conv)
next
case Pr
then show ?thesis using * by (simp add: prod_encode_eq)
next
case Mn
then show ?thesis using * by (simp add: prod_encode_eq)
qed
next
case Pr
then obtain z where *: "encode f = prod_encode (4, z)" by simp
show ?case
using Pr by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq)
next
case Mn
then obtain z where *: "encode f = prod_encode (5, z)" by simp
show ?case
using Mn by (cases f) (simp_all add: * encode_not_Z_or_S prod_encode_eq)
qed
definition encode_kind :: "nat \<Rightarrow> nat" where
"encode_kind e \<equiv> if e = 0 then 0 else if e = 1 then 1 else pdec1 e"
lemma encode_kind_0: "encode_kind (encode Z) = 0"
unfolding encode_kind_def by simp
lemma encode_kind_1: "encode_kind (encode S) = 1"
unfolding encode_kind_def by simp
lemma encode_kind_2: "encode_kind (encode (Id m n)) = 2"
unfolding encode_kind_def
by (metis encode.simps(1-3) encode_injective fst_conv prod_encode_inverse
recf.simps(16) recf.simps(8))
lemma encode_kind_3: "encode_kind (encode (Cn n f gs)) = 3"
unfolding encode_kind_def
by (metis encode.simps(1,2,4) encode_injective fst_conv prod_encode_inverse
recf.simps(10) recf.simps(18))
lemma encode_kind_4: "encode_kind (encode (Pr n f g)) = 4"
unfolding encode_kind_def
by (metis encode.simps(1,2,5) encode_injective fst_conv prod_encode_inverse
recf.simps(12) recf.simps(20))
lemma encode_kind_5: "encode_kind (encode (Mn n f)) = 5"
unfolding encode_kind_def
by (metis encode.simps(1,2,6) encode_injective fst_conv prod_encode_inverse
recf.simps(14) recf.simps(22))
lemmas encode_kind_n =
encode_kind_0 encode_kind_1 encode_kind_2 encode_kind_3 encode_kind_4 encode_kind_5
lemma encode_kind_Cn:
assumes "encode_kind (encode f) = 3"
shows "\<exists>n f' gs. f = Cn n f' gs"
using assms encode_kind_n by (cases f) auto
lemma encode_kind_Pr:
assumes "encode_kind (encode f) = 4"
shows "\<exists>n f' g. f = Pr n f' g"
using assms encode_kind_n by (cases f) auto
lemma encode_kind_Mn:
assumes "encode_kind (encode f) = 5"
shows "\<exists>n g. f = Mn n g"
using assms encode_kind_n by (cases f) auto
lemma pdec2_encode_Id: "pdec2 (encode (Id m n)) = prod_encode (m, n)"
by simp
lemma pdec2_encode_Pr: "pdec2 (encode (Pr n f g)) = triple_encode n (encode f) (encode g)"
by simp
subsection \<open>The step function on encoded configurations\label{s:step_enc}\<close>
text \<open>In this section we construct a function @{text "estep :: nat
\<Rightarrow> nat"} that is equivalent to the function @{text "step ::
configuration \<Rightarrow> configuration"} except that it applies to encoded
configurations. We start by defining an encoding for configurations.\<close>
definition encode_frame :: "frame \<Rightarrow> nat" where
"encode_frame s \<equiv>
triple_encode (encode (fst s)) (list_encode (fst (snd s))) (list_encode (snd (snd s)))"
lemma encode_frame:
"encode_frame (f, xs, ls) = triple_encode (encode f) (list_encode xs) (list_encode ls)"
unfolding encode_frame_def by simp
abbreviation encode_option :: "nat option \<Rightarrow> nat" where
"encode_option x \<equiv> if x = None then 0 else Suc (the x)"
definition encode_config :: "configuration \<Rightarrow> nat" where
"encode_config cfg \<equiv>
prod_encode (list_encode (map encode_frame (fst cfg)), encode_option (snd cfg))"
lemma encode_config:
"encode_config (ss, rv) = prod_encode (list_encode (map encode_frame ss), encode_option rv)"
unfolding encode_config_def by simp
text \<open>Various projections from encoded configurations:\<close>
definition e2stack where "e2stack e \<equiv> pdec1 e"
definition e2rv where "e2rv e \<equiv> pdec2 e"
definition e2tail where "e2tail e \<equiv> e_tl (e2stack e)"
definition e2frame where "e2frame e \<equiv> e_hd (e2stack e)"
definition e2i where "e2i e \<equiv> pdec1 (e2frame e)"
definition e2xs where "e2xs e \<equiv> pdec12 (e2frame e)"
definition e2ls where "e2ls e \<equiv> pdec22 (e2frame e)"
definition e2lenas where "e2lenas e \<equiv> e_length (e2xs e)"
definition e2lenls where "e2lenls e \<equiv> e_length (e2ls e)"
lemma e2rv_rv [simp]:
"e2rv (encode_config (ss, rv)) = (if rv \<up> then 0 else Suc (the rv))"
unfolding e2rv_def using encode_config by simp
lemma e2stack_stack [simp]:
"e2stack (encode_config (ss, rv)) = list_encode (map encode_frame ss)"
unfolding e2stack_def using encode_config by simp
lemma e2tail_tail [simp]:
"e2tail (encode_config (s # ss, rv)) = list_encode (map encode_frame ss)"
unfolding e2tail_def using encode_config by fastforce
lemma e2frame_frame [simp]:
"e2frame (encode_config (s # ss, rv)) = encode_frame s"
unfolding e2frame_def using encode_config by fastforce
lemma e2i_f [simp]:
"e2i (encode_config ((f, xs, ls) # ss, rv)) = encode f"
unfolding e2i_def using encode_config e2frame_frame encode_frame by force
lemma e2xs_xs [simp]:
"e2xs (encode_config ((f, xs, ls) # ss, rv)) = list_encode xs"
using e2xs_def e2frame_frame encode_frame by force
lemma e2ls_ls [simp]:
"e2ls (encode_config ((f, xs, ls) # ss, rv)) = list_encode ls"
using e2ls_def e2frame_frame encode_frame by force
lemma e2lenas_lenas [simp]:
"e2lenas (encode_config ((f, xs, ls) # ss, rv)) = length xs"
using e2lenas_def e2frame_frame encode_frame by simp
lemma e2lenls_lenls [simp]:
"e2lenls (encode_config ((f, xs, ls) # ss, rv)) = length ls"
using e2lenls_def e2frame_frame encode_frame by simp
lemma e2stack_0_iff_Nil:
assumes "e = encode_config (ss, rv)"
shows "e2stack e = 0 \<longleftrightarrow> ss = []"
using assms
by (metis list_encode.simps(1) e2stack_stack list_encode_0 map_is_Nil_conv)
lemma e2ls_0_iff_Nil [simp]: "list_decode (e2ls e) = [] \<longleftrightarrow> e2ls e = 0"
by (metis list_decode.simps(1) list_encode_decode)
text \<open>We now define @{text eterm} piecemeal by considering the more
complicated cases @{text Cn}, @{text Pr}, and @{text Mn} separately.\<close>
definition "estep_Cn e \<equiv>
if e2lenls e = e_length (pdec222 (e2i e))
then if e2rv e = 0
then prod_encode (e_cons (triple_encode (pdec122 (e2i e)) (e2ls e) 0) (e2stack e), 0)
else prod_encode (e2tail e, e2rv e)
else if e2rv e = 0
then if e2lenls e < e_length (pdec222 (e2i e))
then prod_encode
(e_cons
(triple_encode (e_nth (pdec222 (e2i e)) (e2lenls e)) (e2xs e) 0)
(e2stack e),
0)
else prod_encode (e2tail e, e2rv e)
else prod_encode
(e_cons
(triple_encode (e2i e) (e2xs e) (e_snoc (e2ls e) (e2rv e - 1)))
(e2tail e),
0)"
lemma estep_Cn:
assumes "c = (((Cn n f gs, xs, ls) # fs), rv)"
shows "estep_Cn (encode_config c) = encode_config (step c)"
using encode_frame by (simp add: assms estep_Cn_def, simp add: encode_config assms)
definition "estep_Pr e \<equiv>
if e2ls e = 0
then if e2rv e = 0
then prod_encode
(e_cons (triple_encode (pdec122 (e2i e)) (e_tl (e2xs e)) 0) (e2stack e),
0)
else prod_encode
(e_cons (triple_encode (e2i e) (e2xs e) (singleton_encode (e2rv e - 1))) (e2tail e),
0)
else if e2lenls e = Suc (e_hd (e2xs e))
then prod_encode (e2tail e, Suc (e_hd (e2ls e)))
else if e2rv e = 0
then prod_encode
(e_cons
(triple_encode
(pdec222 (e2i e))
(e_cons (e2lenls e - 1) (e_cons (e_hd (e2ls e)) (e_tl (e2xs e))))
0)
(e2stack e),
0)
else prod_encode
(e_cons
(triple_encode (e2i e) (e2xs e) (e_cons (e2rv e - 1) (e2ls e))) (e2tail e),
0)"
lemma estep_Pr1:
assumes "c = (((Pr n f g, xs, ls) # fs), rv)"
and "ls \<noteq> []"
and "length ls \<noteq> Suc (hd xs)"
and "rv \<noteq> None"
and "recfn (length xs) (Pr n f g)"
shows "estep_Pr (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
from assms(5) have "length xs > 0" by auto
then have eq: "hd xs = e_hd (e2xs ?e)"
using assms e_hd_def by auto
have "step c = ((Pr n f g, xs, (the rv) # ls) # fs, None)"
(is "step c = (?t # ?ss, None)")
using assms by simp
then have "encode_config (step c) =
prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)"
using encode_config by simp
also have "... =
prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)"
by simp
also have "... = prod_encode (e_cons (encode_frame ?t) (e2tail ?e), 0)"
using assms(1) by simp
also have "... = prod_encode
(e_cons
(triple_encode (e2i ?e) (e2xs ?e) (e_cons (e2rv ?e - 1) (e2ls ?e)))
(e2tail ?e),
0)"
by (simp add: assms encode_frame)
finally show ?thesis
using assms eq estep_Pr_def by auto
qed
lemma estep_Pr2:
assumes "c = (((Pr n f g, xs, ls) # fs), rv)"
and "ls \<noteq> []"
and "length ls \<noteq> Suc (hd xs)"
and "rv = None"
and "recfn (length xs) (Pr n f g)"
shows "estep_Pr (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
from assms(5) have "length xs > 0" by auto
then have eq: "hd xs = e_hd (e2xs ?e)"
using assms e_hd_def by auto
have "step c = ((g, (length ls - 1) # hd ls # tl xs, []) # (Pr n f g, xs, ls) # fs, None)"
(is "step c = (?t # ?ss, None)")
using assms by simp
then have "encode_config (step c) =
prod_encode (list_encode (map encode_frame (?t # ?ss)), 0)"
using encode_config by simp
also have "... =
prod_encode (e_cons (encode_frame ?t) (list_encode (map encode_frame (?ss))), 0)"
by simp
also have "... = prod_encode (e_cons (encode_frame ?t) (e2stack ?e), 0)"
using assms(1) by simp
also have "... = prod_encode
(e_cons
(triple_encode
(pdec222 (e2i ?e))
(e_cons (e2lenls ?e - 1) (e_cons (e_hd (e2ls ?e)) (e_tl (e2xs ?e))))
0)
(e2stack ?e),
0)"
using assms(1,2) encode_frame[of g "(length ls - 1) # hd ls # tl xs" "[]"]
pdec2_encode_Pr[of n f g] e2xs_xs e2i_f e2lenls_lenls e2ls_ls e_hd
by (metis list_encode.simps(1) list.collapse list_decode_encode
prod_encode_inverse snd_conv)
finally show ?thesis
using assms eq estep_Pr_def by auto
qed
lemma estep_Pr3:
assumes "c = (((Pr n f g, xs, ls) # fs), rv)"
and "ls \<noteq> []"
and "length ls = Suc (hd xs)"
and "recfn (length xs) (Pr n f g)"
shows "estep_Pr (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
from assms(4) have "length xs > 0" by auto
then have "hd xs = e_hd (e2xs ?e)"
using assms e_hd_def by auto
then have "(length ls = Suc (hd xs)) = (e2lenls ?e = Suc (e_hd (e2xs ?e)))"
using assms by simp
then have *: "estep_Pr ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))"
using assms estep_Pr_def by auto
have "step c = (fs, Some (hd ls))"
using assms(1,2,3) by simp
then have "encode_config (step c) =
prod_encode (list_encode (map encode_frame fs), encode_option (Some (hd ls)))"
using encode_config by simp
also have "... =
prod_encode (list_encode (map encode_frame fs), encode_option (Some (e_hd (e2ls ?e))))"
using assms(1,2) e_hd_def by auto
also have "... = prod_encode (list_encode (map encode_frame fs), Suc (e_hd (e2ls ?e)))"
by simp
also have "... = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))"
using assms(1) by simp
finally have "encode_config (step c) = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))" .
then show ?thesis
using estep_Pr_def * by presburger
qed
lemma estep_Pr4:
assumes "c = (((Pr n f g, xs, ls) # fs), rv)" and "ls = []"
shows "estep_Pr (encode_config c) = encode_config (step c)"
using encode_frame
by (simp add: assms estep_Pr_def, simp add: encode_config assms)
lemma estep_Pr:
assumes "c = (((Pr n f g, xs, ls) # fs), rv)"
and "recfn (length xs) (Pr n f g)"
shows "estep_Pr (encode_config c) = encode_config (step c)"
using assms estep_Pr1 estep_Pr2 estep_Pr3 estep_Pr4 by auto
definition "estep_Mn e \<equiv>
if e2ls e = 0
then prod_encode
(e_cons
(triple_encode (pdec22 (e2i e)) (e_cons 0 (e2xs e)) 0)
(e_cons
(triple_encode (e2i e) (e2xs e) (singleton_encode 0))
(e2tail e)),
0)
else if e2rv e = 1
then prod_encode (e2tail e, Suc (e_hd (e2ls e)))
else prod_encode
(e_cons
(triple_encode (pdec22 (e2i e)) (e_cons (Suc (e_hd (e2ls e))) (e2xs e)) 0)
(e_cons
(triple_encode (e2i e) (e2xs e) (singleton_encode (Suc (e_hd (e2ls e)))))
(e2tail e)),
0)"
lemma estep_Mn:
assumes "c = (((Mn n f, xs, ls) # fs), rv)"
shows "estep_Mn (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
consider "ls \<noteq> []" and "rv \<noteq> Some 0" | "ls \<noteq> []" and "rv = Some 0" | "ls = []"
by auto
then show ?thesis
proof (cases)
case 1
then have step_c: "step c =
((f, (Suc (hd ls)) # xs, []) # (Mn n f, xs, [Suc (hd ls)]) # fs, None)"
(is "step c = ?cfg")
using assms by simp
have "estep_Mn ?e =
prod_encode
(e_cons
(triple_encode (encode f) (e_cons (Suc (hd ls)) (list_encode xs)) 0)
(e_cons
(triple_encode (encode (Mn n f)) (list_encode xs) (singleton_encode (Suc (hd ls))))
(list_encode (map encode_frame fs))),
0)"
using 1 assms e_hd_def estep_Mn_def by auto
also have "... = encode_config ?cfg"
using encode_config by (simp add: encode_frame)
finally show ?thesis
using step_c by simp
next
case 2
have "estep_Mn ?e = prod_encode (e2tail ?e, Suc (e_hd (e2ls ?e)))"
using 2 assms estep_Mn_def by auto
also have "... = prod_encode (e2tail ?e, Suc (hd ls))"
using 2 assms e_hd_def by auto
also have "... = prod_encode (list_encode (map encode_frame fs), Suc (hd ls))"
using assms by simp
also have "... = encode_config (fs, Some (hd ls))"
using encode_config by simp
finally show ?thesis
using 2 assms by simp
next
case 3
then show ?thesis
using assms encode_frame by (simp add: estep_Mn_def, simp add: encode_config)
qed
qed
definition "estep e \<equiv>
if e2stack e = 0 then prod_encode (0, e2rv e)
else if e2i e = 0 then prod_encode (e2tail e, 1)
else if e2i e = 1 then prod_encode (e2tail e, Suc (Suc (e_hd (e2xs e))))
else if encode_kind (e2i e) = 2 then
prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e))))
else if encode_kind (e2i e) = 3 then estep_Cn e
else if encode_kind (e2i e) = 4 then estep_Pr e
else if encode_kind (e2i e) = 5 then estep_Mn e
else 0"
lemma estep_Z:
assumes "c = (((Z, xs, ls) # fs), rv)"
shows "estep (encode_config c) = encode_config (step c)"
using encode_frame by (simp add: assms estep_def, simp add: encode_config assms)
lemma estep_S:
assumes "c = (((S, xs, ls) # fs), rv)"
and "recfn (length xs) (fst (hd (fst c)))"
shows "estep (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
from assms have "length xs > 0" by auto
then have eq: "hd xs = e_hd (e2xs ?e)"
using assms(1) e_hd_def by auto
then have "estep ?e = prod_encode (e2tail ?e, Suc (Suc (e_hd (e2xs ?e))))"
using assms(1) estep_def by simp
moreover have "step c = (fs, Some (Suc (hd xs)))"
using assms(1) by simp
ultimately show ?thesis
using assms(1) eq estep_def encode_config[of fs "Some (Suc (hd xs))"] by simp
qed
lemma estep_Id:
assumes "c = (((Id m n, xs, ls) # fs), rv)"
and "recfn (length xs) (fst (hd (fst c)))"
shows "estep (encode_config c) = encode_config (step c)"
proof -
let ?e = "encode_config c"
from assms have "length xs = m" and "m > 0" by auto
then have eq: "xs ! n = e_nth (e2xs ?e) n"
using assms e_hd_def by auto
moreover have "encode_kind (e2i ?e) = 2"
using assms(1) encode_kind_2 by auto
ultimately have "estep ?e =
prod_encode (e2tail ?e, Suc (e_nth (e2xs ?e) (pdec22 (e2i ?e))))"
using assms estep_def encode_kind_def by auto
moreover have "step c = (fs, Some (xs ! n))"
using assms(1) by simp
ultimately show ?thesis
using assms(1) eq encode_config[of fs "Some (xs ! n)"] by simp
qed
lemma estep:
assumes "valid (fst c)"
shows "estep (encode_config c) = encode_config (step c)"
proof (cases "fst c")
case Nil
then show ?thesis
using estep_def
by (metis list_encode.simps(1) e2rv_def e2stack_stack encode_config_def
map_is_Nil_conv prod.collapse prod_encode_inverse snd_conv step.simps(1))
next
case (Cons s fs)
then obtain f xs ls rv where c: "c = ((f, xs, ls) # fs, rv)"
by (metis prod.exhaust_sel)
with assms valid_def have lenas: "recfn (length xs) f" by simp
show ?thesis
proof (cases f)
case Z
then show ?thesis using estep_Z c by simp
next
case S
then show ?thesis using estep_S c lenas by simp
next
case Id
then show ?thesis using estep_Id c lenas by simp
next
case Cn
then show ?thesis
using estep_Cn c
by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2
encode_kind_3 encode_kind_Cn estep_def list.distinct(1) recf.distinct(13)
recf.distinct(19) recf.distinct(5))
next
case Pr
then show ?thesis
using estep_Pr c lenas
by (metis e2i_f e2stack_0_iff_Nil encode.simps(1) encode.simps(2) encode_kind_2
encode_kind_4 encode_kind_Cn encode_kind_Pr estep_def list.distinct(1) recf.distinct(15)
recf.distinct(21) recf.distinct(25) recf.distinct(7))
next
case Mn
then show ?thesis
using estep_Pr c lenas
by (metis (no_types, lifting) e2i_f e2stack_0_iff_Nil encode.simps(1)
encode.simps(2) encode_kind_2 encode_kind_5 encode_kind_Cn encode_kind_Mn encode_kind_Pr
estep_Mn estep_def list.distinct(1) recf.distinct(17) recf.distinct(23)
recf.distinct(27) recf.distinct(9))
qed
qed
subsection \<open>The step function as a partial recursive function\label{s:step_recf}\<close>
text \<open>In this section we construct a primitive recursive function
@{term r_step} computing @{term estep}. This will entail defining @{typ
recf}s for many functions defined in the previous section.\<close>
definition "r_e2stack \<equiv> r_pdec1"
lemma r_e2stack_prim: "prim_recfn 1 r_e2stack"
unfolding r_e2stack_def using r_pdec1_prim by simp
lemma r_e2stack [simp]: "eval r_e2stack [e] \<down>= e2stack e"
unfolding r_e2stack_def e2stack_def using r_pdec1_prim by simp
definition "r_e2rv \<equiv> r_pdec2"
lemma r_e2rv_prim: "prim_recfn 1 r_e2rv"
unfolding r_e2rv_def using r_pdec2_prim by simp
lemma r_e2rv [simp]: "eval r_e2rv [e] \<down>= e2rv e"
unfolding r_e2rv_def e2rv_def using r_pdec2_prim by simp
definition "r_e2tail \<equiv> Cn 1 r_tl [r_e2stack]"
lemma r_e2tail_prim: "prim_recfn 1 r_e2tail"
unfolding r_e2tail_def using r_e2stack_prim r_tl_prim by simp
lemma r_e2tail [simp]: "eval r_e2tail [e] \<down>= e2tail e"
unfolding r_e2tail_def e2tail_def using r_e2stack_prim r_tl_prim by simp
definition "r_e2frame \<equiv> Cn 1 r_hd [r_e2stack]"
lemma r_e2frame_prim: "prim_recfn 1 r_e2frame"
unfolding r_e2frame_def using r_hd_prim r_e2stack_prim by simp
lemma r_e2frame [simp]: "eval r_e2frame [e] \<down>= e2frame e"
unfolding r_e2frame_def e2frame_def using r_hd_prim r_e2stack_prim by simp
definition "r_e2i \<equiv> Cn 1 r_pdec1 [r_e2frame]"
lemma r_e2i_prim: "prim_recfn 1 r_e2i"
unfolding r_e2i_def using r_pdec12_prim r_e2frame_prim by simp
lemma r_e2i [simp]: "eval r_e2i [e] \<down>= e2i e"
unfolding r_e2i_def e2i_def using r_pdec12_prim r_e2frame_prim by simp
definition "r_e2xs \<equiv> Cn 1 r_pdec12 [r_e2frame]"
lemma r_e2xs_prim: "prim_recfn 1 r_e2xs"
unfolding r_e2xs_def using r_pdec122_prim r_e2frame_prim by simp
lemma r_e2xs [simp]: "eval r_e2xs [e] \<down>= e2xs e"
unfolding r_e2xs_def e2xs_def using r_pdec122_prim r_e2frame_prim by simp
definition "r_e2ls \<equiv> Cn 1 r_pdec22 [r_e2frame]"
lemma r_e2ls_prim: "prim_recfn 1 r_e2ls"
unfolding r_e2ls_def using r_pdec222_prim r_e2frame_prim by simp
lemma r_e2ls [simp]: "eval r_e2ls [e] \<down>= e2ls e"
unfolding r_e2ls_def e2ls_def using r_pdec222_prim r_e2frame_prim by simp
definition "r_e2lenls \<equiv> Cn 1 r_length [r_e2ls]"
lemma r_e2lenls_prim: "prim_recfn 1 r_e2lenls"
unfolding r_e2lenls_def using r_length_prim r_e2ls_prim by simp
lemma r_e2lenls [simp]: "eval r_e2lenls [e] \<down>= e2lenls e"
unfolding r_e2lenls_def e2lenls_def using r_length_prim r_e2ls_prim by simp
definition "r_kind \<equiv>
Cn 1 r_ifz [Id 1 0, Z, Cn 1 r_ifeq [Id 1 0, r_const 1, r_const 1, r_pdec1]]"
lemma r_kind_prim: "prim_recfn 1 r_kind"
unfolding r_kind_def by simp
lemma r_kind: "eval r_kind [e] \<down>= encode_kind e"
unfolding r_kind_def encode_kind_def by simp
lemmas helpers_for_r_step_prim =
r_e2i_prim
r_e2lenls_prim
r_e2ls_prim
r_e2rv_prim
r_e2xs_prim
r_e2stack_prim
r_e2tail_prim
r_e2frame_prim
text \<open>We define primitive recursive functions @{term r_step_Id}, @{term
r_step_Cn}, @{term r_step_Pr}, and @{term r_step_Mn}. The last three
correspond to @{term estep_Cn}, @{term estep_Pr}, and @{term estep_Mn} from
the previous section.\<close>
definition "r_step_Id \<equiv>
Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]]"
lemma r_step_Id:
"eval r_step_Id [e] \<down>= prod_encode (e2tail e, Suc (e_nth (e2xs e) (pdec22 (e2i e))))"
unfolding r_step_Id_def using helpers_for_r_step_prim by simp
abbreviation r_triple_encode :: "recf \<Rightarrow> recf \<Rightarrow> recf \<Rightarrow> recf" where
"r_triple_encode x y z \<equiv> Cn 1 r_prod_encode [x, Cn 1 r_prod_encode [y, z]]"
definition "r_step_Cn \<equiv>
Cn 1 r_ifeq
[r_e2lenls,
Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]],
Cn 1 r_ifz
[r_e2rv,
Cn 1 r_prod_encode
[Cn 1 r_cons [r_triple_encode (Cn 1 r_pdec122 [r_e2i]) r_e2ls Z, r_e2stack],
Z],
Cn 1 r_prod_encode [r_e2tail, r_e2rv]],
Cn 1 r_ifz
[r_e2rv,
Cn 1 r_ifless
[r_e2lenls,
Cn 1 r_length [Cn 1 r_pdec222 [r_e2i]],
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode (Cn 1 r_nth [Cn 1 r_pdec222 [r_e2i], r_e2lenls]) r_e2xs Z,
r_e2stack],
Z],
Cn 1 r_prod_encode [r_e2tail, r_e2rv]],
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode r_e2i r_e2xs (Cn 1 r_snoc [r_e2ls, Cn 1 r_dec [r_e2rv]]),
r_e2tail],
Z]]]"
lemma r_step_Cn_prim: "prim_recfn 1 r_step_Cn"
unfolding r_step_Cn_def using helpers_for_r_step_prim by simp
lemma r_step_Cn: "eval r_step_Cn [e] \<down>= estep_Cn e"
unfolding r_step_Cn_def estep_Cn_def using helpers_for_r_step_prim by simp
definition "r_step_Pr \<equiv>
Cn 1 r_ifz
[r_e2ls,
Cn 1 r_ifz
[r_e2rv,
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode (Cn 1 r_pdec122 [r_e2i]) (Cn 1 r_tl [r_e2xs]) Z,
r_e2stack],
Z],
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 r_dec [r_e2rv]]),
r_e2tail],
Z]],
Cn 1 r_ifeq
[r_e2lenls,
Cn 1 S [Cn 1 r_hd [r_e2xs]],
Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]],
Cn 1 r_ifz
[r_e2rv,
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode
(Cn 1 r_pdec222 [r_e2i])
(Cn 1 r_cons
[Cn 1 r_dec [r_e2lenls],
Cn 1 r_cons [Cn 1 r_hd [r_e2ls],
Cn 1 r_tl [r_e2xs]]])
Z,
r_e2stack],
Z],
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode r_e2i r_e2xs (Cn 1 r_cons [Cn 1 r_dec [r_e2rv], r_e2ls]),
r_e2tail],
Z]]]]"
lemma r_step_Pr_prim: "prim_recfn 1 r_step_Pr"
unfolding r_step_Pr_def using helpers_for_r_step_prim by simp
lemma r_step_Pr: "eval r_step_Pr [e] \<down>= estep_Pr e"
unfolding r_step_Pr_def estep_Pr_def using helpers_for_r_step_prim by simp
definition "r_step_Mn \<equiv>
Cn 1 r_ifz
[r_e2ls,
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode (Cn 1 r_pdec22 [r_e2i]) (Cn 1 r_cons [Z, r_e2xs]) Z,
Cn 1 r_cons
[r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Z]),
r_e2tail]],
Z],
Cn 1 r_ifeq
[r_e2rv,
r_const 1,
Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_hd [r_e2ls]]],
Cn 1 r_prod_encode
[Cn 1 r_cons
[r_triple_encode
(Cn 1 r_pdec22 [r_e2i])
(Cn 1 r_cons [Cn 1 S [Cn 1 r_hd [r_e2ls]], r_e2xs])
Z,
Cn 1 r_cons
[r_triple_encode r_e2i r_e2xs (Cn 1 r_singleton_encode [Cn 1 S [Cn 1 r_hd [r_e2ls]]]),
r_e2tail]],
Z]]]"
lemma r_step_Mn_prim: "prim_recfn 1 r_step_Mn"
unfolding r_step_Mn_def using helpers_for_r_step_prim by simp
lemma r_step_Mn: "eval r_step_Mn [e] \<down>= estep_Mn e"
unfolding r_step_Mn_def estep_Mn_def using helpers_for_r_step_prim by simp
definition "r_step \<equiv>
Cn 1 r_ifz
[r_e2stack,
Cn 1 r_prod_encode [Z, r_e2rv],
Cn 1 r_ifz
[r_e2i,
Cn 1 r_prod_encode [r_e2tail, r_const 1],
Cn 1 r_ifeq
[r_e2i,
r_const 1,
Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 S [Cn 1 r_hd [r_e2xs]]]],
Cn 1 r_ifeq
[Cn 1 r_kind [r_e2i],
r_const 2,
Cn 1 r_prod_encode [r_e2tail, Cn 1 S [Cn 1 r_nth [r_e2xs, Cn 1 r_pdec22 [r_e2i]]]],
Cn 1 r_ifeq
[Cn 1 r_kind [r_e2i],
r_const 3,
r_step_Cn,
Cn 1 r_ifeq
[Cn 1 r_kind [r_e2i],
r_const 4,
r_step_Pr,
Cn 1 r_ifeq
[Cn 1 r_kind [r_e2i], r_const 5, r_step_Mn, Z]]]]]]]"
lemma r_step_prim: "prim_recfn 1 r_step"
unfolding r_step_def
using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim
by simp
lemma r_step: "eval r_step [e] \<down>= estep e"
unfolding r_step_def estep_def
using r_kind_prim r_step_Mn_prim r_step_Pr_prim r_step_Cn_prim helpers_for_r_step_prim
r_kind r_step_Cn r_step_Pr r_step_Mn
by simp
theorem r_step_equiv_step:
assumes "valid (fst c)"
shows "eval r_step [encode_config c] \<down>= encode_config (step c)"
using r_step estep assms by simp
subsection \<open>The universal function\label{s:the_universal}\<close>
text \<open>The next function computes the configuration after arbitrarily
many steps.\<close>
definition "r_leap \<equiv>
Pr 2
(Cn 2 r_prod_encode
[Cn 2 r_singleton_encode
[Cn 2 r_prod_encode [Id 2 0, Cn 2 r_prod_encode [Id 2 1, r_constn 1 0]]],
r_constn 1 0])
(Cn 4 r_step [Id 4 1])"
lemma r_leap_prim [simp]: "prim_recfn 3 r_leap"
unfolding r_leap_def using r_step_prim by simp
lemma r_leap_total: "eval r_leap [t, i, x] \<down>"
using prim_recfn_total[OF r_leap_prim] by simp
lemma r_leap:
assumes "i = encode f" and "recfn (e_length x) f"
shows "eval r_leap [t, i, x] \<down>= encode_config (iterate t step ([(f, list_decode x, [])], None))"
proof (induction t)
case 0
then show ?case
unfolding r_leap_def using r_step_prim assms encode_config encode_frame by simp
next
case (Suc t)
let ?c = "([(f, list_decode x, [])], None)"
let ?tc = "iterate t step ?c"
have "valid (fst ?c)"
using valid_def assms by simp
then have valid: "valid (fst ?tc)"
using iterate_step_valid by simp
have "eval r_leap [Suc t, i, x] =
eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]"
by (smt One_nat_def Suc_eq_plus1 eq_numeral_Suc eval_Pr_converg_Suc list.size(3) list.size(4) nat_1_add_1 pred_numeral_simps(3) r_leap_def r_leap_prim r_leap_total)
then have "eval r_leap [Suc t, i, x] = eval (Cn 4 r_step [Id 4 1]) [t, encode_config ?tc, i, x]"
using Suc by simp
then have "eval r_leap [Suc t, i, x] = eval r_step [encode_config ?tc]"
using r_step_prim by simp
then have "eval r_leap [Suc t, i, x] \<down>= encode_config (step ?tc)"
by (simp add: r_step_equiv_step valid)
then show ?case by simp
qed
lemma step_leaves_empty_stack_empty:
assumes "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)"
shows "iterate (t + t') step ([(f, list_decode x, [])], None) = ([], Some v)"
using assms by (induction t') simp_all
text \<open>The next function is essentially a convenience wrapper around
@{term r_leap}. It returns zero if the configuration returned by @{term
r_leap} is non-final, and @{term "Suc v"} if the configuration is final with
return value $v$.\<close>
definition "r_result \<equiv>
Cn 3 r_ifz [Cn 3 r_pdec1 [r_leap], Cn 3 r_pdec2 [r_leap], r_constn 2 0]"
lemma r_result_prim [simp]: "prim_recfn 3 r_result"
unfolding r_result_def using r_leap_prim by simp
lemma r_result_total: "total r_result"
using r_result_prim by blast
lemma r_result_empty_stack_None:
assumes "i = encode f"
and "recfn (e_length x) f"
and "iterate t step ([(f, list_decode x, [])], None) = ([], None)"
shows "eval r_result [t, i, x] \<down>= 0"
unfolding r_result_def
using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim
e2rv_def e2rv_rv
by simp
lemma r_result_empty_stack_Some:
assumes "i = encode f"
and "recfn (e_length x) f"
and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)"
shows "eval r_result [t, i, x] \<down>= Suc v"
unfolding r_result_def
using assms r_leap e2stack_0_iff_Nil e2stack_def e2stack_stack r_leap_total r_leap_prim
e2rv_def e2rv_rv
by simp
lemma r_result_empty_stack_stays:
assumes "i = encode f"
and "recfn (e_length x) f"
and "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)"
shows "eval r_result [t + t', i, x] \<down>= Suc v"
using assms step_leaves_empty_stack_empty r_result_empty_stack_Some by simp
lemma r_result_nonempty_stack:
assumes "i = encode f"
and "recfn (e_length x) f"
and "fst (iterate t step ([(f, list_decode x, [])], None)) \<noteq> []"
shows "eval r_result [t, i, x] \<down>= 0"
proof -
obtain ss rv where "iterate t step ([(f, list_decode x, [])], None) = (ss, rv)"
by fastforce
moreover from this assms(3) have "ss \<noteq> []" by simp
ultimately have "eval r_leap [t, i, x] \<down>= encode_config (ss, rv)"
using assms r_leap by simp
then have "eval (Cn 3 r_pdec1 [r_leap]) [t, i, x] \<down>\<noteq> 0"
- using `ss \<noteq> []` r_leap_prim encode_config r_leap_total list_encode_0
+ using \<open>ss \<noteq> []\<close> r_leap_prim encode_config r_leap_total list_encode_0
by (auto, blast)
then show ?thesis unfolding r_result_def using r_leap_prim by auto
qed
lemma r_result_Suc:
assumes "i = encode f"
and "recfn (e_length x) f"
and "eval r_result [t, i, x] \<down>= Suc v"
shows "iterate t step ([(f, list_decode x, [])], None) = ([], Some v)"
(is "?cfg = _")
proof (cases "fst ?cfg")
case Nil
then show ?thesis
using assms r_result_empty_stack_None r_result_empty_stack_Some
by (metis Zero_not_Suc nat.inject option.collapse option.inject prod.exhaust_sel)
next
case Cons
then show ?thesis using assms r_result_nonempty_stack by simp
qed
lemma r_result_converg:
assumes "i = encode f"
and "recfn (e_length x) f"
and "eval f (list_decode x) \<down>= v"
shows "\<exists>t.
(\<forall>t'\<ge>t. eval r_result [t', i, x] \<down>= Suc v) \<and>
(\<forall>t'<t. eval r_result [t', i, x] \<down>= 0)"
proof -
let ?xs = "list_decode x"
let ?stack = "[(f, ?xs, [])]"
have "wellf f" using assms(2) by simp
moreover have "length ?xs = arity f"
using assms(2) by simp
ultimately have "correct (?stack, None)"
using step_correct valid_def by simp
with assms(3) have "reachable (?stack, None) ([], Some v)"
by simp
then obtain t where
"iterate t step (?stack, None) = ([], Some v)"
"\<forall>t'<t. fst (iterate t' step (?stack, None)) \<noteq> []"
using reachable_iterate_step_empty_stack by blast
then have t:
"eval r_result [t, i, x] \<down>= Suc v"
"\<forall>t'<t. eval r_result [t', i, x] \<down>= 0"
using r_result_empty_stack_Some r_result_nonempty_stack assms(1,2)
by simp_all
then have "eval r_result [t + t', i, x] \<down>= Suc v" for t'
using r_result_empty_stack_stays assms r_result_Suc by simp
then have "\<forall>t'\<ge>t. eval r_result [t', i, x] \<down>= Suc v"
using le_Suc_ex by blast
with t(2) show ?thesis by auto
qed
lemma r_result_diverg:
assumes "i = encode f"
and "recfn (e_length x) f"
and "eval f (list_decode x) \<up>"
shows "eval r_result [t, i, x] \<down>= 0"
proof -
let ?xs = "list_decode x"
let ?stack = "[(f, ?xs, [])]"
have "recfn (length ?xs) f"
using assms(2) by auto
then have "correct (?stack, None)"
using step_correct valid_def by simp
with assms(3) have "nonterminating (?stack, None)"
by simp
then show ?thesis
using r_result_nonempty_stack assms(1,2) by simp
qed
text \<open>Now we can define the universal partial recursive function. This
function executes @{term r_result} for increasing time bounds, waits for it
to reach a final configuration, and then extracts its result value. If no
final configuration is reached, the universal function diverges.\<close>
definition "r_univ \<equiv>
Cn 2 r_dec [Cn 2 r_result [Mn 2 (Cn 3 r_not [r_result]), Id 2 0, Id 2 1]]"
lemma r_univ_recfn [simp]: "recfn 2 r_univ"
unfolding r_univ_def by simp
theorem r_univ:
assumes "i = encode f" and "recfn (e_length x) f"
shows "eval r_univ [i, x] = eval f (list_decode x)"
proof -
let ?cond = "Cn 3 r_not [r_result]"
let ?while = "Mn 2 ?cond"
let ?res = "Cn 2 r_result [?while, Id 2 0, Id 2 1]"
let ?xs = "list_decode x"
have *: "eval ?cond [t, i, x] \<down>= (if eval r_result [t, i, x] \<down>= 0 then 1 else 0)" for t
proof -
have "eval ?cond [t, i, x] = eval r_not [the (eval r_result [t, i, x])]"
using r_result_total by simp
moreover have "eval r_result [t, i, x] \<down>"
by (simp add: r_result_total)
ultimately show ?thesis by auto
qed
show ?thesis
proof (cases "eval f ?xs \<up>")
case True
then show ?thesis
unfolding r_univ_def using * r_result_diverg[OF assms] eval_Mn_diverg by simp
next
case False
then obtain v where v: "eval f ?xs \<down>= v" by auto
then obtain t where t:
"\<forall>t'\<ge>t. eval r_result [t', i, x] \<down>= Suc v"
"\<forall>t'<t. eval r_result [t', i, x] \<down>= 0"
using r_result_converg[OF assms] by blast
then have
"\<forall>t'\<ge>t. eval ?cond [t', i, x] \<down>= 0"
"\<forall>t'<t. eval ?cond [t', i, x] \<down>= 1"
using * by simp_all
then have "eval ?while [i, x] \<down>= t"
using eval_Mn_convergI[of 2 ?cond "[i, x]" t] by simp
then have "eval ?res [i, x] = eval r_result [t, i, x]"
by simp
then have "eval ?res [i, x] \<down>= Suc v"
using t(1) by simp
then show ?thesis
unfolding r_univ_def using v by simp
qed
qed
theorem r_univ':
assumes "recfn (e_length x) f"
shows "eval r_univ [encode f, x] = eval f (list_decode x)"
using r_univ assms by simp
text \<open>Universal functions for every arity can be built from @{term "r_univ"}.\<close>
definition r_universal :: "nat \<Rightarrow> recf" where
"r_universal n \<equiv> Cn (Suc n) r_univ [Id (Suc n) 0, r_shift (r_list_encode (n - 1))]"
lemma r_universal_recfn [simp]: "n > 0 \<Longrightarrow> recfn (Suc n) (r_universal n)"
unfolding r_universal_def by simp
lemma r_universal:
assumes "recfn n f" and "length xs = n"
shows "eval (r_universal n) (encode f # xs) = eval f xs"
unfolding r_universal_def using wellf_arity_nonzero assms r_list_encode r_univ'
by fastforce
text \<open>We will mostly be concerned with computing unary functions. Hence
we introduce separate functions for this case.\<close>
definition "r_result1 \<equiv>
Cn 3 r_result [Id 3 0, Id 3 1, Cn 3 r_singleton_encode [Id 3 2]]"
lemma r_result1_prim [simp]: "prim_recfn 3 r_result1"
unfolding r_result1_def by simp
lemma r_result1_total: "total r_result1"
using Mn_free_imp_total by simp
lemma r_result1 [simp]:
"eval r_result1 [t, i, x] = eval r_result [t, i, singleton_encode x]"
unfolding r_result1_def by simp
text \<open>The following function will be our standard Gödel numbering
of all unary partial recursive functions.\<close>
definition "r_phi \<equiv> r_universal 1"
lemma r_phi_recfn [simp]: "recfn 2 r_phi"
unfolding r_phi_def by simp
theorem r_phi:
assumes "i = encode f" and "recfn 1 f"
shows "eval r_phi [i, x] = eval f [x]"
unfolding r_phi_def using r_universal assms by force
corollary r_phi':
assumes "recfn 1 f"
shows "eval r_phi [encode f, x] = eval f [x]"
using assms r_phi by simp
lemma r_phi'': "eval r_phi [i, x] = eval r_univ [i, singleton_encode x]"
unfolding r_universal_def r_phi_def using r_list_encode by simp
section \<open>Applications of the universal function\<close>
text \<open>In this section we shall see some ways @{term r_univ} and @{term r_result} can
be used.\<close>
subsection \<open>Lazy conditional evaluation\<close>
text \<open>With the help of @{term r_univ} we can now define a
\hypertarget{p:r_lifz}{lazy variant} of @{term r_ifz}, in which only one
branch is evaluated.\<close>
definition r_lazyifzero :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> recf" where
"r_lazyifzero n j\<^sub>1 j\<^sub>2 \<equiv>
Cn (Suc (Suc n)) r_univ
[Cn (Suc (Suc n)) r_ifz [Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2],
r_shift (r_list_encode n)]"
lemma r_lazyifzero_recfn: "recfn (Suc (Suc n)) (r_lazyifzero n j\<^sub>1 j\<^sub>2)"
using r_lazyifzero_def by simp
lemma r_lazyifzero:
assumes "length xs = Suc n"
and "j\<^sub>1 = encode f\<^sub>1"
and "j\<^sub>2 = encode f\<^sub>2"
and "recfn (Suc n) f\<^sub>1"
and "recfn (Suc n) f\<^sub>2"
shows "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) = (if c = 0 then eval f\<^sub>1 xs else eval f\<^sub>2 xs)"
proof -
let ?a = "r_constn (Suc n) n"
let ?b = "Cn (Suc (Suc n)) r_ifz
[Id (Suc (Suc n)) 0, r_constn (Suc n) j\<^sub>1, r_constn (Suc n) j\<^sub>2]"
let ?c = "r_shift (r_list_encode n)"
have "eval ?a (c # xs) \<down>= n"
using assms(1) by simp
moreover have "eval ?b (c # xs) \<down>= (if c = 0 then j\<^sub>1 else j\<^sub>2)"
using assms(1) by simp
moreover have "eval ?c (c # xs) \<down>= list_encode xs"
using assms(1) r_list_encode r_shift by simp
ultimately have "eval (r_lazyifzero n j\<^sub>1 j\<^sub>2) (c # xs) =
eval r_univ [if c = 0 then j\<^sub>1 else j\<^sub>2, list_encode xs]"
unfolding r_lazyifzero_def using r_lazyifzero_recfn assms(1) by simp
then show ?thesis using assms r_univ by simp
qed
definition r_lifz :: "recf \<Rightarrow> recf \<Rightarrow> recf" where
"r_lifz f g \<equiv> r_lazyifzero (arity f - 1) (encode f) (encode g)"
lemma r_lifz_recfn [simp]:
assumes "recfn n f" and "recfn n g"
shows "recfn (Suc n) (r_lifz f g)"
using assms r_lazyifzero_recfn r_lifz_def wellf_arity_nonzero by auto
lemma r_lifz [simp]:
assumes "length xs = n" and "recfn n f" and "recfn n g"
shows "eval (r_lifz f g) (c # xs) = (if c = 0 then eval f xs else eval g xs)"
using assms r_lazyifzero r_lifz_def wellf_arity_nonzero
by (metis One_nat_def Suc_pred)
subsection \<open>Enumerating the domains of partial recursive functions\<close>
text \<open>In this section we define a binary function $\mathit{enumdom}$
such that for all $i$, the domain of $\varphi_i$ equals
$\{\mathit{enumdom}(i, x) \mid \mathit{enumdom}(i, x)\!\downarrow\}$. In
other words, the image of $\mathit{enumdom}_i$ is the domain of $\varphi_i$.
First we need some more properties of @{term r_leap} and @{term r_result}.\<close>
lemma r_leap_Suc: "eval r_leap [Suc t, i, x] = eval r_step [the (eval r_leap [t, i, x])]"
proof -
have "eval r_leap [Suc t, i, x] =
eval (Cn 4 r_step [Id 4 1]) [t, the (eval r_leap [t, i, x]), i, x]"
using r_leap_total eval_Pr_converg_Suc r_leap_def
by (metis length_Cons list.size(3) numeral_2_eq_2 numeral_3_eq_3 r_leap_prim)
then show ?thesis using r_step_prim by auto
qed
lemma r_leap_Suc_saturating:
assumes "pdec1 (the (eval r_leap [t, i, x])) = 0"
shows "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]"
proof -
let ?e = "eval r_leap [t, i, x]"
have "eval r_step [the ?e] \<down>= estep (the ?e)"
using r_step by simp
then have "eval r_step [the ?e] \<down>= prod_encode (0, e2rv (the ?e))"
using estep_def assms by (simp add: e2stack_def)
then have "eval r_step [the ?e] \<down>= prod_encode (pdec1 (the ?e), pdec2 (the ?e))"
using assms by (simp add: e2rv_def)
then have "eval r_step [the ?e] \<down>= the ?e" by simp
then show ?thesis using r_leap_total r_leap_Suc by simp
qed
lemma r_result_Suc_saturating:
assumes "eval r_result [t, i, x] \<down>= Suc v"
shows "eval r_result [Suc t, i, x] \<down>= Suc v"
proof -
let ?r = "\<lambda>t. eval r_ifz [pdec1 (the (eval r_leap [t, i, x])), pdec2 (the (eval r_leap [t, i, x])), 0]"
have "?r t \<down>= Suc v"
using assms unfolding r_result_def using r_leap_total r_leap_prim by simp
then have "pdec1 (the (eval r_leap [t, i, x])) = 0"
using option.sel by fastforce
then have "eval r_leap [Suc t, i, x] = eval r_leap [t, i, x]"
using r_leap_Suc_saturating by simp
moreover have "eval r_result [t, i, x] = ?r t"
unfolding r_result_def using r_leap_total r_leap_prim by simp
moreover have "eval r_result [Suc t, i, x] = ?r (Suc t)"
unfolding r_result_def using r_leap_total r_leap_prim by simp
ultimately have "eval r_result [Suc t, i, x] = eval r_result [t, i, x]"
by simp
with assms show ?thesis by simp
qed
lemma r_result_saturating:
assumes "eval r_result [t, i, x] \<down>= Suc v"
shows "eval r_result [t + d, i, x] \<down>= Suc v"
using r_result_Suc_saturating assms by (induction d) simp_all
lemma r_result_converg':
assumes "eval r_univ [i, x] \<down>= v"
shows "\<exists>t. (\<forall>t'\<ge>t. eval r_result [t', i, x] \<down>= Suc v) \<and> (\<forall>t'<t. eval r_result [t', i, x] \<down>= 0)"
proof -
let ?f = "Cn 3 r_not [r_result]"
let ?m = "Mn 2 ?f"
have "recfn 2 ?m" by simp
have eval_m: "eval ?m [i, x] \<down>"
proof
assume "eval ?m [i, x] \<up>"
then have "eval r_univ [i, x] \<up>"
unfolding r_univ_def by simp
with assms show False by simp
qed
then obtain t where t: "eval ?m [i, x] \<down>= t"
by auto
then have f_t: "eval ?f [t, i, x] \<down>= 0" and f_less_t: "\<And>y. y < t \<Longrightarrow> eval ?f [y, i, x] \<down>\<noteq> 0"
- using eval_Mn_convergE[of 2 ?f "[i, x]" t] `recfn 2 ?m`
+ using eval_Mn_convergE[of 2 ?f "[i, x]" t] \<open>recfn 2 ?m\<close>
by (metis (no_types, lifting) One_nat_def Suc_1 length_Cons list.size(3))+
have eval_Cn2: "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \<down>"
proof
assume "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \<up>"
then have "eval r_univ [i, x] \<up>"
unfolding r_univ_def by simp
with assms show False by simp
qed
have "eval r_result [t, i, x] \<down>= Suc v"
proof (rule ccontr)
assume neq_Suc: "\<not> eval r_result [t, i, x] \<down>= Suc v"
show False
proof (cases "eval r_result [t, i, x] = None")
case True
then show ?thesis using f_t by simp
next
case False
then obtain w where w: "eval r_result [t, i, x] \<down>= w" "w \<noteq> Suc v"
using neq_Suc by auto
moreover have "eval r_result [t, i, x] \<down>\<noteq> 0"
by (rule ccontr; use f_t in auto)
ultimately have "w \<noteq> 0" by simp
have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] =
eval r_result [the (eval ?m [i, x]), i, x]"
using eval_m by simp
with w t have "eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x] \<down>= w"
by simp
moreover have "eval r_univ [i, x] =
eval r_dec [the (eval (Cn 2 r_result [?m, Id 2 0, Id 2 1]) [i, x])]"
unfolding r_univ_def using eval_Cn2 by simp
ultimately have "eval r_univ [i, x] = eval r_dec [w]" by simp
then have "eval r_univ [i, x] \<down>= w - 1" by simp
- with assms `w \<noteq> 0` w show ?thesis by simp
+ with assms \<open>w \<noteq> 0\<close> w show ?thesis by simp
qed
qed
then have "\<forall>t'\<ge>t. eval r_result [t', i, x] \<down>= Suc v"
using r_result_saturating le_Suc_ex by blast
moreover have "eval r_result [y, i, x] \<down>= 0" if "y < t" for y
proof (rule ccontr)
assume neq0: "eval r_result [y, i, x] \<noteq> Some 0"
then show False
proof (cases "eval r_result [y, i, x] = None")
case True
- then show ?thesis using f_less_t `y < t` by fastforce
+ then show ?thesis using f_less_t \<open>y < t\<close> by fastforce
next
case False
then obtain v where "eval r_result [y, i, x] \<down>= v" "v \<noteq> 0"
using neq0 by auto
then have "eval ?f [y, i, x] \<down>= 0" by simp
- then show ?thesis using f_less_t `y < t` by simp
+ then show ?thesis using f_less_t \<open>y < t\<close> by simp
qed
qed
ultimately show ?thesis by auto
qed
lemma r_result_diverg':
assumes "eval r_univ [i, x] \<up>"
shows "eval r_result [t, i, x] \<down>= 0"
proof (rule ccontr)
let ?f = "Cn 3 r_not [r_result]"
let ?m = "Mn 2 ?f"
assume "eval r_result [t, i, x] \<noteq> Some 0"
with r_result_total have "eval r_result [t, i, x] \<down>\<noteq> 0" by simp
then have "eval ?f [t, i, x] \<down>= 0" by auto
moreover have "eval ?f [y, i, x] \<down>" if "y < t" for y
using r_result_total by simp
ultimately have "\<exists>z. eval ?f (z # [i, x]) \<down>= 0 \<and> (\<forall>y<z. eval ?f (y # [i, x]) \<down>)"
by blast
then have "eval ?m [i, x] \<down>" by simp
then have "eval r_univ [i, x] \<down>"
unfolding r_univ_def using r_result_total by simp
with assms show False by simp
qed
lemma r_result_bivalent':
assumes "eval r_univ [i, x] \<down>= v"
shows "eval r_result [t, i, x] \<down>= Suc v \<or> eval r_result [t, i, x] \<down>= 0"
using r_result_converg'[OF assms] not_less by blast
lemma r_result_Some':
assumes "eval r_result [t, i, x] \<down>= Suc v"
shows "eval r_univ [i, x] \<down>= v"
proof (rule ccontr)
assume not_v: "\<not> eval r_univ [i, x] \<down>= v"
show False
proof (cases "eval r_univ [i, x] \<up>")
case True
then show ?thesis
using assms r_result_diverg' by simp
next
case False
then obtain w where w: "eval r_univ [i, x] \<down>= w" "w \<noteq> v"
using not_v by auto
then have "eval r_result [t, i, x] \<down>= Suc w \<or> eval r_result [t, i, x] \<down>= 0"
using r_result_bivalent' by simp
then show ?thesis using assms not_v w by simp
qed
qed
lemma r_result1_converg':
assumes "eval r_phi [i, x] \<down>= v"
shows "\<exists>t.
(\<forall>t'\<ge>t. eval r_result1 [t', i, x] \<down>= Suc v) \<and>
(\<forall>t'<t. eval r_result1 [t', i, x] \<down>= 0)"
using assms r_result1 r_result_converg' r_phi'' by simp
lemma r_result1_diverg':
assumes "eval r_phi [i, x] \<up>"
shows "eval r_result1 [t, i, x] \<down>= 0"
using assms r_result1 r_result_diverg' r_phi'' by simp
lemma r_result1_Some':
assumes "eval r_result1 [t, i, x] \<down>= Suc v"
shows "eval r_phi [i, x] \<down>= v"
using assms r_result1 r_result_Some' r_phi'' by simp
text \<open>The next function performs dovetailing in order to evaluate
$\varphi_i$ for every argument for arbitrarily many steps. Given $i$ and $z$,
the function decodes $z$ into a pair $(x, t$) and outputs zero (meaning
``true'') iff.\ the computation of $\varphi_i$ on input $x$ halts after at most
$t$ steps. Fixing $i$ and varying $z$ will eventually compute $\varphi_i$
for every argument in the domain of $\varphi_i$ sufficiently long for it to
converge.\<close>
definition "r_dovetail \<equiv>
Cn 2 r_not [Cn 2 r_result1 [Cn 2 r_pdec2 [Id 2 1], Id 2 0, Cn 2 r_pdec1 [Id 2 1]]]"
lemma r_dovetail_prim: "prim_recfn 2 r_dovetail"
by (simp add: r_dovetail_def)
lemma r_dovetail:
"eval r_dovetail [i, z] \<down>=
(if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 then 0 else 1)"
unfolding r_dovetail_def using r_result_total by simp
text \<open>The function $\mathit{enumdom}$ works as follows in order to
enumerate exactly the domain of $\varphi_i$. Given $i$ and $y$ it searches
for the minimum $z \geq y$ for which the dovetail function returns true. This
$z$ is decoded into $(x, t)$ and the $x$ is output. In this way every value
output by $\mathit{enumdom}$ is in the domain of $\varphi_i$ by construction
of @{term r_dovetail}. Conversely an $x$ in the domain will be output for $y
= (x, t)$ where $t$ is such that $\varphi_i$ halts on $x$ within $t$
steps.\<close>
definition "r_dovedelay \<equiv>
Cn 3 r_and
[Cn 3 r_dovetail [Id 3 1, Id 3 0],
Cn 3 r_ifle [Id 3 2, Id 3 0, r_constn 2 0, r_constn 2 1]]"
lemma r_dovedelay_prim: "prim_recfn 3 r_dovedelay"
unfolding r_dovedelay_def using r_dovetail_prim by simp
lemma r_dovedelay:
"eval r_dovedelay [z, i, y] \<down>=
(if the (eval r_result1 [pdec2 z, i, pdec1 z]) > 0 \<and> y \<le> z then 0 else 1)"
by (simp add: r_dovedelay_def r_dovetail r_dovetail_prim)
definition "r_enumdom \<equiv> Cn 2 r_pdec1 [Mn 2 r_dovedelay]"
lemma r_enumdom_recfn [simp]: "recfn 2 r_enumdom"
by (simp add: r_enumdom_def r_dovedelay_prim)
lemma r_enumdom [simp]:
"eval r_enumdom [i, y] =
(if \<exists>z. eval r_dovedelay [z, i, y] \<down>= 0
then Some (pdec1 (LEAST z. eval r_dovedelay [z, i, y] \<down>= 0))
else None)"
proof -
let ?h = "Mn 2 r_dovedelay"
have "total r_dovedelay"
using r_dovedelay_prim by blast
then have "eval ?h [i, y] =
(if (\<exists>z. eval r_dovedelay [z, i, y] \<down>= 0)
then Some (LEAST z. eval r_dovedelay [z, i, y] \<down>= 0)
else None)"
using r_dovedelay_prim r_enumdom_recfn eval_Mn_convergI by simp
then show ?thesis
unfolding r_enumdom_def using r_dovedelay_prim by simp
qed
text \<open>If @{term i} is the code of the empty function, @{term r_enumdom}
has an empty domain, too.\<close>
lemma r_enumdom_empty_domain:
assumes "\<And>x. eval r_phi [i, x] \<up>"
shows "\<And>y. eval r_enumdom [i, y] \<up>"
using assms r_result1_diverg' r_dovedelay by simp
text \<open>If @{term i} is the code of a function with non-empty domain,
@{term r_enumdom} enumerates its domain.\<close>
lemma r_enumdom_nonempty_domain:
assumes "eval r_phi [i, x\<^sub>0] \<down>"
shows "\<And>y. eval r_enumdom [i, y] \<down>"
and "\<And>x. eval r_phi [i, x] \<down> \<longleftrightarrow> (\<exists>y. eval r_enumdom [i, y] \<down>= x)"
proof -
show "eval r_enumdom [i, y] \<down>" for y
proof -
obtain t where t: "\<forall>t'\<ge>t. the (eval r_result1 [t', i, x\<^sub>0]) > 0"
using assms r_result1_converg' by fastforce
let ?z = "prod_encode (x\<^sub>0, max t y)"
have "y \<le> ?z"
using le_prod_encode_2 max.bounded_iff by blast
moreover have "pdec2 ?z \<ge> t" by simp
ultimately have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0"
using t by simp
- with `y \<le> ?z` r_dovedelay have "eval r_dovedelay [?z, i, y] \<down>= 0"
+ with \<open>y \<le> ?z\<close> r_dovedelay have "eval r_dovedelay [?z, i, y] \<down>= 0"
by presburger
then show "eval r_enumdom [i, y] \<down>"
using r_enumdom by auto
qed
show "eval r_phi [i, x] \<down> = (\<exists>y. eval r_enumdom [i, y] \<down>= x)" for x
proof
show "\<exists>y. eval r_enumdom [i, y] \<down>= x" if "eval r_phi [i, x] \<down>" for x
proof -
from that obtain v where "eval r_phi [i, x] \<down>= v" by auto
then obtain t where t: "the (eval r_result1 [t, i, x]) > 0"
using r_result1_converg' assms
by (metis Zero_not_Suc dual_order.refl option.sel zero_less_iff_neq_zero)
let ?y = "prod_encode (x, t)"
have "eval r_dovedelay [?y, i, ?y] \<down>= 0"
using r_dovedelay t by simp
moreover from this have "(LEAST z. eval r_dovedelay [z, i, ?y] \<down>= 0) = ?y"
using gr_implies_not_zero r_dovedelay by (intro Least_equality; fastforce)
ultimately have "eval r_enumdom [i, ?y] \<down>= x"
using r_enumdom by auto
then show ?thesis by blast
qed
show "eval r_phi [i, x] \<down>" if "\<exists>y. eval r_enumdom [i, y] \<down>= x" for x
proof -
from that obtain y where y: "eval r_enumdom [i, y] \<down>= x"
by auto
then have "eval r_enumdom [i, y] \<down>"
by simp
then have
"\<exists>z. eval r_dovedelay [z, i, y] \<down>= 0" and
*: "eval r_enumdom [i, y] \<down>= pdec1 (LEAST z. eval r_dovedelay [z, i, y] \<down>= 0)"
(is "_ \<down>= pdec1 ?z")
using r_enumdom by metis+
then have z: "eval r_dovedelay [?z, i, y] \<down>= 0"
by (meson wellorder_Least_lemma(1))
have "the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0"
proof (rule ccontr)
assume "\<not> (the (eval r_result1 [pdec2 ?z, i, pdec1 ?z]) > 0)"
then show False
using r_dovedelay z by simp
qed
then have "eval r_phi [i, pdec1 ?z] \<down>"
using r_result1_diverg' assms by fastforce
then show ?thesis using y * by auto
qed
qed
qed
text \<open>For every $\varphi_i$ with non-empty domain there is a total
recursive function that enumerates the domain of $\varphi_i$.\<close>
lemma nonempty_domain_enumerable:
assumes "eval r_phi [i, x\<^sub>0] \<down>"
shows "\<exists>g. recfn 1 g \<and> total g \<and> (\<forall>x. eval r_phi [i, x] \<down> \<longleftrightarrow> (\<exists>y. eval g [y] \<down>= x))"
proof -
define g where "g \<equiv> Cn 1 r_enumdom [r_const i, Id 1 0]"
then have "recfn 1 g" by simp
moreover from this have "total g"
using totalI1[of g] g_def assms r_enumdom_nonempty_domain(1) by simp
moreover have "eval r_phi [i, x] \<down> \<longleftrightarrow> (\<exists>y. eval g [y] \<down>= x)" for x
unfolding g_def using r_enumdom_nonempty_domain(2)[OF assms] by simp
ultimately show ?thesis by auto
qed
subsection \<open>Concurrent evaluation of functions\<close>
text \<open>We define a function that simulates two @{typ recf}s
``concurrently'' for the same argument and returns the result of the one
converging first. If both diverge, so does the simulation function.\<close>
definition "r_both \<equiv>
Cn 4 r_ifz
[Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3],
Cn 4 r_ifz
[Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3],
Cn 4 r_prod_encode [r_constn 3 2, r_constn 3 0],
Cn 4 r_prod_encode
[r_constn 3 1, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 2, Id 4 3]]]],
Cn 4 r_prod_encode
[r_constn 3 0, Cn 4 r_dec [Cn 4 r_result1 [Id 4 0, Id 4 1, Id 4 3]]]]"
lemma r_both_prim [simp]: "prim_recfn 4 r_both"
unfolding r_both_def by simp
lemma r_both:
assumes "\<And>x. eval r_phi [i, x] = eval f [x]"
and "\<And>x. eval r_phi [j, x] = eval g [x]"
shows "eval f [x] \<up> \<and> eval g [x] \<up> \<Longrightarrow> eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)"
and "\<lbrakk>eval r_result1 [t, i, x] \<down>= 0; eval r_result1 [t, j, x] \<down>= 0\<rbrakk> \<Longrightarrow>
eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)"
and "eval r_result1 [t, i, x] \<down>= Suc v \<Longrightarrow>
eval r_both [t, i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
and "\<lbrakk>eval r_result1 [t, i, x] \<down>= 0; eval r_result1 [t, j, x] \<down>= Suc v\<rbrakk> \<Longrightarrow>
eval r_both [t, i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
proof -
have r_result_total [simp]: "eval r_result [t, k, x] \<down>" for t k x
using r_result_total by simp
{
assume "eval f [x] \<up> \<and> eval g [x] \<up>"
then have "eval r_result1 [t, i, x] \<down>= 0" and "eval r_result1 [t, j, x] \<down>= 0"
using assms r_result1_diverg' by auto
then show "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)"
unfolding r_both_def by simp
next
assume "eval r_result1 [t, i, x] \<down>= 0" and "eval r_result1 [t, j, x] \<down>= 0"
then show "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)"
unfolding r_both_def by simp
next
assume "eval r_result1 [t, i, x] \<down>= Suc v"
moreover from this have "eval r_result1 [t, i, x] \<down>= Suc (the (eval f [x]))"
using assms r_result1_Some' by fastforce
ultimately show "eval r_both [t, i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
unfolding r_both_def by auto
next
assume "eval r_result1 [t, i, x] \<down>= 0" and "eval r_result1 [t, j, x] \<down>= Suc v"
moreover from this have "eval r_result1 [t, j, x] \<down>= Suc (the (eval g [x]))"
using assms r_result1_Some' by fastforce
ultimately show "eval r_both [t, i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
unfolding r_both_def by auto
}
qed
definition "r_parallel \<equiv>
Cn 3 r_both [Mn 3 (Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]), Id 3 0, Id 3 1, Id 3 2]"
lemma r_parallel_recfn [simp]: "recfn 3 r_parallel"
unfolding r_parallel_def by simp
lemma r_parallel:
assumes "\<And>x. eval r_phi [i, x] = eval f [x]"
and "\<And>x. eval r_phi [j, x] = eval g [x]"
shows "eval f [x] \<up> \<and> eval g [x] \<up> \<Longrightarrow> eval r_parallel [i, j, x] \<up>"
and "eval f [x] \<down> \<and> eval g [x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
and "eval g [x] \<down> \<and> eval f [x] \<up> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
and "eval f [x] \<down> \<and> eval g [x] \<down> \<Longrightarrow>
eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval f [x])) \<or>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
proof -
let ?cond = "Cn 4 r_le [Cn 4 r_pdec1 [r_both], r_constn 3 1]"
define m where "m = Mn 3 ?cond"
then have m: "r_parallel = Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]"
unfolding r_parallel_def by simp
from m_def have "recfn 3 m" by simp
{
assume "eval f [x] \<up> \<and> eval g [x] \<up>"
then have "\<forall>t. eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)"
using assms r_both by simp
then have "eval ?cond [t, i, j, x] \<down>= 1" for t
by simp
then have "eval m [i, j, x] \<up>"
unfolding m_def using eval_Mn_diverg by simp
then have "eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x] \<up>"
- using `recfn 3 m` by simp
+ using \<open>recfn 3 m\<close> by simp
then show "eval r_parallel [i, j, x] \<up>"
using m by simp
next
assume "eval f [x] \<down> \<and> eval g [x] \<down>"
then obtain vf vg where v: "eval f [x] \<down>= vf" "eval g [x] \<down>= vg"
by auto
then obtain tf where tf:
"\<forall>t\<ge>tf. eval r_result1 [t, i, x] \<down>= Suc vf"
"\<forall>t<tf. eval r_result1 [t, i, x] \<down>= 0"
using r_result1_converg' assms by metis
from v obtain tg where tg:
"\<forall>t\<ge>tg. eval r_result1 [t, j, x] \<down>= Suc vg"
"\<forall>t<tg. eval r_result1 [t, j, x] \<down>= 0"
using r_result1_converg' assms by metis
show "eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval f [x])) \<or>
eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
proof (cases "tf \<le> tg")
case True
with tg(2) have j0: "\<forall>t<tf. eval r_result1 [t, j, x] \<down>= 0"
by simp
have *: "eval r_both [tf, i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
using r_both(3) assms tf(1) by simp
have "eval m [i, j, x] \<down>= tf"
unfolding m_def
proof (rule eval_Mn_convergI)
show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp
have "eval (Cn 4 r_pdec1 [r_both]) [tf, i, j, x] \<down>= 0"
using * by simp
then show "eval ?cond [tf, i, j, x] \<down>= 0" by simp
have "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)" if "t < tf" for t
using tf(2) r_both(2) assms that j0 by simp
then have "eval ?cond [t, i, j, x] \<down>= 1" if "t < tf" for t
using that by simp
then show "\<And>y. y < tf \<Longrightarrow> eval ?cond [y, i, j, x] \<down>\<noteq> 0" by simp
qed
moreover have "eval r_parallel [i, j, x] =
eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]"
using m by simp
ultimately have "eval r_parallel [i, j, x] = eval r_both [tf, i, j, x]"
- using `recfn 3 m` by simp
+ using \<open>recfn 3 m\<close> by simp
with * have "eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
by simp
then show ?thesis by simp
next
case False
with tf(2) have i0: "\<forall>t\<le>tg. eval r_result1 [t, i, x] \<down>= 0"
by simp
then have *: "eval r_both [tg, i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
using assms r_both(4) tg(1) by auto
have "eval m [i, j, x] \<down>= tg"
unfolding m_def
proof (rule eval_Mn_convergI)
show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp
have "eval (Cn 4 r_pdec1 [r_both]) [tg, i, j, x] \<down>= 1"
using * by simp
then show "eval ?cond [tg, i, j, x] \<down>= 0" by simp
have "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)" if "t < tg" for t
using tg(2) r_both(2) assms that i0 by simp
then have "eval ?cond [t, i, j, x] \<down>= 1" if "t < tg" for t
using that by simp
then show "\<And>y. y < tg \<Longrightarrow> eval ?cond [y, i, j, x] \<down>\<noteq> 0" by simp
qed
moreover have "eval r_parallel [i, j, x] =
eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]"
using m by simp
ultimately have "eval r_parallel [i, j, x] = eval r_both [tg, i, j, x]"
- using `recfn 3 m` by simp
+ using \<open>recfn 3 m\<close> by simp
with * have "eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
by simp
then show ?thesis by simp
qed
next
assume eval_fg: "eval g [x] \<down> \<and> eval f [x] \<up>"
then have i0: "\<forall>t. eval r_result1 [t, i, x] \<down>= 0"
using r_result1_diverg' assms by auto
from eval_fg obtain v where "eval g [x] \<down>= v"
by auto
then obtain t\<^sub>0 where t0:
"\<forall>t\<ge>t\<^sub>0. eval r_result1 [t, j, x] \<down>= Suc v"
"\<forall>t<t\<^sub>0. eval r_result1 [t, j, x] \<down>= 0"
using r_result1_converg' assms by metis
then have *: "eval r_both [t\<^sub>0, i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
using r_both(4) assms i0 by simp
have "eval m [i, j, x] \<down>= t\<^sub>0"
unfolding m_def
proof (rule eval_Mn_convergI)
show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp
have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \<down>= 1"
using * by simp
then show "eval ?cond [t\<^sub>0, i, j, x] \<down>= 0" by simp
have "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)" if "t < t\<^sub>0" for t
using t0(2) r_both(2) assms that i0 by simp
then have "eval ?cond [t, i, j, x] \<down>= 1" if "t < t\<^sub>0" for t
using that by simp
then show "\<And>y. y < t\<^sub>0 \<Longrightarrow> eval ?cond [y, i, j, x] \<down>\<noteq> 0" by simp
qed
moreover have "eval r_parallel [i, j, x] =
eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]"
using m by simp
ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]"
- using `recfn 3 m` by simp
+ using \<open>recfn 3 m\<close> by simp
with * show "eval r_parallel [i, j, x] \<down>= prod_encode (1, the (eval g [x]))"
by simp
next
assume eval_fg: "eval f [x] \<down> \<and> eval g [x] \<up>"
then have j0: "\<forall>t. eval r_result1 [t, j, x] \<down>= 0"
using r_result1_diverg' assms by auto
from eval_fg obtain v where "eval f [x] \<down>= v"
by auto
then obtain t\<^sub>0 where t0:
"\<forall>t\<ge>t\<^sub>0. eval r_result1 [t, i, x] \<down>= Suc v"
"\<forall>t<t\<^sub>0. eval r_result1 [t, i, x] \<down>= 0"
using r_result1_converg' assms by metis
then have *: "eval r_both [t\<^sub>0, i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
using r_both(3) assms by blast
have "eval m [i, j, x] \<down>= t\<^sub>0"
unfolding m_def
proof (rule eval_Mn_convergI)
show "recfn (length [i, j, x]) (Mn 3 ?cond)" by simp
have "eval (Cn 4 r_pdec1 [r_both]) [t\<^sub>0, i, j, x] \<down>= 0"
using * by simp
then show "eval ?cond [t\<^sub>0, i, j, x] \<down>= 0"
by simp
have "eval r_both [t, i, j, x] \<down>= prod_encode (2, 0)" if "t < t\<^sub>0" for t
using t0(2) r_both(2) assms that j0 by simp
then have "eval ?cond [t, i, j, x] \<down>= 1" if "t < t\<^sub>0" for t
using that by simp
then show "\<And>y. y < t\<^sub>0 \<Longrightarrow> eval ?cond [y, i, j, x] \<down>\<noteq> 0" by simp
qed
moreover have "eval r_parallel [i, j, x] =
eval (Cn 3 r_both [m, Id 3 0, Id 3 1, Id 3 2]) [i, j, x]"
using m by simp
ultimately have "eval r_parallel [i, j, x] = eval r_both [t\<^sub>0, i, j, x]"
- using `recfn 3 m` by simp
+ using \<open>recfn 3 m\<close> by simp
with * show "eval r_parallel [i, j, x] \<down>= prod_encode (0, the (eval f [x]))"
by simp
}
qed
end
\ No newline at end of file
diff --git a/thys/Jacobson_Basic_Algebra/Set_Theory.thy b/thys/Jacobson_Basic_Algebra/Set_Theory.thy
--- a/thys/Jacobson_Basic_Algebra/Set_Theory.thy
+++ b/thys/Jacobson_Basic_Algebra/Set_Theory.thy
@@ -1,495 +1,495 @@
(*
Copyright (c) 2014-2019 by Clemens Ballarin
This file is licensed under the 3-clause BSD license.
*)
theory Set_Theory imports "HOL-Library.FuncSet" begin
hide_const map
hide_const partition
no_notation divide (infixl "'/" 70)
no_notation inverse_divide (infixl "'/" 70)
text \<open>
Each statement in the formal text is annotated with the location of the originating statement
in Jacobson's text @{cite Jacobson1985}. Each fact that Jacobson states explicitly is marked
as @{command theorem} unless it is translated to a @{command sublocale} declaration.
Literal quotations from Jacobson's text are reproduced in double quotes.
Auxiliary results needed for the formalisation that cannot be found in Jacobson's text are marked
as @{command lemma} or are @{command interpretation}s. Such results are annotated with the
location of a related statement. For example, the introduction rule of a constant is annotated
with the location of the definition of the corresponding operation.
\<close>
section \<open>Concepts from Set Theory. The Integers\<close>
subsection \<open>The Cartesian Product Set. Maps\<close>
text \<open>Maps as extensional HOL functions\<close>
text \<open>p 5, ll 21--25\<close>
locale map =
fixes \<alpha> and S and T
assumes graph [intro, simp]: "\<alpha> \<in> S \<rightarrow>\<^sub>E T"
begin
text \<open>p 5, ll 21--25\<close>
lemma map_closed [intro, simp]:
"a \<in> S \<Longrightarrow> \<alpha> a \<in> T"
using graph by fast
text \<open>p 5, ll 21--25\<close>
lemma map_undefined [intro, simp]:
"a \<notin> S \<Longrightarrow> \<alpha> a = undefined"
using graph by fast
end (* map *)
text \<open>p 7, ll 7--8\<close>
locale surjective_map = map + assumes surjective [intro]: "\<alpha> ` S = T"
text \<open>p 7, ll 8--9\<close>
locale injective_map = map + assumes injective [intro, simp]: "inj_on \<alpha> S"
text \<open>Enables locale reasoning about the inverse @{term "restrict (inv_into S \<alpha>) T"} of @{term \<alpha>}.\<close>
text \<open>p 7, ll 9--10\<close>
locale bijective =
fixes \<alpha> and S and T
assumes bijective [intro, simp]: "bij_betw \<alpha> S T"
text \<open>
Exploit existing knowledge about @{term bij_betw} rather than extending @{locale surjective_map}
and @{locale injective_map}.
\<close>
text \<open>p 7, ll 9--10\<close>
locale bijective_map = map + bijective begin
text \<open>p 7, ll 9--10\<close>
sublocale surjective_map by unfold_locales (simp add: bij_betw_imp_surj_on)
text \<open>p 7, ll 9--10\<close>
sublocale injective_map using bij_betw_def by unfold_locales fast
text \<open>p 9, ll 12--13\<close>
sublocale inverse: map "restrict (inv_into S \<alpha>) T" T S
by unfold_locales (simp add: inv_into_into surjective)
text \<open>p 9, ll 12--13\<close>
sublocale inverse: bijective "restrict (inv_into S \<alpha>) T" T S
by unfold_locales (simp add: bij_betw_inv_into surjective)
end (* bijective_map *)
text \<open>p 8, ll 14--15\<close>
abbreviation "identity S \<equiv> (\<lambda>x \<in> S. x)"
context map begin
text \<open>p 8, ll 18--20; p 9, ll 1--8\<close>
theorem bij_betw_iff_has_inverse:
"bij_betw \<alpha> S T \<longleftrightarrow> (\<exists>\<beta> \<in> T \<rightarrow>\<^sub>E S. compose S \<beta> \<alpha> = identity S \<and> compose T \<alpha> \<beta> = identity T)"
(is "_ \<longleftrightarrow> (\<exists>\<beta> \<in> T \<rightarrow>\<^sub>E S. ?INV \<beta>)")
proof
let ?inv = "restrict (inv_into S \<alpha>) T"
assume "bij_betw \<alpha> S T"
then have "?INV ?inv" and "?inv \<in> T \<rightarrow>\<^sub>E S"
by (simp_all add: compose_inv_into_id bij_betw_imp_surj_on compose_id_inv_into bij_betw_imp_funcset bij_betw_inv_into)
then show "\<exists>\<beta> \<in> T \<rightarrow>\<^sub>E S. ?INV \<beta>" ..
next
assume "\<exists>\<beta> \<in> T \<rightarrow>\<^sub>E S. ?INV \<beta>"
then obtain \<beta> where "\<alpha> \<in> S \<rightarrow> T" "\<beta> \<in> T \<rightarrow> S" "\<And>x. x \<in> S \<Longrightarrow> \<beta> (\<alpha> x) = x" "\<And>y. y \<in> T \<Longrightarrow> \<alpha> (\<beta> y) = y"
by (metis PiE_restrict compose_eq graph restrict_PiE restrict_apply)
then show "bij_betw \<alpha> S T" by (rule bij_betwI)
qed
end (* map *)
subsection \<open>Equivalence Relations. Factoring a Map Through an Equivalence Relation\<close>
text \<open>p 11, ll 6--11\<close>
locale equivalence =
fixes S and E
assumes closed [intro, simp]: "E \<subseteq> S \<times> S"
and reflexive [intro, simp]: "a \<in> S \<Longrightarrow> (a, a) \<in> E"
and symmetric [sym]: "(a, b) \<in> E \<Longrightarrow> (b, a) \<in> E"
and transitive [trans]: "\<lbrakk> (a, b) \<in> E; (b, c) \<in> E \<rbrakk> \<Longrightarrow> (a, c) \<in> E"
begin
text \<open>p 11, ll 6--11\<close>
lemma left_closed [intro]: (* inefficient as a simp rule *)
"(a, b) \<in> E \<Longrightarrow> a \<in> S"
using closed by blast
text \<open>p 11, ll 6--11\<close>
lemma right_closed [intro]: (* inefficient as a simp rule *)
"(a, b) \<in> E \<Longrightarrow> b \<in> S"
using closed by blast
end (* equivalence *)
text \<open>p 11, ll 16--20\<close>
locale partition =
fixes S and P
assumes subset: "P \<subseteq> Pow S"
and non_vacuous: "{} \<notin> P"
and complete: "\<Union>P = S"
and disjoint: "\<lbrakk> A \<in> P; B \<in> P; A \<noteq> B \<rbrakk> \<Longrightarrow> A \<inter> B = {}"
context equivalence begin
text \<open>p 11, ll 24--26\<close>
definition "Class = (\<lambda>a \<in> S. {b \<in> S. (b, a) \<in> E})"
text \<open>p 11, ll 24--26\<close>
lemma Class_closed [dest]:
"\<lbrakk> a \<in> Class b; b \<in> S \<rbrakk> \<Longrightarrow> a \<in> S"
unfolding Class_def by auto
text \<open>p 11, ll 24--26\<close>
lemma Class_closed2 [intro, simp]:
"a \<in> S \<Longrightarrow> Class a \<subseteq> S"
unfolding Class_def by auto
text \<open>p 11, ll 24--26\<close>
lemma Class_undefined [intro, simp]:
"a \<notin> S \<Longrightarrow> Class a = undefined"
unfolding Class_def by simp
text \<open>p 11, ll 24--26\<close>
lemma ClassI [intro, simp]:
"(a, b) \<in> E \<Longrightarrow> a \<in> Class b"
unfolding Class_def by (simp add: left_closed right_closed)
text \<open>p 11, ll 24--26\<close>
lemma Class_revI [intro, simp]:
"(a, b) \<in> E \<Longrightarrow> b \<in> Class a"
by (blast intro: symmetric)
text \<open>p 11, ll 24--26\<close>
lemma ClassD [dest]:
"\<lbrakk> b \<in> Class a; a \<in> S \<rbrakk> \<Longrightarrow> (b, a) \<in> E"
unfolding Class_def by simp
text \<open>p 11, ll 30--31\<close>
theorem Class_self [intro, simp]:
"a \<in> S \<Longrightarrow> a \<in> Class a"
unfolding Class_def by simp
text \<open>p 11, l 31; p 12, l 1\<close>
theorem Class_Union [simp]:
"(\<Union>a\<in>S. Class a) = S"
by blast
text \<open>p 11, ll 2--3\<close>
theorem Class_subset:
"(a, b) \<in> E \<Longrightarrow> Class a \<subseteq> Class b"
proof
fix a and b and c
assume "(a, b) \<in> E" and "c \<in> Class a"
then have "(c, a) \<in> E" by auto
- also note `(a, b) \<in> E`
+ also note \<open>(a, b) \<in> E\<close>
finally have "(c, b) \<in> E" by simp
then show "c \<in> Class b" by auto
qed
text \<open>p 11, ll 3--4\<close>
theorem Class_eq:
"(a, b) \<in> E \<Longrightarrow> Class a = Class b"
by (auto intro!: Class_subset intro: symmetric)
text \<open>p 12, ll 1--5\<close>
theorem Class_equivalence:
"\<lbrakk> a \<in> S; b \<in> S \<rbrakk> \<Longrightarrow> Class a = Class b \<longleftrightarrow> (a, b) \<in> E"
proof
fix a and b
assume "a \<in> S" "b \<in> S" "Class a = Class b"
then have "a \<in> Class a" by auto
- also note `Class a = Class b`
- finally show "(a, b) \<in> E" by (fast intro: `b \<in> S`)
+ also note \<open>Class a = Class b\<close>
+ finally show "(a, b) \<in> E" by (fast intro: \<open>b \<in> S\<close>)
qed (rule Class_eq)
text \<open>p 12, ll 5--7\<close>
theorem not_disjoint_implies_equal:
assumes not_disjoint: "Class a \<inter> Class b \<noteq> {}"
assumes closed: "a \<in> S" "b \<in> S"
shows "Class a = Class b"
proof -
from not_disjoint and closed obtain c where witness: "c \<in> Class a \<inter> Class b" by fast
with closed have "(a, c) \<in> E" by (blast intro: symmetric)
also from witness and closed have "(c, b) \<in> E" by fast
finally show ?thesis by (rule Class_eq)
qed
text \<open>p 12, ll 7--8\<close>
definition "Partition = {Class a | a. a \<in> S}"
text \<open>p 12, ll 7--8\<close>
lemma Class_in_Partition [intro, simp]:
"a \<in> S \<Longrightarrow> Class a \<in> Partition"
unfolding Partition_def by fast
text \<open>p 12, ll 7--8\<close>
theorem partition:
"partition S Partition"
proof
fix A and B
assume closed: "A \<in> Partition" "B \<in> Partition"
then obtain a and b where eq: "A = Class a" "B = Class b" and ab: "a \<in> S" "b \<in> S"
unfolding Partition_def by auto
assume distinct: "A \<noteq> B"
then show "A \<inter> B = {}"
proof (rule contrapos_np)
assume not_disjoint: "A \<inter> B \<noteq> {}"
then show "A = B" by (simp add: eq) (metis not_disjoint_implies_equal ab)
qed
qed (auto simp: Partition_def)
end (* equivalence *)
context partition begin
text \<open>p 12, ll 9--10\<close>
theorem block_exists:
"a \<in> S \<Longrightarrow> \<exists>A. a \<in> A \<and> A \<in> P"
using complete by blast
text \<open>p 12, ll 9--10\<close>
theorem block_unique:
"\<lbrakk> a \<in> A; A \<in> P; a \<in> B; B \<in> P \<rbrakk> \<Longrightarrow> A = B"
using disjoint by blast
text \<open>p 12, ll 9--10\<close>
lemma block_closed [intro]: (* inefficient as a simp rule *)
"\<lbrakk> a \<in> A; A \<in> P \<rbrakk> \<Longrightarrow> a \<in> S"
using complete by blast
text \<open>p 12, ll 9--10\<close>
lemma element_exists:
"A \<in> P \<Longrightarrow> \<exists>a \<in> S. a \<in> A"
by (metis non_vacuous block_closed all_not_in_conv)
text \<open>p 12, ll 9--10\<close>
definition "Block = (\<lambda>a \<in> S. THE A. a \<in> A \<and> A \<in> P)"
text \<open>p 12, ll 9--10\<close>
lemma Block_closed [intro, simp]:
assumes [intro]: "a \<in> S" shows "Block a \<in> P"
proof -
obtain A where "a \<in> A" "A \<in> P" using block_exists by blast
then show ?thesis
apply (auto simp: Block_def)
apply (rule theI2)
apply (auto dest: block_unique)
done
qed
text \<open>p 12, ll 9--10\<close>
lemma Block_undefined [intro, simp]:
"a \<notin> S \<Longrightarrow> Block a = undefined"
unfolding Block_def by simp
text \<open>p 12, ll 9--10\<close>
lemma Block_self:
"\<lbrakk> a \<in> A; A \<in> P \<rbrakk> \<Longrightarrow> Block a = A"
apply (simp add: Block_def block_closed)
apply (rule the_equality)
apply (auto dest: block_unique)
done
text \<open>p 12, ll 10--11\<close>
definition "Equivalence = {(a, b) . \<exists>A \<in> P. a \<in> A \<and> b \<in> A}"
text \<open>p 12, ll 11--12\<close>
theorem equivalence: "equivalence S Equivalence"
proof
show "Equivalence \<subseteq> S \<times> S" unfolding Equivalence_def using subset by auto
next
fix a
assume "a \<in> S"
then show "(a, a) \<in> Equivalence" unfolding Equivalence_def using complete by auto
next
fix a and b
assume "(a, b) \<in> Equivalence"
then show "(b, a) \<in> Equivalence" unfolding Equivalence_def by auto
next
fix a and b and c
assume "(a, b) \<in> Equivalence" "(b, c) \<in> Equivalence"
then show "(a, c) \<in> Equivalence" unfolding Equivalence_def using disjoint by auto
qed
text \<open>Temporarily introduce equivalence associated to partition.\<close>
text \<open>p 12, ll 12--14\<close>
interpretation equivalence S Equivalence by (rule equivalence)
text \<open>p 12, ll 12--14\<close>
theorem Class_is_Block:
assumes "a \<in> S" shows "Class a = Block a"
proof -
- from `a \<in> S` and block_exists obtain A where block: "a \<in> A \<and> A \<in> P" by blast
+ from \<open>a \<in> S\<close> and block_exists obtain A where block: "a \<in> A \<and> A \<in> P" by blast
show ?thesis
apply (simp add: Block_def Class_def)
apply (rule theI2)
apply (rule block)
apply (metis block block_unique)
apply (auto dest: block_unique simp: Equivalence_def)
done
qed
text \<open>p 12, l 14\<close>
lemma Class_equals_Block:
"Class = Block"
proof
fix x show "Class x = Block x"
by (cases "x \<in> S") (auto simp: Class_is_Block)
qed
text \<open>p 12, l 14\<close>
theorem partition_of_equivalence:
"Partition = P"
by (auto simp add: Partition_def Class_equals_Block) (metis Block_self element_exists)
end (* partition *)
context equivalence begin
text \<open>p 12, ll 14--17\<close>
interpretation partition S Partition by (rule partition)
text \<open>p 12, ll 14--17\<close>
theorem equivalence_of_partition:
"Equivalence = E"
unfolding Equivalence_def unfolding Partition_def by auto (metis ClassD Class_closed Class_eq)
end (* equivalence *)
text \<open>p 12, l 14\<close>
sublocale partition \<subseteq> equivalence S Equivalence
rewrites "equivalence.Partition S Equivalence = P" and "equivalence.Class S Equivalence = Block"
apply (rule equivalence)
apply (rule partition_of_equivalence)
apply (rule Class_equals_Block)
done
text \<open>p 12, ll 14--17\<close>
sublocale equivalence \<subseteq> partition S Partition
rewrites "partition.Equivalence Partition = E" and "partition.Block S Partition = Class"
apply (rule partition)
apply (rule equivalence_of_partition)
apply (metis equivalence_of_partition partition partition.Class_equals_Block)
done
text \<open>Unfortunately only effective on input\<close>
text \<open>p 12, ll 18--20\<close>
notation equivalence.Partition (infixl "'/" 75)
context equivalence begin
text \<open>p 12, ll 18--20\<close>
lemma representant_exists [dest]: "A \<in> S / E \<Longrightarrow> \<exists>a\<in>S. a \<in> A \<and> A = Class a"
by (metis Block_self element_exists)
text \<open>p 12, ll 18--20\<close>
lemma quotient_ClassE: "A \<in> S / E \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> P (Class a)) \<Longrightarrow> P A"
using representant_exists by blast
end (* equivalence *)
text \<open>p 12, ll 21--23\<close>
sublocale equivalence \<subseteq> natural: surjective_map Class S "S / E"
by unfold_locales force+
text \<open>Technical device to achieve Jacobson's syntax; context where @{text \<alpha>} is not a parameter.\<close>
text \<open>p 12, ll 25--26\<close>
locale fiber_relation_notation = fixes S :: "'a set" begin
text \<open>p 12, ll 25--26\<close>
definition Fiber_Relation ("E'(_')") where "Fiber_Relation \<alpha> = {(x, y). x \<in> S \<and> y \<in> S \<and> \<alpha> x = \<alpha> y}"
end (* fiber_relation_notation *)
text \<open>
Context where classes and the induced map are defined through the fiber relation.
This will be the case for monoid homomorphisms but not group homomorphisms.
\<close>
text \<open>Avoids infinite interpretation chain.\<close>
text \<open>p 12, ll 25--26\<close>
locale fiber_relation = map begin
text \<open>Install syntax\<close>
text \<open>p 12, ll 25--26\<close>
sublocale fiber_relation_notation .
text \<open>p 12, ll 26--27\<close>
sublocale equivalence where E = "E(\<alpha>)"
unfolding Fiber_Relation_def by unfold_locales auto
text \<open>``define $\bar{\alpha}$ by $\bar{\alpha}(\bar{a}) = \alpha(a)$''\<close>
text \<open>p 13, ll 8--9\<close>
definition "induced = (\<lambda>A \<in> S / E(\<alpha>). THE b. \<exists>a \<in> A. b = \<alpha> a)"
text \<open>p 13, l 10\<close>
theorem Fiber_equality:
"\<lbrakk> a \<in> S; b \<in> S \<rbrakk> \<Longrightarrow> Class a = Class b \<longleftrightarrow> \<alpha> a = \<alpha> b"
unfolding Class_equivalence unfolding Fiber_Relation_def by simp
text \<open>p 13, ll 8--9\<close>
theorem induced_Fiber_simp [simp]:
assumes [intro, simp]: "a \<in> S" shows "induced (Class a) = \<alpha> a"
proof -
have "(THE b. \<exists>a\<in>Class a. b = \<alpha> a) = \<alpha> a"
by (rule the_equality) (auto simp: Fiber_equality [symmetric] Block_self block_closed)
then show ?thesis unfolding induced_def by simp
qed
text \<open>p 13, ll 10--11\<close>
interpretation induced: map induced "S / E(\<alpha>)" T
proof (unfold_locales, rule)
fix A
assume [intro, simp]: "A \<in> S / E(\<alpha>)"
obtain a where a: "a \<in> S" "a \<in> A" using element_exists by auto
have "(THE b. \<exists>a\<in>A. b = \<alpha> a) \<in> T"
apply (rule theI2)
using a apply (auto simp: Fiber_equality [symmetric] Block_self block_closed)
done
then show "induced A \<in> T" unfolding induced_def by simp
qed (simp add: induced_def)
text \<open>p 13, ll 12--13\<close>
sublocale induced: injective_map induced "S / E(\<alpha>)" T
apply unfold_locales
apply (rule inj_onI)
apply (metis Fiber_equality Block_self element_exists induced_Fiber_simp)
done
text \<open>p 13, ll 16--19\<close>
theorem factorization_lemma:
"a \<in> S \<Longrightarrow> compose S induced Class a = \<alpha> a"
by (simp add: compose_eq)
text \<open>p 13, ll 16--19\<close>
theorem factorization [simp]: "compose S induced Class = \<alpha>"
by (rule ext) (simp add: compose_def)
text \<open>p 14, ll 2--4\<close>
theorem uniqueness:
assumes map: "\<beta> \<in> S / E(\<alpha>) \<rightarrow>\<^sub>E T"
and factorization: "compose S \<beta> Class = \<alpha>"
shows "\<beta> = induced"
proof
fix A
show "\<beta> A = induced A"
proof (cases "A \<in> S / E(\<alpha>)")
case True
then obtain a where [simp]: "A = Class a" "a \<in> S" by fast
then have "\<beta> (Class a) = \<alpha> a" by (metis compose_eq factorization)
also have "\<dots> = induced (Class a)" by simp
finally show ?thesis by simp
qed (simp add: induced_def PiE_arb [OF map])
qed
end (* fiber_relation *)
end
diff --git a/thys/LTL/Disjunctive_Normal_Form.thy b/thys/LTL/Disjunctive_Normal_Form.thy
--- a/thys/LTL/Disjunctive_Normal_Form.thy
+++ b/thys/LTL/Disjunctive_Normal_Form.thy
@@ -1,1092 +1,1092 @@
(*
Author: Benedikt Seidl
Author: Salomon Sickert
License: BSD
*)
section \<open>Disjunctive Normal Form of LTL formulas\<close>
theory Disjunctive_Normal_Form
imports
LTL Equivalence_Relations "HOL-Library.FSet"
begin
text \<open>
We use the propositional representation of LTL formulas to define
the minimal disjunctive normal form of our formulas. For this purpose
we define the minimal product \<open>\<otimes>\<^sub>m\<close> and union \<open>\<union>\<^sub>m\<close>.
In the end we show that for a set \<open>\<A>\<close> of literals,
@{term "\<A> \<Turnstile>\<^sub>P \<phi>"} if, and only if, there exists a subset
of \<open>\<A>\<close> in the minimal DNF of \<open>\<phi>\<close>.
\<close>
subsection \<open>Definition of Minimum Sets\<close>
definition (in ord) min_set :: "'a set \<Rightarrow> 'a set" where
"min_set X = {y \<in> X. \<forall>x \<in> X. x \<le> y \<longrightarrow> x = y}"
lemma min_set_iff:
"x \<in> min_set X \<longleftrightarrow> x \<in> X \<and> (\<forall>y \<in> X. y \<le> x \<longrightarrow> y = x)"
unfolding min_set_def by blast
lemma min_set_subset:
"min_set X \<subseteq> X"
by (auto simp: min_set_def)
lemma min_set_idem[simp]:
"min_set (min_set X) = min_set X"
by (auto simp: min_set_def)
lemma min_set_empty[simp]:
"min_set {} = {}"
using min_set_subset by blast
lemma min_set_singleton[simp]:
"min_set {x} = {x}"
by (auto simp: min_set_def)
lemma min_set_finite:
"finite X \<Longrightarrow> finite (min_set X)"
by (simp add: min_set_def)
lemma min_set_obtains_helper:
"A \<in> B \<Longrightarrow> \<exists>C. C |\<subseteq>| A \<and> C \<in> min_set B"
proof (induction "fcard A" arbitrary: A rule: less_induct)
case less
then have "(\<forall>A'. A' \<notin> B \<or> \<not> A' |\<subseteq>| A \<or> A' = A) \<or> (\<exists>A'. A' |\<subseteq>| A \<and> A' \<in> min_set B)"
by (metis (no_types) dual_order.trans order.not_eq_order_implies_strict pfsubset_fcard_mono)
then show ?case
using less.prems min_set_def by auto
qed
lemma min_set_obtains:
assumes "A \<in> B"
obtains C where "C |\<subseteq>| A" and "C \<in> min_set B"
using min_set_obtains_helper assms by metis
subsection \<open>Minimal operators on sets\<close>
definition product :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<otimes>" 65)
where "A \<otimes> B = {a |\<union>| b | a b. a \<in> A \<and> b \<in> B}"
definition min_product :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<otimes>\<^sub>m" 65)
where "A \<otimes>\<^sub>m B = min_set (A \<otimes> B)"
definition min_union :: "'a fset set \<Rightarrow> 'a fset set \<Rightarrow> 'a fset set" (infixr "\<union>\<^sub>m" 65)
where "A \<union>\<^sub>m B = min_set (A \<union> B)"
definition product_set :: "'a fset set set \<Rightarrow> 'a fset set" ("\<Otimes>")
where "\<Otimes> X = Finite_Set.fold product {{||}} X"
definition min_product_set :: "'a fset set set \<Rightarrow> 'a fset set" ("\<Otimes>\<^sub>m")
where "\<Otimes>\<^sub>m X = Finite_Set.fold min_product {{||}} X"
lemma min_product_idem[simp]:
"A \<otimes>\<^sub>m A = min_set A"
by (auto simp: min_product_def product_def min_set_def) fastforce
lemma min_union_idem[simp]:
"A \<union>\<^sub>m A = min_set A"
by (simp add: min_union_def)
lemma product_empty[simp]:
"A \<otimes> {} = {}"
"{} \<otimes> A = {}"
by (simp_all add: product_def)
lemma min_product_empty[simp]:
"A \<otimes>\<^sub>m {} = {}"
"{} \<otimes>\<^sub>m A = {}"
by (simp_all add: min_product_def)
lemma min_union_empty[simp]:
"A \<union>\<^sub>m {} = min_set A"
"{} \<union>\<^sub>m A = min_set A"
by (simp_all add: min_union_def)
lemma product_empty_singleton[simp]:
"A \<otimes> {{||}} = A"
"{{||}} \<otimes> A = A"
by (simp_all add: product_def)
lemma min_product_empty_singleton[simp]:
"A \<otimes>\<^sub>m {{||}} = min_set A"
"{{||}} \<otimes>\<^sub>m A = min_set A"
by (simp_all add: min_product_def)
lemma product_singleton_singleton:
"A \<otimes> {{|x|}} = finsert x ` A"
"{{|x|}} \<otimes> A = finsert x ` A"
unfolding product_def by blast+
lemma product_mono:
"A \<subseteq> B \<Longrightarrow> A \<otimes> C \<subseteq> B \<otimes> C"
"B \<subseteq> C \<Longrightarrow> A \<otimes> B \<subseteq> A \<otimes> C"
unfolding product_def by auto
lemma product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes> B)"
by (simp add: product_def finite_image_set2)
lemma min_product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes>\<^sub>m B)"
by (metis min_product_def product_finite min_set_finite)
lemma min_union_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<union>\<^sub>m B)"
by (simp add: min_union_def min_set_finite)
lemma product_set_infinite[simp]:
"infinite X \<Longrightarrow> \<Otimes> X = {{||}}"
by (simp add: product_set_def)
lemma min_product_set_infinite[simp]:
"infinite X \<Longrightarrow> \<Otimes>\<^sub>m X = {{||}}"
by (simp add: min_product_set_def)
lemma product_comm:
"A \<otimes> B = B \<otimes> A"
unfolding product_def by blast
lemma min_product_comm:
"A \<otimes>\<^sub>m B = B \<otimes>\<^sub>m A"
unfolding min_product_def
by (simp add: product_comm)
lemma min_union_comm:
"A \<union>\<^sub>m B = B \<union>\<^sub>m A"
unfolding min_union_def
by (simp add: sup.commute)
lemma product_iff:
"x \<in> A \<otimes> B \<longleftrightarrow> (\<exists>a \<in> A. \<exists>b \<in> B. x = a |\<union>| b)"
unfolding product_def by blast
lemma min_product_iff:
"x \<in> A \<otimes>\<^sub>m B \<longleftrightarrow> (\<exists>a \<in> A. \<exists>b \<in> B. x = a |\<union>| b) \<and> (\<forall>a \<in> A. \<forall>b \<in> B. a |\<union>| b |\<subseteq>| x \<longrightarrow> a |\<union>| b = x)"
unfolding min_product_def min_set_iff product_iff product_def by blast
lemma min_union_iff:
"x \<in> A \<union>\<^sub>m B \<longleftrightarrow> x \<in> A \<union> B \<and> (\<forall>a \<in> A. a |\<subseteq>| x \<longrightarrow> a = x) \<and> (\<forall>b \<in> B. b |\<subseteq>| x \<longrightarrow> b = x)"
unfolding min_union_def min_set_iff by blast
lemma min_set_min_product_helper:
"x \<in> (min_set A) \<otimes>\<^sub>m B \<longleftrightarrow> x \<in> A \<otimes>\<^sub>m B"
proof
fix x
assume "x \<in> (min_set A) \<otimes>\<^sub>m B"
then obtain a b where "a \<in> min_set A" and "b \<in> B" and "x = a |\<union>| b" and 1: "\<forall>a \<in> min_set A. \<forall>b \<in> B. a |\<union>| b |\<subseteq>| x \<longrightarrow> a |\<union>| b = x"
unfolding min_product_iff by blast
moreover
{
fix a' b'
assume "a' \<in> A" and "b' \<in> B" and "a' |\<union>| b' |\<subseteq>| x"
then obtain a'' where "a'' |\<subseteq>| a'" and "a'' \<in> min_set A"
using min_set_obtains by metis
then have "a'' |\<union>| b' = x"
by (metis (full_types) 1 \<open>b' \<in> B\<close> \<open>a' |\<union>| b' |\<subseteq>| x\<close> dual_order.trans le_sup_iff)
then have "a' |\<union>| b' = x"
using \<open>a' |\<union>| b' |\<subseteq>| x\<close> \<open>a'' |\<subseteq>| a'\<close> by blast
}
ultimately show "x \<in> A \<otimes>\<^sub>m B"
by (metis min_product_iff min_set_iff)
next
fix x
assume "x \<in> A \<otimes>\<^sub>m B"
then have 1: "x \<in> A \<otimes> B" and "\<forall>y \<in> A \<otimes> B. y |\<subseteq>| x \<longrightarrow> y = x"
unfolding min_product_def min_set_iff by simp+
then have 2: "\<forall>y\<in>min_set A \<otimes> B. y |\<subseteq>| x \<longrightarrow> y = x"
by (metis product_iff min_set_iff)
then have "x \<in> min_set A \<otimes> B"
by (metis 1 funion_mono min_set_obtains order_refl product_iff)
then show "x \<in> min_set A \<otimes>\<^sub>m B"
by (simp add: 2 min_product_def min_set_iff)
qed
lemma min_set_min_product[simp]:
"(min_set A) \<otimes>\<^sub>m B = A \<otimes>\<^sub>m B"
"A \<otimes>\<^sub>m (min_set B) = A \<otimes>\<^sub>m B"
using min_product_comm min_set_min_product_helper by blast+
lemma min_set_min_union[simp]:
"(min_set A) \<union>\<^sub>m B = A \<union>\<^sub>m B"
"A \<union>\<^sub>m (min_set B) = A \<union>\<^sub>m B"
proof (unfold min_union_def min_set_def, safe)
show "\<And>x xa xb. \<lbrakk>\<forall>xa\<in>{y \<in> A. \<forall>x\<in>A. x |\<subseteq>| y \<longrightarrow> x = y} \<union> B. xa |\<subseteq>| x \<longrightarrow> xa = x; x \<in> B; xa |\<subseteq>| x; xb |\<in>| x; xa \<in> A\<rbrakk> \<Longrightarrow> xb |\<in>| xa"
by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains)
next
show "\<And>x xa xb. \<lbrakk>\<forall>xa\<in>A \<union> {y \<in> B. \<forall>x\<in>B. x |\<subseteq>| y \<longrightarrow> x = y}. xa |\<subseteq>| x \<longrightarrow> xa = x; x \<in> A; xa |\<subseteq>| x; xb |\<in>| x; xa \<in> B\<rbrakk> \<Longrightarrow> xb |\<in>| xa"
by (metis (mono_tags) UnCI dual_order.trans fequalityI min_set_def min_set_obtains)
qed blast+
lemma product_assoc[simp]:
"(A \<otimes> B) \<otimes> C = A \<otimes> (B \<otimes> C)"
proof (unfold product_def, safe)
fix a b c
assume "a \<in> A" and "c \<in> C" and "b \<in> B"
then have "b |\<union>| c \<in> {b |\<union>| c |b c. b \<in> B \<and> c \<in> C}"
by blast
then show "\<exists>a' bc. a |\<union>| b |\<union>| c = a' |\<union>| bc \<and> a' \<in> A \<and> bc \<in> {b |\<union>| c |b c. b \<in> B \<and> c \<in> C}"
- using `a \<in> A` by (metis (no_types) inf_sup_aci(5) sup_left_commute)
+ using \<open>a \<in> A\<close> by (metis (no_types) inf_sup_aci(5) sup_left_commute)
qed (metis (mono_tags, lifting) mem_Collect_eq sup_assoc)
lemma min_product_assoc[simp]:
"(A \<otimes>\<^sub>m B) \<otimes>\<^sub>m C = A \<otimes>\<^sub>m (B \<otimes>\<^sub>m C)"
unfolding min_product_def[of A B] min_product_def[of B C]
by simp (simp add: min_product_def)
lemma min_union_assoc[simp]:
"(A \<union>\<^sub>m B) \<union>\<^sub>m C = A \<union>\<^sub>m (B \<union>\<^sub>m C)"
unfolding min_union_def[of A B] min_union_def[of B C]
by simp (simp add: min_union_def sup_assoc)
lemma min_product_comp:
"a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> \<exists>c. c |\<subseteq>| (a |\<union>| b) \<and> c \<in> A \<otimes>\<^sub>m B"
by (metis (mono_tags, lifting) mem_Collect_eq min_product_def product_def min_set_obtains)
lemma min_union_comp:
"a \<in> A \<Longrightarrow> \<exists>c. c |\<subseteq>| a \<and> c \<in> A \<union>\<^sub>m B"
by (metis Un_iff min_set_obtains min_union_def)
interpretation product_set_thms: Finite_Set.comp_fun_commute product
proof unfold_locales
have "\<And>x y z. x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)"
by (simp only: product_assoc[symmetric]) (simp only: product_comm)
then show "\<And>x y. (\<otimes>) y \<circ> (\<otimes>) x = (\<otimes>) x \<circ> (\<otimes>) y"
by fastforce
qed
interpretation min_product_set_thms: Finite_Set.comp_fun_idem min_product
proof unfold_locales
have "\<And>x y z. x \<otimes>\<^sub>m (y \<otimes>\<^sub>m z) = y \<otimes>\<^sub>m (x \<otimes>\<^sub>m z)"
by (simp only: min_product_assoc[symmetric]) (simp only: min_product_comm)
then show "\<And>x y. (\<otimes>\<^sub>m) y \<circ> (\<otimes>\<^sub>m) x = (\<otimes>\<^sub>m) x \<circ> (\<otimes>\<^sub>m) y"
by fastforce
next
have "\<And>x y. x \<otimes>\<^sub>m (x \<otimes>\<^sub>m y) = x \<otimes>\<^sub>m y"
by (simp add: min_product_assoc[symmetric])
then show "\<And>x. (\<otimes>\<^sub>m) x \<circ> (\<otimes>\<^sub>m) x = (\<otimes>\<^sub>m) x"
by fastforce
qed
interpretation min_union_set_thms: Finite_Set.comp_fun_idem min_union
proof unfold_locales
have "\<And>x y z. x \<union>\<^sub>m (y \<union>\<^sub>m z) = y \<union>\<^sub>m (x \<union>\<^sub>m z)"
by (simp only: min_union_assoc[symmetric]) (simp only: min_union_comm)
then show "\<And>x y. (\<union>\<^sub>m) y \<circ> (\<union>\<^sub>m) x = (\<union>\<^sub>m) x \<circ> (\<union>\<^sub>m) y"
by fastforce
next
have "\<And>x y. x \<union>\<^sub>m (x \<union>\<^sub>m y) = x \<union>\<^sub>m y"
by (simp add: min_union_assoc[symmetric])
then show "\<And>x. (\<union>\<^sub>m) x \<circ> (\<union>\<^sub>m) x = (\<union>\<^sub>m) x"
by fastforce
qed
lemma product_set_empty[simp]:
"\<Otimes> {} = {{||}}"
"\<Otimes> {{}} = {}"
"\<Otimes> {{{||}}} = {{||}}"
by (simp_all add: product_set_def)
lemma min_product_set_empty[simp]:
"\<Otimes>\<^sub>m {} = {{||}}"
"\<Otimes>\<^sub>m {{}} = {}"
"\<Otimes>\<^sub>m {{{||}}} = {{||}}"
by (simp_all add: min_product_set_def)
lemma product_set_code[code]:
"\<Otimes> (set xs) = fold product (remdups xs) {{||}}"
by (simp add: product_set_def product_set_thms.fold_set_fold_remdups)
lemma min_product_set_code[code]:
"\<Otimes>\<^sub>m (set xs) = fold min_product (remdups xs) {{||}}"
by (simp add: min_product_set_def min_product_set_thms.fold_set_fold_remdups)
lemma product_set_insert[simp]:
"finite X \<Longrightarrow> \<Otimes> (insert x X) = x \<otimes> (\<Otimes> (X - {x}))"
unfolding product_set_def product_set_thms.fold_insert_remove ..
lemma min_product_set_insert[simp]:
"finite X \<Longrightarrow> \<Otimes>\<^sub>m (insert x X) = x \<otimes>\<^sub>m (\<Otimes>\<^sub>m X)"
unfolding min_product_set_def min_product_set_thms.fold_insert_idem ..
lemma min_product_subseteq:
"x \<in> A \<otimes>\<^sub>m B \<Longrightarrow> \<exists>a. a |\<subseteq>| x \<and> a \<in> A"
by (metis funion_upper1 min_product_iff)
lemma min_product_set_subseteq:
"finite X \<Longrightarrow> x \<in> \<Otimes>\<^sub>m X \<Longrightarrow> A \<in> X \<Longrightarrow> \<exists>a \<in> A. a |\<subseteq>| x"
by (induction X rule: finite_induct) (blast, metis finite_insert insert_absorb min_product_set_insert min_product_subseteq)
lemma min_set_product_set:
"\<Otimes>\<^sub>m A = min_set (\<Otimes> A)"
by (cases "finite A", induction A rule: finite_induct) (simp_all add: min_product_set_def product_set_def, metis min_product_def)
lemma min_product_min_set[simp]:
"min_set (A \<otimes>\<^sub>m B) = A \<otimes>\<^sub>m B"
by (simp add: min_product_def)
lemma min_union_min_set[simp]:
"min_set (A \<union>\<^sub>m B) = A \<union>\<^sub>m B"
by (simp add: min_union_def)
lemma min_product_set_min_set[simp]:
"finite X \<Longrightarrow> min_set (\<Otimes>\<^sub>m X) = \<Otimes>\<^sub>m X"
by (induction X rule: finite_induct, auto simp add: min_product_set_def min_set_iff)
lemma min_set_min_product_set[simp]:
"finite X \<Longrightarrow> \<Otimes>\<^sub>m (min_set ` X) = \<Otimes>\<^sub>m X"
by (induction X rule: finite_induct) simp_all
lemma min_product_set_union[simp]:
"finite X \<Longrightarrow> finite Y \<Longrightarrow> \<Otimes>\<^sub>m (X \<union> Y) = (\<Otimes>\<^sub>m X) \<otimes>\<^sub>m (\<Otimes>\<^sub>m Y)"
by (induction X rule: finite_induct) simp_all
lemma product_set_finite:
"(\<And>x. x \<in> X \<Longrightarrow> finite x) \<Longrightarrow> finite (\<Otimes> X)"
by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: product_set_def, insert product_finite, blast)
lemma min_product_set_finite:
"(\<And>x. x \<in> X \<Longrightarrow> finite x) \<Longrightarrow> finite (\<Otimes>\<^sub>m X)"
by (cases "finite X", rotate_tac, induction X rule: finite_induct) (simp_all add: min_product_set_def, insert min_product_finite, blast)
subsection \<open>Disjunctive Normal Form\<close>
fun dnf :: "'a ltln \<Rightarrow> 'a ltln fset set"
where
"dnf true\<^sub>n = {{||}}"
| "dnf false\<^sub>n = {}"
| "dnf (\<phi> and\<^sub>n \<psi>) = (dnf \<phi>) \<otimes> (dnf \<psi>)"
| "dnf (\<phi> or\<^sub>n \<psi>) = (dnf \<phi>) \<union> (dnf \<psi>)"
| "dnf \<phi> = {{|\<phi>|}}"
fun min_dnf :: "'a ltln \<Rightarrow> 'a ltln fset set"
where
"min_dnf true\<^sub>n = {{||}}"
| "min_dnf false\<^sub>n = {}"
| "min_dnf (\<phi> and\<^sub>n \<psi>) = (min_dnf \<phi>) \<otimes>\<^sub>m (min_dnf \<psi>)"
| "min_dnf (\<phi> or\<^sub>n \<psi>) = (min_dnf \<phi>) \<union>\<^sub>m (min_dnf \<psi>)"
| "min_dnf \<phi> = {{|\<phi>|}}"
lemma dnf_min_set:
"min_dnf \<phi> = min_set (dnf \<phi>)"
by (induction \<phi>) (simp_all, simp_all only: min_product_def min_union_def)
lemma dnf_finite:
"finite (dnf \<phi>)"
by (induction \<phi>) (auto simp: product_finite)
lemma min_dnf_finite:
"finite (min_dnf \<phi>)"
by (induction \<phi>) (auto simp: min_product_finite min_union_finite)
lemma dnf_Abs_fset[simp]:
"fset (Abs_fset (dnf \<phi>)) = dnf \<phi>"
by (simp add: dnf_finite Abs_fset_inverse)
lemma min_dnf_Abs_fset[simp]:
"fset (Abs_fset (min_dnf \<phi>)) = min_dnf \<phi>"
by (simp add: min_dnf_finite Abs_fset_inverse)
lemma dnf_prop_atoms:
"\<Phi> \<in> dnf \<phi> \<Longrightarrow> fset \<Phi> \<subseteq> prop_atoms \<phi>"
by (induction \<phi> arbitrary: \<Phi>) (auto simp: product_def, blast+)
lemma min_dnf_prop_atoms:
"\<Phi> \<in> min_dnf \<phi> \<Longrightarrow> fset \<Phi> \<subseteq> prop_atoms \<phi>"
using dnf_min_set dnf_prop_atoms min_set_subset by blast
lemma min_dnf_atoms_dnf:
"\<Phi> \<in> min_dnf \<psi> \<Longrightarrow> \<phi> \<in> fset \<Phi> \<Longrightarrow> dnf \<phi> = {{|\<phi>|}}"
proof (induction \<phi>)
case True_ltln
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(1) by blast
next
case False_ltln
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(2) by blast
next
case (And_ltln \<phi>1 \<phi>2)
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(3) by force
next
case (Or_ltln \<phi>1 \<phi>2)
then show ?case
using min_dnf_prop_atoms prop_atoms_notin(4) by force
qed auto
lemma min_dnf_min_set[simp]:
"min_set (min_dnf \<phi>) = min_dnf \<phi>"
by (induction \<phi>) (simp_all add: min_set_def min_product_def min_union_def, blast+)
lemma min_dnf_iff_prop_assignment_subset:
"\<A> \<Turnstile>\<^sub>P \<phi> \<longleftrightarrow> (\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>)"
proof
assume "\<A> \<Turnstile>\<^sub>P \<phi>"
then show "\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>"
proof (induction \<phi> arbitrary: \<A>)
case (And_ltln \<phi>\<^sub>1 \<phi>\<^sub>2)
then obtain B\<^sub>1 B\<^sub>2 where 1: "fset B\<^sub>1 \<subseteq> \<A> \<and> B\<^sub>1 \<in> min_dnf \<phi>\<^sub>1" and 2: "fset B\<^sub>2 \<subseteq> \<A> \<and> B\<^sub>2 \<in> min_dnf \<phi>\<^sub>2"
by fastforce
then obtain C where "C |\<subseteq>| B\<^sub>1 |\<union>| B\<^sub>2" and "C \<in> min_dnf \<phi>\<^sub>1 \<otimes>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_product_comp by metis
then show ?case
by (metis 1 2 le_sup_iff min_dnf.simps(3) sup.absorb_iff1 sup_fset.rep_eq)
next
case (Or_ltln \<phi>\<^sub>1 \<phi>\<^sub>2)
{
assume "\<A> \<Turnstile>\<^sub>P \<phi>\<^sub>1"
then obtain B where 1: "fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>\<^sub>1"
using Or_ltln by fastforce
then obtain C where "C |\<subseteq>| B" and "C \<in> min_dnf \<phi>\<^sub>1 \<union>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_union_comp by metis
then have ?case
by (metis 1 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4))
}
moreover
{
assume "\<A> \<Turnstile>\<^sub>P \<phi>\<^sub>2"
then obtain B where 2: "fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>\<^sub>2"
using Or_ltln by fastforce
then obtain C where "C |\<subseteq>| B" and "C \<in> min_dnf \<phi>\<^sub>1 \<union>\<^sub>m min_dnf \<phi>\<^sub>2"
using min_union_comp min_union_comm by metis
then have ?case
by (metis 2 dual_order.trans less_eq_fset.rep_eq min_dnf.simps(4))
}
ultimately show ?case
using Or_ltln.prems by auto
qed simp_all
next
assume "\<exists>B. fset B \<subseteq> \<A> \<and> B \<in> min_dnf \<phi>"
then obtain B where "fset B \<subseteq> \<A>" and "B \<in> min_dnf \<phi>"
by auto
then have "fset B \<Turnstile>\<^sub>P \<phi>"
by (induction \<phi> arbitrary: B) (auto simp: min_set_def min_product_def product_def min_union_def, blast+)
then show "\<A> \<Turnstile>\<^sub>P \<phi>"
using \<open>fset B \<subseteq> \<A>\<close> by blast
qed
lemma ltl_prop_implies_min_dnf:
"\<phi> \<longrightarrow>\<^sub>P \<psi> = (\<forall>A \<in> min_dnf \<phi>. \<exists>B \<in> min_dnf \<psi>. B |\<subseteq>| A)"
by (meson less_eq_fset.rep_eq ltl_prop_implies_def min_dnf_iff_prop_assignment_subset order_refl dual_order.trans)
lemma ltl_prop_equiv_min_dnf:
"\<phi> \<sim>\<^sub>P \<psi> = (min_dnf \<phi> = min_dnf \<psi>)"
proof
assume "\<phi> \<sim>\<^sub>P \<psi>"
then have "\<And>x. x \<in> min_set (min_dnf \<phi>) \<longleftrightarrow> x \<in> min_set (min_dnf \<psi>)"
unfolding ltl_prop_implies_equiv ltl_prop_implies_min_dnf min_set_iff
by fastforce
then show "min_dnf \<phi> = min_dnf \<psi>"
by auto
qed (simp add: ltl_prop_equiv_def min_dnf_iff_prop_assignment_subset)
lemma min_dnf_rep_abs[simp]:
"min_dnf (rep_ltln\<^sub>P (abs_ltln\<^sub>P \<phi>)) = min_dnf \<phi>"
by (simp add: ltl_prop_equiv_min_dnf[symmetric] Quotient3_ltln\<^sub>P rep_abs_rsp_left)
subsection \<open>Folding of \<open>and\<^sub>n\<close> and \<open>or\<^sub>n\<close> over Finite Sets\<close>
definition And\<^sub>n :: "'a ltln set \<Rightarrow> 'a ltln"
where
"And\<^sub>n \<Phi> \<equiv> SOME \<phi>. fold_graph And_ltln True_ltln \<Phi> \<phi>"
definition Or\<^sub>n :: "'a ltln set \<Rightarrow> 'a ltln"
where
"Or\<^sub>n \<Phi> \<equiv> SOME \<phi>. fold_graph Or_ltln False_ltln \<Phi> \<phi>"
lemma fold_graph_And\<^sub>n:
"finite \<Phi> \<Longrightarrow> fold_graph And_ltln True_ltln \<Phi> (And\<^sub>n \<Phi>)"
unfolding And\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph])
lemma fold_graph_Or\<^sub>n:
"finite \<Phi> \<Longrightarrow> fold_graph Or_ltln False_ltln \<Phi> (Or\<^sub>n \<Phi>)"
unfolding Or\<^sub>n_def by (rule someI2_ex[OF finite_imp_fold_graph])
lemma Or\<^sub>n_empty[simp]:
"Or\<^sub>n {} = False_ltln"
by (metis empty_fold_graphE finite.emptyI fold_graph_Or\<^sub>n)
lemma And\<^sub>n_empty[simp]:
"And\<^sub>n {} = True_ltln"
by (metis empty_fold_graphE finite.emptyI fold_graph_And\<^sub>n)
interpretation dnf_union_thms: Finite_Set.comp_fun_commute "\<lambda>\<phi>. (\<union>) (f \<phi>)"
by unfold_locales fastforce
interpretation dnf_product_thms: Finite_Set.comp_fun_commute "\<lambda>\<phi>. (\<otimes>) (f \<phi>)"
by unfold_locales (simp add: product_set_thms.comp_fun_commute)
\<comment> \<open>Copied from locale @{locale comp_fun_commute}\<close>
lemma fold_graph_finite:
assumes "fold_graph f z A y"
shows "finite A"
using assms by induct simp_all
text \<open>Taking the DNF of @{const And\<^sub>n} and @{const Or\<^sub>n} is the same as folding over the individual DNFs.\<close>
lemma And\<^sub>n_dnf:
"finite \<Phi> \<Longrightarrow> dnf (And\<^sub>n \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} \<Phi>"
proof (drule fold_graph_And\<^sub>n, induction rule: fold_graph.induct)
case (insertI x A y)
then have "finite A"
using fold_graph_finite by fast
then show ?case
using insertI by auto
qed simp
lemma Or\<^sub>n_dnf:
"finite \<Phi> \<Longrightarrow> dnf (Or\<^sub>n \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<union>) (dnf \<phi>)) {} \<Phi>"
proof (drule fold_graph_Or\<^sub>n, induction rule: fold_graph.induct)
case (insertI x A y)
then have "finite A"
using fold_graph_finite by fast
then show ?case
using insertI by auto
qed simp
text \<open>@{const And\<^sub>n} and @{const Or\<^sub>n} are injective on finite sets.\<close>
lemma And\<^sub>n_inj:
"inj_on And\<^sub>n {s. finite s}"
proof (standard, simp)
fix x y :: "'a ltln set"
assume "finite x" and "finite y"
then have 1: "fold_graph And_ltln True_ltln x (And\<^sub>n x)" and 2: "fold_graph And_ltln True_ltln y (And\<^sub>n y)"
using fold_graph_And\<^sub>n by blast+
assume "And\<^sub>n x = And\<^sub>n y"
with 1 show "x = y"
proof (induction rule: fold_graph.induct)
case emptyI
then show ?case
using 2 fold_graph.cases by force
next
case (insertI x A y)
with 2 show ?case
proof (induction arbitrary: x A y rule: fold_graph.induct)
case (insertI x A y)
then show ?case
by (metis fold_graph.cases insertI1 ltln.distinct(7) ltln.inject(3))
qed blast
qed
qed
lemma Or\<^sub>n_inj:
"inj_on Or\<^sub>n {s. finite s}"
proof (standard, simp)
fix x y :: "'a ltln set"
assume "finite x" and "finite y"
then have 1: "fold_graph Or_ltln False_ltln x (Or\<^sub>n x)" and 2: "fold_graph Or_ltln False_ltln y (Or\<^sub>n y)"
using fold_graph_Or\<^sub>n by blast+
assume "Or\<^sub>n x = Or\<^sub>n y"
with 1 show "x = y"
proof (induction rule: fold_graph.induct)
case emptyI
then show ?case
using 2 fold_graph.cases by force
next
case (insertI x A y)
with 2 show ?case
proof (induction arbitrary: x A y rule: fold_graph.induct)
case (insertI x A y)
then show ?case
by (metis fold_graph.cases insertI1 ltln.distinct(27) ltln.inject(4))
qed blast
qed
qed
text \<open>The semantics of @{const And\<^sub>n} and @{const Or\<^sub>n} can be expressed using quantifiers.\<close>
lemma And\<^sub>n_semantics:
"finite \<Phi> \<Longrightarrow> w \<Turnstile>\<^sub>n And\<^sub>n \<Phi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph And_ltln True_ltln \<Phi> \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<psi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_And\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_semantics:
"finite \<Phi> \<Longrightarrow> w \<Turnstile>\<^sub>n Or\<^sub>n \<Phi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph Or_ltln False_ltln \<Phi> \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<psi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_Or\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma And\<^sub>n_prop_semantics:
"finite \<Phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P And\<^sub>n \<Phi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph And_ltln True_ltln \<Phi> \<psi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<psi> \<longleftrightarrow> (\<forall>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_And\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_prop_semantics:
"finite \<Phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P Or\<^sub>n \<Phi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
proof -
assume "finite \<Phi>"
have "\<And>\<psi>. fold_graph Or_ltln False_ltln \<Phi> \<psi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<psi> \<longleftrightarrow> (\<exists>\<phi> \<in> \<Phi>. \<A> \<Turnstile>\<^sub>P \<phi>)"
by (rule fold_graph.induct) auto
then show ?thesis
using fold_graph_Or\<^sub>n[OF \<open>finite \<Phi>\<close>] by simp
qed
lemma Or\<^sub>n_And\<^sub>n_image_semantics:
assumes "finite \<A>" and "\<And>\<Phi>. \<Phi> \<in> \<A> \<Longrightarrow> finite \<Phi>"
shows "w \<Turnstile>\<^sub>n Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
proof -
have "w \<Turnstile>\<^sub>n Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. w \<Turnstile>\<^sub>n And\<^sub>n \<Phi>)"
using Or\<^sub>n_semantics assms by auto
then show ?thesis
using And\<^sub>n_semantics assms by fast
qed
lemma Or\<^sub>n_And\<^sub>n_image_prop_semantics:
assumes "finite \<A>" and "\<And>\<Phi>. \<Phi> \<in> \<A> \<Longrightarrow> finite \<Phi>"
shows "\<I> \<Turnstile>\<^sub>P Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi> \<in> \<Phi>. \<I> \<Turnstile>\<^sub>P \<phi>)"
proof -
have "\<I> \<Turnstile>\<^sub>P Or\<^sub>n (And\<^sub>n ` \<A>) \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<I> \<Turnstile>\<^sub>P And\<^sub>n \<Phi>)"
using Or\<^sub>n_prop_semantics assms by blast
then show ?thesis
using And\<^sub>n_prop_semantics assms by metis
qed
subsection \<open>DNF to LTL conversion\<close>
definition ltln_of_dnf :: "'a ltln fset set \<Rightarrow> 'a ltln"
where
"ltln_of_dnf \<A> = Or\<^sub>n (And\<^sub>n ` fset ` \<A>)"
lemma ltln_of_dnf_semantics:
assumes "finite \<A>"
shows "w \<Turnstile>\<^sub>n ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> w \<Turnstile>\<^sub>n \<phi>)"
proof -
have "finite (fset ` \<A>)"
using assms by blast
then have "w \<Turnstile>\<^sub>n ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> fset ` \<A>. \<forall>\<phi> \<in> \<Phi>. w \<Turnstile>\<^sub>n \<phi>)"
unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_semantics by fastforce
then show ?thesis
by (metis image_iff notin_fset)
qed
lemma ltln_of_dnf_prop_semantics:
assumes "finite \<A>"
shows "\<I> \<Turnstile>\<^sub>P ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> \<A>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> \<I> \<Turnstile>\<^sub>P \<phi>)"
proof -
have "finite (fset ` \<A>)"
using assms by blast
then have "\<I> \<Turnstile>\<^sub>P ltln_of_dnf \<A> \<longleftrightarrow> (\<exists>\<Phi> \<in> fset ` \<A>. \<forall>\<phi> \<in> \<Phi>. \<I> \<Turnstile>\<^sub>P \<phi>)"
unfolding ltln_of_dnf_def using Or\<^sub>n_And\<^sub>n_image_prop_semantics by fastforce
then show ?thesis
by (metis image_iff notin_fset)
qed
lemma ltln_of_dnf_prop_equiv:
"ltln_of_dnf (min_dnf \<phi>) \<sim>\<^sub>P \<phi>"
unfolding ltl_prop_equiv_def
proof
fix \<A>
have "\<A> \<Turnstile>\<^sub>P ltln_of_dnf (min_dnf \<phi>) \<longleftrightarrow> (\<exists>\<Phi> \<in> min_dnf \<phi>. \<forall>\<phi>. \<phi> |\<in>| \<Phi> \<longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>)"
using ltln_of_dnf_prop_semantics min_dnf_finite by metis
also have "\<dots> \<longleftrightarrow> (\<exists>\<Phi> \<in> min_dnf \<phi>. fset \<Phi> \<subseteq> \<A>)"
by (metis min_dnf_prop_atoms prop_atoms_entailment_iff notin_fset subset_eq)
also have "\<dots> \<longleftrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>"
using min_dnf_iff_prop_assignment_subset by blast
finally show "\<A> \<Turnstile>\<^sub>P ltln_of_dnf (min_dnf \<phi>) = \<A> \<Turnstile>\<^sub>P \<phi>" .
qed
lemma min_dnf_ltln_of_dnf[simp]:
"min_dnf (ltln_of_dnf (min_dnf \<phi>)) = min_dnf \<phi>"
using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv by blast
subsection \<open>Substitution in DNF formulas\<close>
definition subst_clause :: "'a ltln fset \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln fset set"
where
"subst_clause \<Phi> m = \<Otimes>\<^sub>m {min_dnf (subst \<phi> m) | \<phi>. \<phi> \<in> fset \<Phi>}"
definition subst_dnf :: "'a ltln fset set \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln fset set"
where
"subst_dnf \<A> m = (\<Union>\<Phi> \<in> \<A>. subst_clause \<Phi> m)"
lemma subst_clause_empty[simp]:
"subst_clause {||} m = {{||}}"
by (simp add: subst_clause_def)
lemma subst_dnf_empty[simp]:
"subst_dnf {} m = {}"
by (simp add: subst_dnf_def)
lemma subst_clause_inner_finite:
"finite {min_dnf (subst \<phi> m) | \<phi>. \<phi> \<in> \<Phi>}" if "finite \<Phi>"
using that by simp
lemma subst_clause_finite:
"finite (subst_clause \<Phi> m)"
unfolding subst_clause_def
by (auto intro: min_dnf_finite min_product_set_finite)
lemma subst_dnf_finite:
"finite \<A> \<Longrightarrow> finite (subst_dnf \<A> m)"
unfolding subst_dnf_def using subst_clause_finite by blast
lemma subst_dnf_mono:
"\<A> \<subseteq> \<B> \<Longrightarrow> subst_dnf \<A> m \<subseteq> subst_dnf \<B> m"
unfolding subst_dnf_def by blast
lemma subst_clause_min_set[simp]:
"min_set (subst_clause \<Phi> m) = subst_clause \<Phi> m"
unfolding subst_clause_def by simp
lemma subst_clause_finsert[simp]:
"subst_clause (finsert \<phi> \<Phi>) m = (min_dnf (subst \<phi> m)) \<otimes>\<^sub>m (subst_clause \<Phi> m)"
proof -
have "{min_dnf (subst \<psi> m) | \<psi>. \<psi> \<in> fset (finsert \<phi> \<Phi>)}
= insert (min_dnf (subst \<phi> m)) {min_dnf (subst \<psi> m) | \<psi>. \<psi> \<in> fset \<Phi>}"
by auto
then show ?thesis
by (simp add: subst_clause_def)
qed
lemma subst_clause_funion[simp]:
"subst_clause (\<Phi> |\<union>| \<Psi>) m = (subst_clause \<Phi> m) \<otimes>\<^sub>m (subst_clause \<Psi> m)"
proof (induction \<Psi>)
case (insert x F)
then show ?case
using min_product_set_thms.fun_left_comm by fastforce
qed simp
text \<open>For the proof of correctness, we redefine the @{const product} operator on lists.\<close>
definition list_product :: "'a list set \<Rightarrow> 'a list set \<Rightarrow> 'a list set" (infixl "\<otimes>\<^sub>l" 65)
where
"A \<otimes>\<^sub>l B = {a @ b | a b. a \<in> A \<and> b \<in> B}"
lemma list_product_fset_of_list[simp]:
"fset_of_list ` (A \<otimes>\<^sub>l B) = (fset_of_list ` A) \<otimes> (fset_of_list ` B)"
unfolding list_product_def product_def image_def by fastforce
lemma list_product_finite:
"finite A \<Longrightarrow> finite B \<Longrightarrow> finite (A \<otimes>\<^sub>l B)"
unfolding list_product_def by (simp add: finite_image_set2)
lemma list_product_iff:
"x \<in> A \<otimes>\<^sub>l B \<longleftrightarrow> (\<exists>a b. a \<in> A \<and> b \<in> B \<and> x = a @ b)"
unfolding list_product_def by blast
lemma list_product_assoc[simp]:
"A \<otimes>\<^sub>l (B \<otimes>\<^sub>l C) = A \<otimes>\<^sub>l B \<otimes>\<^sub>l C"
unfolding set_eq_iff list_product_iff by fastforce
text \<open>Furthermore, we introduct DNFs where the clauses are represented as lists.\<close>
fun list_dnf :: "'a ltln \<Rightarrow> 'a ltln list set"
where
"list_dnf true\<^sub>n = {[]}"
| "list_dnf false\<^sub>n = {}"
| "list_dnf (\<phi> and\<^sub>n \<psi>) = (list_dnf \<phi>) \<otimes>\<^sub>l (list_dnf \<psi>)"
| "list_dnf (\<phi> or\<^sub>n \<psi>) = (list_dnf \<phi>) \<union> (list_dnf \<psi>)"
| "list_dnf \<phi> = {[\<phi>]}"
definition list_dnf_to_dnf :: "'a list set \<Rightarrow> 'a fset set"
where
"list_dnf_to_dnf X = fset_of_list ` X"
lemma list_dnf_to_dnf_list_dnf[simp]:
"list_dnf_to_dnf (list_dnf \<phi>) = dnf \<phi>"
by (induction \<phi>) (simp_all add: list_dnf_to_dnf_def image_Un)
lemma list_dnf_finite:
"finite (list_dnf \<phi>)"
by (induction \<phi>) (simp_all add: list_product_finite)
text \<open>We use this to redefine @{const subst_clause} and @{const subst_dnf} on list DNFs.\<close>
definition subst_clause' :: "'a ltln list \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln list set"
where
"subst_clause' \<Phi> m = fold (\<lambda>\<phi> acc. acc \<otimes>\<^sub>l list_dnf (subst \<phi> m)) \<Phi> {[]}"
definition subst_dnf' :: "'a ltln list set \<Rightarrow> ('a ltln \<rightharpoonup> 'a ltln) \<Rightarrow> 'a ltln list set"
where
"subst_dnf' \<A> m = (\<Union>\<Phi> \<in> \<A>. subst_clause' \<Phi> m)"
lemma subst_clause'_finite:
"finite (subst_clause' \<Phi> m)"
by (induction \<Phi> rule: rev_induct) (simp_all add: subst_clause'_def list_dnf_finite list_product_finite)
lemma subst_clause'_nil[simp]:
"subst_clause' [] m = {[]}"
by (simp add: subst_clause'_def)
lemma subst_clause'_cons[simp]:
"subst_clause' (xs @ [x]) m = subst_clause' xs m \<otimes>\<^sub>l list_dnf (subst x m)"
by (simp add: subst_clause'_def)
lemma subst_clause'_append[simp]:
"subst_clause' (A @ B) m = subst_clause' A m \<otimes>\<^sub>l subst_clause' B m"
proof (induction B rule: rev_induct)
case (snoc x xs)
then show ?case
by simp (metis append_assoc subst_clause'_cons)
qed(simp add: list_product_def)
lemma subst_dnf'_iff:
"x \<in> subst_dnf' A m \<longleftrightarrow> (\<exists>\<Phi> \<in> A. x \<in> subst_clause' \<Phi> m)"
by (simp add: subst_dnf'_def)
lemma subst_dnf'_product:
"subst_dnf' (A \<otimes>\<^sub>l B) m = (subst_dnf' A m) \<otimes>\<^sub>l (subst_dnf' B m)" (is "?lhs = ?rhs")
proof (unfold set_eq_iff, safe)
fix x
assume "x \<in> ?lhs"
then obtain \<Phi> where "\<Phi> \<in> A \<otimes>\<^sub>l B" and "x \<in> subst_clause' \<Phi> m"
unfolding subst_dnf'_iff by blast
then obtain a b where "a \<in> A" and "b \<in> B" and "\<Phi> = a @ b"
unfolding list_product_def by blast
then have "x \<in> (subst_clause' a m) \<otimes>\<^sub>l (subst_clause' b m)"
using \<open>x \<in> subst_clause' \<Phi> m\<close> by simp
then obtain a' b' where "a' \<in> subst_clause' a m" and "b' \<in> subst_clause' b m" and "x = a' @ b'"
unfolding list_product_iff by blast
then have "a' \<in> subst_dnf' A m" and "b' \<in> subst_dnf' B m"
unfolding subst_dnf'_iff using \<open>a \<in> A\<close> \<open>b \<in> B\<close> by auto
then have "\<exists>a\<in>subst_dnf' A m. \<exists>b\<in>subst_dnf' B m. x = a @ b"
using \<open>x = a' @ b'\<close> by blast
then show "x \<in> ?rhs"
unfolding list_product_iff by blast
next
fix x
assume "x \<in> ?rhs"
then obtain a b where "a \<in> subst_dnf' A m" and "b \<in> subst_dnf' B m" and "x = a @ b"
unfolding list_product_iff by blast
then obtain a' b' where "a' \<in> A" and "b' \<in> B" and a: "a \<in> subst_clause' a' m" and b: "b \<in> subst_clause' b' m"
unfolding subst_dnf'_iff by blast
then have "x \<in> (subst_clause' a' m) \<otimes>\<^sub>l (subst_clause' b' m)"
unfolding list_product_iff using \<open>x = a @ b\<close> by blast
moreover
have "a' @ b' \<in> A \<otimes>\<^sub>l B"
unfolding list_product_iff using \<open>a' \<in> A\<close> \<open>b' \<in> B\<close> by blast
ultimately show "x \<in> ?lhs"
unfolding subst_dnf'_iff by force
qed
lemma subst_dnf'_list_dnf:
"subst_dnf' (list_dnf \<phi>) m = list_dnf (subst \<phi> m)"
proof (induction \<phi>)
case (And_ltln \<phi>1 \<phi>2)
then show ?case
by (simp add: subst_dnf'_product)
qed (simp_all add: subst_dnf'_def subst_clause'_def list_product_def)
lemma min_set_Union:
"finite X \<Longrightarrow> min_set (\<Union> (min_set ` X)) = min_set (\<Union> X)" for X :: "'a fset set set"
by (induction X rule: finite_induct) (force, metis Sup_insert image_insert min_set_min_union min_union_def)
lemma min_set_Union_image:
"finite X \<Longrightarrow> min_set (\<Union>x \<in> X. min_set (f x)) = min_set (\<Union>x \<in> X. f x)" for f :: "'b \<Rightarrow> 'a fset set"
proof -
assume "finite X"
then have *: "finite (f ` X)" by auto
with min_set_Union show ?thesis
unfolding image_image by fastforce
qed
lemma subst_clause_fset_of_list:
"subst_clause (fset_of_list \<Phi>) m = min_set (list_dnf_to_dnf (subst_clause' \<Phi> m))"
unfolding list_dnf_to_dnf_def subst_clause'_def
proof (induction \<Phi> rule: rev_induct)
case (snoc x xs)
then show ?case
by simp (metis (no_types, lifting) dnf_min_set list_dnf_to_dnf_def list_dnf_to_dnf_list_dnf min_product_comm min_product_def min_set_min_product(1))
qed simp
lemma min_set_list_dnf_to_dnf_subst_dnf':
"finite X \<Longrightarrow> min_set (list_dnf_to_dnf (subst_dnf' X m)) = min_set (subst_dnf (list_dnf_to_dnf X) m)"
by (simp add: subst_dnf'_def subst_dnf_def subst_clause_fset_of_list list_dnf_to_dnf_def min_set_Union_image image_Union)
lemma subst_dnf_dnf:
"min_set (subst_dnf (dnf \<phi>) m) = min_dnf (subst \<phi> m)"
unfolding dnf_min_set
unfolding list_dnf_to_dnf_list_dnf[symmetric]
unfolding subst_dnf'_list_dnf[symmetric]
unfolding min_set_list_dnf_to_dnf_subst_dnf'[OF list_dnf_finite]
by simp
text \<open>This is almost the lemma we need. However, we need to show that the same holds for @{term "min_dnf \<phi>"}, too.\<close>
lemma fold_product:
"Finite_Set.fold (\<lambda>x. (\<otimes>) {{|x|}}) {{||}} (fset x) = {x}"
by (induction x) (simp_all add: notin_fset, simp add: product_singleton_singleton)
lemma fold_union:
"Finite_Set.fold (\<lambda>x. (\<union>) {x}) {} (fset x) = fset x"
by (induction x) (simp_all add: notin_fset comp_fun_idem.fold_insert_idem comp_fun_idem_insert)
lemma fold_union_fold_product:
assumes "finite X" and "\<And>\<Psi> \<psi>. \<Psi> \<in> X \<Longrightarrow> \<psi> \<in> fset \<Psi> \<Longrightarrow> dnf \<psi> = {{|\<psi>|}}"
shows "Finite_Set.fold (\<lambda>x. (\<union>) (Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} (fset x))) {} X = X" (is "?lhs = X")
proof -
from assms have "?lhs = Finite_Set.fold (\<lambda>x. (\<union>) (Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) {{|\<phi>|}}) {{||}} (fset x))) {} X"
proof (induction X rule: finite_induct)
case (insert \<Phi> X)
from insert.prems have 1: "\<And>\<Psi> \<psi>. \<lbrakk>\<Psi> \<in> X; \<psi> \<in> fset \<Psi>\<rbrakk> \<Longrightarrow> dnf \<psi> = {{|\<psi>|}}"
by force
from insert.prems have "Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) (dnf \<phi>)) {{||}} (fset \<Phi>) = Finite_Set.fold (\<lambda>\<phi>. (\<otimes>) {{|\<phi>|}}) {{||}} (fset \<Phi>)"
by (induction \<Phi>) (force simp: notin_fset)+
with insert 1 show ?case
by simp
qed simp
with \<open>finite X\<close> show ?thesis
unfolding fold_product by (metis fset_to_fset fold_union)
qed
lemma dnf_ltln_of_dnf_min_dnf:
"dnf (ltln_of_dnf (min_dnf \<phi>)) = min_dnf \<phi>"
proof -
have 1: "finite (And\<^sub>n ` fset ` min_dnf \<phi>)"
using min_dnf_finite by blast
have 2: "inj_on And\<^sub>n (fset ` min_dnf \<phi>)"
by (metis (mono_tags, lifting) And\<^sub>n_inj f_inv_into_f fset inj_onI inj_on_contraD)
have 3: "inj_on fset (min_dnf \<phi>)"
by (meson fset_inject inj_onI)
show ?thesis
unfolding ltln_of_dnf_def
unfolding Or\<^sub>n_dnf[OF 1]
unfolding fold_image[OF 2]
unfolding fold_image[OF 3]
unfolding comp_def
unfolding And\<^sub>n_dnf[OF finite_fset]
by (metis fold_union_fold_product min_dnf_finite min_dnf_atoms_dnf)
qed
lemma min_dnf_subst:
"min_set (subst_dnf (min_dnf \<phi>) m) = min_dnf (subst \<phi> m)" (is "?lhs = ?rhs")
proof -
let ?\<phi>' = "ltln_of_dnf (min_dnf \<phi>)"
have "?lhs = min_set (subst_dnf (dnf ?\<phi>') m)"
unfolding dnf_ltln_of_dnf_min_dnf ..
also have "\<dots> = min_dnf (subst ?\<phi>' m)"
unfolding subst_dnf_dnf ..
also have "\<dots> = min_dnf (subst \<phi> m)"
using ltl_prop_equiv_min_dnf ltln_of_dnf_prop_equiv subst_respects_ltl_prop_entailment(2) by blast
finally show ?thesis .
qed
end
diff --git a/thys/LTL/LTL.thy b/thys/LTL/LTL.thy
--- a/thys/LTL/LTL.thy
+++ b/thys/LTL/LTL.thy
@@ -1,1189 +1,1189 @@
(*
Author: Salomon Sickert
Author: Benedikt Seidl
Author: Alexander Schimpf (original entry: CAVA/LTL.thy)
Author: Stephan Merz (original entry: Stuttering_Equivalence/PLTL.thy)
License: BSD
*)
section \<open>Linear Temporal Logic\<close>
theory LTL
imports
Main "HOL-Library.Omega_Words_Fun"
begin
text \<open>This theory provides a formalisation of linear temporal logic. It provides three variants:
\begin{enumerate}
\item LTL with syntactic sugar. This variant is the semantic reference and the included parser
generates ASTs of this datatype.
\item LTL in negation normal form without syntactic sugar. This variant is used by the included
rewriting engine and is used for the translation to automata (implemented in other entries).
\item LTL in restricted negation normal form without the rather uncommon operators ``weak until''
and ``strong release''. It is used by the formalization of Gerth's algorithm.
\item PLTL. A variant with a reduced set of operators.
\end{enumerate}
This theory subsumes (and partly reuses) the existing formalisation found in LTL\_to\_GBA and
Stuttering\_Equivalence and unifies them.\<close>
subsection \<open>LTL with Syntactic Sugar\<close>
text \<open>In this section, we provide a formulation of LTL with explicit syntactic sugar deeply embedded.
This formalization serves as a reference semantics.\<close>
subsubsection \<open>Syntax\<close>
datatype (atoms_ltlc: 'a) ltlc =
True_ltlc ("true\<^sub>c")
| False_ltlc ("false\<^sub>c")
| Prop_ltlc 'a ("prop\<^sub>c'(_')")
| Not_ltlc "'a ltlc" ("not\<^sub>c _" [85] 85)
| And_ltlc "'a ltlc" "'a ltlc" ("_ and\<^sub>c _" [82,82] 81)
| Or_ltlc "'a ltlc" "'a ltlc" ("_ or\<^sub>c _" [81,81] 80)
| Implies_ltlc "'a ltlc" "'a ltlc" ("_ implies\<^sub>c _" [81,81] 80)
| Next_ltlc "'a ltlc" ("X\<^sub>c _" [88] 87)
| Final_ltlc "'a ltlc" ("F\<^sub>c _" [88] 87)
| Global_ltlc "'a ltlc" ("G\<^sub>c _" [88] 87)
| Until_ltlc "'a ltlc" "'a ltlc" ("_ U\<^sub>c _" [84,84] 83)
| Release_ltlc "'a ltlc" "'a ltlc" ("_ R\<^sub>c _" [84,84] 83)
| WeakUntil_ltlc "'a ltlc" "'a ltlc" ("_ W\<^sub>c _" [84,84] 83)
| StrongRelease_ltlc "'a ltlc" "'a ltlc" ("_ M\<^sub>c _" [84,84] 83)
definition Iff_ltlc ("_ iff\<^sub>c _" [81,81] 80)
where
"\<phi> iff\<^sub>c \<psi> \<equiv> (\<phi> implies\<^sub>c \<psi>) and\<^sub>c (\<psi> implies\<^sub>c \<phi>)"
subsubsection \<open>Semantics\<close>
primrec semantics_ltlc :: "['a set word, 'a ltlc] \<Rightarrow> bool" ("_ \<Turnstile>\<^sub>c _" [80,80] 80)
where
"\<xi> \<Turnstile>\<^sub>c true\<^sub>c = True"
| "\<xi> \<Turnstile>\<^sub>c false\<^sub>c = False"
| "\<xi> \<Turnstile>\<^sub>c prop\<^sub>c(q) = (q\<in>\<xi> 0)"
| "\<xi> \<Turnstile>\<^sub>c not\<^sub>c \<phi> = (\<not> \<xi> \<Turnstile>\<^sub>c \<phi>)"
| "\<xi> \<Turnstile>\<^sub>c \<phi> and\<^sub>c \<psi> = (\<xi> \<Turnstile>\<^sub>c \<phi> \<and> \<xi> \<Turnstile>\<^sub>c \<psi>)"
| "\<xi> \<Turnstile>\<^sub>c \<phi> or\<^sub>c \<psi> = (\<xi> \<Turnstile>\<^sub>c \<phi> \<or> \<xi> \<Turnstile>\<^sub>c \<psi>)"
| "\<xi> \<Turnstile>\<^sub>c \<phi> implies\<^sub>c \<psi> = (\<xi> \<Turnstile>\<^sub>c \<phi> \<longrightarrow> \<xi> \<Turnstile>\<^sub>c \<psi>)"
| "\<xi> \<Turnstile>\<^sub>c X\<^sub>c \<phi> = (suffix 1 \<xi> \<Turnstile>\<^sub>c \<phi>)"
| "\<xi> \<Turnstile>\<^sub>c F\<^sub>c \<phi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>c \<phi>)"
| "\<xi> \<Turnstile>\<^sub>c G\<^sub>c \<phi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>c \<phi>)"
| "\<xi> \<Turnstile>\<^sub>c \<phi> U\<^sub>c \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>c \<psi> \<and> (\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>c \<phi>))"
| "\<xi> \<Turnstile>\<^sub>c \<phi> R\<^sub>c \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>c \<psi> \<or> (\<exists>j<i. suffix j \<xi> \<Turnstile>\<^sub>c \<phi>))"
| "\<xi> \<Turnstile>\<^sub>c \<phi> W\<^sub>c \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>c \<phi> \<or> (\<exists>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>c \<psi>))"
| "\<xi> \<Turnstile>\<^sub>c \<phi> M\<^sub>c \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>c \<phi> \<and> (\<forall>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>c \<psi>))"
lemma semantics_ltlc_sugar [simp]:
"\<xi> \<Turnstile>\<^sub>c \<phi> iff\<^sub>c \<psi> = (\<xi> \<Turnstile>\<^sub>c \<phi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c \<psi>)"
"\<xi> \<Turnstile>\<^sub>c F\<^sub>c \<phi> = \<xi> \<Turnstile>\<^sub>c (true\<^sub>c U\<^sub>c \<phi>)"
"\<xi> \<Turnstile>\<^sub>c G\<^sub>c \<phi> = \<xi> \<Turnstile>\<^sub>c (false\<^sub>c R\<^sub>c \<phi>)"
by (auto simp add: Iff_ltlc_def)
definition "language_ltlc \<phi> \<equiv> {\<xi>. \<xi> \<Turnstile>\<^sub>c \<phi>}"
lemma language_ltlc_negate[simp]:
"language_ltlc (not\<^sub>c \<phi>) = - language_ltlc \<phi>"
unfolding language_ltlc_def by auto
lemma ltl_true_or_con[simp]:
"\<xi> \<Turnstile>\<^sub>c prop\<^sub>c(p) or\<^sub>c (not\<^sub>c prop\<^sub>c(p))"
by auto
lemma ltl_false_true_con[simp]:
"\<xi> \<Turnstile>\<^sub>c not\<^sub>c true\<^sub>c \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c false\<^sub>c"
by auto
lemma ltl_Next_Neg_con[simp]:
"\<xi> \<Turnstile>\<^sub>c X\<^sub>c (not\<^sub>c \<phi>) \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c not\<^sub>c X\<^sub>c \<phi>"
by auto
\<comment> \<open>The connection between dual operators\<close>
lemma ltl_Until_Release_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> R\<^sub>c \<psi> \<longleftrightarrow> (\<not> \<xi> \<Turnstile>\<^sub>c (not\<^sub>c \<phi>) U\<^sub>c (not\<^sub>c \<psi>))"
"\<xi> \<Turnstile>\<^sub>c \<phi> U\<^sub>c \<psi> \<longleftrightarrow> (\<not> \<xi> \<Turnstile>\<^sub>c (not\<^sub>c \<phi>) R\<^sub>c (not\<^sub>c \<psi>))"
by auto
lemma ltl_WeakUntil_StrongRelease_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> W\<^sub>c \<psi> \<longleftrightarrow> (\<not> \<xi> \<Turnstile>\<^sub>c (not\<^sub>c \<phi>) M\<^sub>c (not\<^sub>c \<psi>))"
"\<xi> \<Turnstile>\<^sub>c \<phi> M\<^sub>c \<psi> \<longleftrightarrow> (\<not> \<xi> \<Turnstile>\<^sub>c (not\<^sub>c \<phi>) W\<^sub>c (not\<^sub>c \<psi>))"
by auto
\<comment> \<open>The connection between weak and strong operators\<close>
lemma ltl_Release_StrongRelease_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> R\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c (\<phi> M\<^sub>c \<psi>) or\<^sub>c (G\<^sub>c \<psi>)"
"\<xi> \<Turnstile>\<^sub>c \<phi> M\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c (\<phi> R\<^sub>c \<psi>) and\<^sub>c (F\<^sub>c \<phi>)"
proof safe
assume asm: "\<xi> \<Turnstile>\<^sub>c \<phi> R\<^sub>c \<psi>"
show "\<xi> \<Turnstile>\<^sub>c (\<phi> M\<^sub>c \<psi>) or\<^sub>c (G\<^sub>c \<psi>)"
proof (cases "\<xi> \<Turnstile>\<^sub>c G\<^sub>c \<psi>")
case False
then obtain i where "\<not> suffix i \<xi> \<Turnstile>\<^sub>c \<psi>" and "\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>c \<psi>"
using exists_least_iff[of "\<lambda>i. \<not> suffix i \<xi> \<Turnstile>\<^sub>c \<psi>"] by force
then show ?thesis
using asm by force
qed simp
next
assume asm: "\<xi> \<Turnstile>\<^sub>c (\<phi> R\<^sub>c \<psi>) and\<^sub>c (F\<^sub>c \<phi>)"
then show "\<xi> \<Turnstile>\<^sub>c \<phi> M\<^sub>c \<psi>"
proof (cases "\<xi> \<Turnstile>\<^sub>c F\<^sub>c \<phi>")
case True
then obtain i where "suffix i \<xi> \<Turnstile>\<^sub>c \<phi>" and "\<forall>j<i. \<not> suffix j \<xi> \<Turnstile>\<^sub>c \<phi>"
using exists_least_iff[of "\<lambda>i. suffix i \<xi> \<Turnstile>\<^sub>c \<phi>"] by force
then show ?thesis
using asm by force
qed simp
qed (unfold semantics_ltlc.simps; insert not_less, blast)+
lemma ltl_Until_WeakUntil_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> U\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c (\<phi> W\<^sub>c \<psi>) and\<^sub>c (F\<^sub>c \<psi>)"
"\<xi> \<Turnstile>\<^sub>c \<phi> W\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c (\<phi> U\<^sub>c \<psi>) or\<^sub>c (G\<^sub>c \<phi>)"
proof safe
assume asm: "\<xi> \<Turnstile>\<^sub>c (\<phi> W\<^sub>c \<psi>) and\<^sub>c (F\<^sub>c \<psi>)"
then show "\<xi> \<Turnstile>\<^sub>c \<phi> U\<^sub>c \<psi>"
proof (cases "\<xi> \<Turnstile>\<^sub>c F\<^sub>c \<psi>")
case True
then obtain i where "suffix i \<xi> \<Turnstile>\<^sub>c \<psi>" and "\<forall>j<i. \<not> suffix j \<xi> \<Turnstile>\<^sub>c \<psi>"
using exists_least_iff[of "\<lambda>i. suffix i \<xi> \<Turnstile>\<^sub>c \<psi>"] by force
then show ?thesis
using asm by force
qed simp
next
assume asm: "\<xi> \<Turnstile>\<^sub>c \<phi> W\<^sub>c \<psi>"
then show "\<xi> \<Turnstile>\<^sub>c (\<phi> U\<^sub>c \<psi>) or\<^sub>c (G\<^sub>c \<phi>)"
proof (cases "\<xi> \<Turnstile>\<^sub>c G\<^sub>c \<phi>")
case False
then obtain i where "\<not> suffix i \<xi> \<Turnstile>\<^sub>c \<phi>" and "\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>c \<phi>"
using exists_least_iff[of "\<lambda>i. \<not> suffix i \<xi> \<Turnstile>\<^sub>c \<phi>"] by force
then show ?thesis
using asm by force
qed simp
qed (unfold semantics_ltlc.simps; insert not_less, blast)+
lemma ltl_StrongRelease_Until_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> M\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c \<psi> U\<^sub>c (\<phi> and\<^sub>c \<psi>)"
using order.order_iff_strict by auto
lemma ltl_WeakUntil_Release_con:
"\<xi> \<Turnstile>\<^sub>c \<phi> R\<^sub>c \<psi> \<longleftrightarrow> \<xi> \<Turnstile>\<^sub>c \<psi> W\<^sub>c (\<phi> and\<^sub>c \<psi>)"
by (meson ltl_Release_StrongRelease_con(1) ltl_StrongRelease_Until_con ltl_Until_WeakUntil_con(2) semantics_ltlc.simps(6))
definition "pw_eq_on S w w' \<equiv> \<forall>i. w i \<inter> S = w' i \<inter> S"
lemma pw_eq_on_refl[simp]: "pw_eq_on S w w"
and pw_eq_on_sym: "pw_eq_on S w w' \<Longrightarrow> pw_eq_on S w' w"
and pw_eq_on_trans[trans]: "\<lbrakk>pw_eq_on S w w'; pw_eq_on S w' w''\<rbrakk> \<Longrightarrow> pw_eq_on S w w''"
unfolding pw_eq_on_def by auto
lemma pw_eq_on_suffix:
"pw_eq_on S w w' \<Longrightarrow> pw_eq_on S (suffix k w) (suffix k w')"
by (simp add: pw_eq_on_def)
lemma pw_eq_on_subset:
"S \<subseteq> S' \<Longrightarrow> pw_eq_on S' w w' \<Longrightarrow> pw_eq_on S w w'"
by (auto simp add: pw_eq_on_def)
lemma ltlc_eq_on_aux:
"pw_eq_on (atoms_ltlc \<phi>) w w' \<Longrightarrow> w \<Turnstile>\<^sub>c \<phi> \<Longrightarrow> w' \<Turnstile>\<^sub>c \<phi>"
proof (induction \<phi> arbitrary: w w')
case Until_ltlc
thus ?case
by simp (meson Un_upper1 Un_upper2 pw_eq_on_subset pw_eq_on_suffix)
next
case Release_ltlc
thus ?case
by simp (metis Un_upper1 pw_eq_on_subset pw_eq_on_suffix sup_commute)
next
case WeakUntil_ltlc
thus ?case
by simp (meson pw_eq_on_subset pw_eq_on_suffix sup.cobounded1 sup_ge2)
next
case StrongRelease_ltlc
thus ?case
by simp (metis Un_upper1 pw_eq_on_subset pw_eq_on_suffix pw_eq_on_sym sup_ge2)
next
case (And_ltlc \<phi> \<psi>)
thus ?case
by simp (meson Un_upper1 inf_sup_ord(4) pw_eq_on_subset)
next
case (Or_ltlc \<phi> \<psi>)
thus ?case
by simp (meson Un_upper2 pw_eq_on_subset sup_ge1)
next
case (Implies_ltlc \<phi> \<psi>)
thus ?case
by simp (meson Un_upper1 Un_upper2 pw_eq_on_subset[of "atoms_ltlc _" "atoms_ltlc \<phi> \<union> atoms_ltlc \<psi>"] pw_eq_on_sym)
qed (auto simp add: pw_eq_on_def; metis suffix_nth)+
lemma ltlc_eq_on:
"pw_eq_on (atoms_ltlc \<phi>) w w' \<Longrightarrow> w \<Turnstile>\<^sub>c \<phi> \<longleftrightarrow> w' \<Turnstile>\<^sub>c \<phi>"
using ltlc_eq_on_aux pw_eq_on_sym by blast
lemma suffix_comp: "(\<lambda>i. f (suffix k w i)) = suffix k (f o w)"
by auto
lemma suffix_range: "\<Union>(range \<xi>) \<subseteq> APs \<Longrightarrow> \<Union>(range (suffix k \<xi>)) \<subseteq> APs"
by auto
lemma map_semantics_ltlc_aux:
assumes "inj_on f APs"
assumes "\<Union>(range w) \<subseteq> APs"
assumes "atoms_ltlc \<phi> \<subseteq> APs"
shows "w \<Turnstile>\<^sub>c \<phi> \<longleftrightarrow> (\<lambda>i. f ` w i) \<Turnstile>\<^sub>c map_ltlc f \<phi>"
using assms(2,3)
proof (induction \<phi> arbitrary: w)
case (Prop_ltlc x)
thus ?case using assms(1)
by (simp add: SUP_le_iff inj_on_image_mem_iff)
next
case (Next_ltlc \<phi>)
show ?case
using Next_ltlc(1)[of "suffix 1 w", unfolded suffix_comp comp_def] Next_ltlc(2,3) apply simp
by (metis Next_ltlc.prems(1) One_nat_def \<open>\<lbrakk>\<Union>(range (suffix 1 w)) \<subseteq> APs; atoms_ltlc \<phi> \<subseteq> APs\<rbrakk> \<Longrightarrow> suffix 1 w \<Turnstile>\<^sub>c \<phi> = suffix 1 (\<lambda>x. f ` w x) \<Turnstile>\<^sub>c map_ltlc f \<phi>\<close> suffix_range)
next
case (Final_ltlc \<phi>)
thus ?case
using Final_ltlc(1)[of "suffix _ _", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
next
case (Global_ltlc)
thus ?case
using Global_ltlc(1)[of "suffix _ w", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
next
case (Until_ltlc)
thus ?case
using Until_ltlc(1,2)[of "suffix _ w", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
next
case (Release_ltlc)
thus ?case
using Release_ltlc(1,2)[of "suffix _ w", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
next
case (WeakUntil_ltlc)
thus ?case
using WeakUntil_ltlc(1,2)[of "suffix _ w", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
next
case (StrongRelease_ltlc)
thus ?case
using StrongRelease_ltlc(1,2)[of "suffix _ w", unfolded suffix_comp comp_def, OF suffix_range] by fastforce
qed simp+
definition "map_props f APs \<equiv> {i. \<exists>p\<in>APs. f p = Some i}"
lemma map_semantics_ltlc:
assumes INJ: "inj_on f (dom f)" and DOM: "atoms_ltlc \<phi> \<subseteq> dom f"
shows "\<xi> \<Turnstile>\<^sub>c \<phi> \<longleftrightarrow> (map_props f o \<xi>) \<Turnstile>\<^sub>c map_ltlc (the o f) \<phi>"
proof -
let ?\<xi>r = "\<lambda>i. \<xi> i \<inter> atoms_ltlc \<phi>"
let ?\<xi>r' = "\<lambda>i. \<xi> i \<inter> dom f"
have 1: "\<Union>(range ?\<xi>r) \<subseteq> atoms_ltlc \<phi>" by auto
have INJ_the_dom: "inj_on (the o f) (dom f)"
using assms
by (auto simp: inj_on_def domIff)
note 2 = subset_inj_on[OF this DOM]
have 3: "(\<lambda>i. (the o f) ` ?\<xi>r' i) = map_props f o \<xi>" using DOM INJ
apply (auto intro!: ext simp: map_props_def domIff image_iff)
by (metis Int_iff domI option.sel)
have "\<xi> \<Turnstile>\<^sub>c \<phi> \<longleftrightarrow> ?\<xi>r \<Turnstile>\<^sub>c \<phi>"
apply (rule ltlc_eq_on)
apply (auto simp: pw_eq_on_def)
done
also from map_semantics_ltlc_aux[OF 2 1 subset_refl]
have "\<dots> \<longleftrightarrow> (\<lambda>i. (the o f) ` ?\<xi>r i) \<Turnstile>\<^sub>c map_ltlc (the o f) \<phi>" .
also have "\<dots> \<longleftrightarrow> (\<lambda>i. (the o f) ` ?\<xi>r' i) \<Turnstile>\<^sub>c map_ltlc (the o f) \<phi>"
apply (rule ltlc_eq_on) using DOM INJ
apply (auto simp: pw_eq_on_def ltlc.set_map domIff image_iff)
by (metis Int_iff contra_subsetD domD domI inj_on_eq_iff option.sel)
also note 3
finally show ?thesis .
qed
lemma map_semantics_ltlc_inv:
assumes INJ: "inj_on f (dom f)" and DOM: "atoms_ltlc \<phi> \<subseteq> dom f"
shows "\<xi> \<Turnstile>\<^sub>c map_ltlc (the o f) \<phi> \<longleftrightarrow> (\<lambda>i. (the o f) -` \<xi> i) \<Turnstile>\<^sub>c \<phi>"
using map_semantics_ltlc[OF assms]
apply simp
apply (intro ltlc_eq_on)
apply (auto simp add: pw_eq_on_def ltlc.set_map map_props_def)
by (metis DOM comp_apply contra_subsetD domD option.sel vimage_eq)
subsection \<open>LTL in Negation Normal Form\<close>
text \<open>We define a type of LTL formula in negation normal form (NNF).\<close>
subsubsection \<open>Syntax\<close>
datatype (atoms_ltln: 'a) ltln =
True_ltln ("true\<^sub>n")
| False_ltln ("false\<^sub>n")
| Prop_ltln 'a ("prop\<^sub>n'(_')")
| Nprop_ltln 'a ("nprop\<^sub>n'(_')")
| And_ltln "'a ltln" "'a ltln" ("_ and\<^sub>n _" [82,82] 81)
| Or_ltln "'a ltln" "'a ltln" ("_ or\<^sub>n _" [84,84] 83)
| Next_ltln "'a ltln" ("X\<^sub>n _" [88] 87)
| Until_ltln "'a ltln" "'a ltln" ("_ U\<^sub>n _" [84,84] 83)
| Release_ltln "'a ltln" "'a ltln" ("_ R\<^sub>n _" [84,84] 83)
| WeakUntil_ltln "'a ltln" "'a ltln" ("_ W\<^sub>n _" [84,84] 83)
| StrongRelease_ltln "'a ltln" "'a ltln" ("_ M\<^sub>n _" [84,84] 83)
abbreviation finally\<^sub>n :: "'a ltln \<Rightarrow> 'a ltln" ("F\<^sub>n _" [88] 87)
where
"F\<^sub>n \<phi> \<equiv> true\<^sub>n U\<^sub>n \<phi>"
notation (input) finally\<^sub>n ("\<diamondsuit>\<^sub>n _" [88] 87)
abbreviation globally\<^sub>n :: "'a ltln \<Rightarrow> 'a ltln" ("G\<^sub>n _" [88] 87)
where
"G\<^sub>n \<phi> \<equiv> false\<^sub>n R\<^sub>n \<phi>"
notation (input) globally\<^sub>n ("\<box>\<^sub>n _" [88] 87)
subsubsection \<open>Semantics\<close>
primrec semantics_ltln :: "['a set word, 'a ltln] \<Rightarrow> bool" ("_ \<Turnstile>\<^sub>n _" [80,80] 80)
where
"\<xi> \<Turnstile>\<^sub>n true\<^sub>n = True"
| "\<xi> \<Turnstile>\<^sub>n false\<^sub>n = False"
| "\<xi> \<Turnstile>\<^sub>n prop\<^sub>n(q) = (q \<in> \<xi> 0)"
| "\<xi> \<Turnstile>\<^sub>n nprop\<^sub>n(q) = (q \<notin> \<xi> 0)"
| "\<xi> \<Turnstile>\<^sub>n \<phi> and\<^sub>n \<psi> = (\<xi> \<Turnstile>\<^sub>n \<phi> \<and> \<xi> \<Turnstile>\<^sub>n \<psi>)"
| "\<xi> \<Turnstile>\<^sub>n \<phi> or\<^sub>n \<psi> = (\<xi> \<Turnstile>\<^sub>n \<phi> \<or> \<xi> \<Turnstile>\<^sub>n \<psi>)"
| "\<xi> \<Turnstile>\<^sub>n X\<^sub>n \<phi> = (suffix 1 \<xi> \<Turnstile>\<^sub>n \<phi>)"
| "\<xi> \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>n \<psi> \<and> (\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>n \<phi>))"
| "\<xi> \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>n \<psi> \<or> (\<exists>j<i. suffix j \<xi> \<Turnstile>\<^sub>n \<phi>))"
| "\<xi> \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>n \<phi> \<or> (\<exists>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>n \<psi>))"
| "\<xi> \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>n \<phi> \<and> (\<forall>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>n \<psi>))"
definition "language_ltln \<phi> \<equiv> {\<xi>. \<xi> \<Turnstile>\<^sub>n \<phi>}"
lemma semantics_ltln_ite_simps[simp]:
"w \<Turnstile>\<^sub>n (if P then true\<^sub>n else false\<^sub>n) = P"
"w \<Turnstile>\<^sub>n (if P then false\<^sub>n else true\<^sub>n) = (\<not>P)"
by simp_all
subsubsection \<open>Conversion\<close>
fun ltlc_to_ltln' :: "bool \<Rightarrow> 'a ltlc \<Rightarrow> 'a ltln"
where
"ltlc_to_ltln' False true\<^sub>c = true\<^sub>n"
| "ltlc_to_ltln' False false\<^sub>c = false\<^sub>n"
| "ltlc_to_ltln' False prop\<^sub>c(q) = prop\<^sub>n(q)"
| "ltlc_to_ltln' False (\<phi> and\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) and\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (\<phi> or\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) or\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (\<phi> implies\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) or\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (F\<^sub>c \<phi>) = true\<^sub>n U\<^sub>n (ltlc_to_ltln' False \<phi>)"
| "ltlc_to_ltln' False (G\<^sub>c \<phi>) = false\<^sub>n R\<^sub>n (ltlc_to_ltln' False \<phi>)"
| "ltlc_to_ltln' False (\<phi> U\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) U\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (\<phi> R\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) R\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (\<phi> W\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) W\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' False (\<phi> M\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) M\<^sub>n (ltlc_to_ltln' False \<psi>)"
| "ltlc_to_ltln' True true\<^sub>c = false\<^sub>n"
| "ltlc_to_ltln' True false\<^sub>c = true\<^sub>n"
| "ltlc_to_ltln' True prop\<^sub>c(q) = nprop\<^sub>n(q)"
| "ltlc_to_ltln' True (\<phi> and\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) or\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (\<phi> or\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) and\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (\<phi> implies\<^sub>c \<psi>) = (ltlc_to_ltln' False \<phi>) and\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (F\<^sub>c \<phi>) = false\<^sub>n R\<^sub>n (ltlc_to_ltln' True \<phi>)"
| "ltlc_to_ltln' True (G\<^sub>c \<phi>) = true\<^sub>n U\<^sub>n (ltlc_to_ltln' True \<phi>)"
| "ltlc_to_ltln' True (\<phi> U\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) R\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (\<phi> R\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) U\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (\<phi> W\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) M\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' True (\<phi> M\<^sub>c \<psi>) = (ltlc_to_ltln' True \<phi>) W\<^sub>n (ltlc_to_ltln' True \<psi>)"
| "ltlc_to_ltln' b (not\<^sub>c \<phi>) = ltlc_to_ltln' (\<not> b) \<phi>"
| "ltlc_to_ltln' b (X\<^sub>c \<phi>) = X\<^sub>n (ltlc_to_ltln' b \<phi>)"
fun ltlc_to_ltln :: "'a ltlc \<Rightarrow> 'a ltln"
where
"ltlc_to_ltln \<phi> = ltlc_to_ltln' False \<phi>"
fun ltln_to_ltlc :: "'a ltln \<Rightarrow> 'a ltlc"
where
"ltln_to_ltlc true\<^sub>n = true\<^sub>c"
| "ltln_to_ltlc false\<^sub>n = false\<^sub>c"
| "ltln_to_ltlc prop\<^sub>n(q) = prop\<^sub>c(q)"
| "ltln_to_ltlc nprop\<^sub>n(q) = not\<^sub>c (prop\<^sub>c(q))"
| "ltln_to_ltlc (\<phi> and\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> and\<^sub>c ltln_to_ltlc \<psi>)"
| "ltln_to_ltlc (\<phi> or\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> or\<^sub>c ltln_to_ltlc \<psi>)"
| "ltln_to_ltlc (X\<^sub>n \<phi>) = (X\<^sub>c ltln_to_ltlc \<phi>)"
| "ltln_to_ltlc (\<phi> U\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> U\<^sub>c ltln_to_ltlc \<psi>)"
| "ltln_to_ltlc (\<phi> R\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> R\<^sub>c ltln_to_ltlc \<psi>)"
| "ltln_to_ltlc (\<phi> W\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> W\<^sub>c ltln_to_ltlc \<psi>)"
| "ltln_to_ltlc (\<phi> M\<^sub>n \<psi>) = (ltln_to_ltlc \<phi> M\<^sub>c ltln_to_ltlc \<psi>)"
lemma ltlc_to_ltln'_correct:
"w \<Turnstile>\<^sub>n (ltlc_to_ltln' True \<phi>) \<longleftrightarrow> \<not> w \<Turnstile>\<^sub>c \<phi>"
"w \<Turnstile>\<^sub>n (ltlc_to_ltln' False \<phi>) \<longleftrightarrow> w \<Turnstile>\<^sub>c \<phi>"
"size (ltlc_to_ltln' True \<phi>) \<le> 2 * size \<phi>"
"size (ltlc_to_ltln' False \<phi>) \<le> 2 * size \<phi>"
by (induction \<phi> arbitrary: w) simp+
lemma ltlc_to_ltln_semantics [simp]:
"w \<Turnstile>\<^sub>n ltlc_to_ltln \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>c \<phi>"
using ltlc_to_ltln'_correct by auto
lemma ltlc_to_ltln_size:
"size (ltlc_to_ltln \<phi>) \<le> 2 * size \<phi>"
using ltlc_to_ltln'_correct by simp
lemma ltln_to_ltlc_semantics [simp]:
"w \<Turnstile>\<^sub>c ltln_to_ltlc \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>"
by (induction \<phi> arbitrary: w) simp+
lemma ltlc_to_ltln_atoms:
"atoms_ltln (ltlc_to_ltln \<phi>) = atoms_ltlc \<phi>"
proof -
have "atoms_ltln (ltlc_to_ltln' True \<phi>) = atoms_ltlc \<phi>"
"atoms_ltln (ltlc_to_ltln' False \<phi>) = atoms_ltlc \<phi>"
by (induction \<phi>) simp+
thus ?thesis
by simp
qed
subsubsection \<open>Negation\<close>
fun not\<^sub>n
where
"not\<^sub>n true\<^sub>n = false\<^sub>n"
| "not\<^sub>n false\<^sub>n = true\<^sub>n"
| "not\<^sub>n prop\<^sub>n(a) = nprop\<^sub>n(a)"
| "not\<^sub>n nprop\<^sub>n(a) = prop\<^sub>n(a)"
| "not\<^sub>n (\<phi> and\<^sub>n \<psi>) = (not\<^sub>n \<phi>) or\<^sub>n (not\<^sub>n \<psi>)"
| "not\<^sub>n (\<phi> or\<^sub>n \<psi>) = (not\<^sub>n \<phi>) and\<^sub>n (not\<^sub>n \<psi>)"
| "not\<^sub>n (X\<^sub>n \<phi>) = X\<^sub>n (not\<^sub>n \<phi>)"
| "not\<^sub>n (\<phi> U\<^sub>n \<psi>) = (not\<^sub>n \<phi>) R\<^sub>n (not\<^sub>n \<psi>)"
| "not\<^sub>n (\<phi> R\<^sub>n \<psi>) = (not\<^sub>n \<phi>) U\<^sub>n (not\<^sub>n \<psi>)"
| "not\<^sub>n (\<phi> W\<^sub>n \<psi>) = (not\<^sub>n \<phi>) M\<^sub>n (not\<^sub>n \<psi>)"
| "not\<^sub>n (\<phi> M\<^sub>n \<psi>) = (not\<^sub>n \<phi>) W\<^sub>n (not\<^sub>n \<psi>)"
lemma not\<^sub>n_semantics[simp]:
"w \<Turnstile>\<^sub>n not\<^sub>n \<phi> \<longleftrightarrow> \<not> w \<Turnstile>\<^sub>n \<phi>"
by (induction \<phi> arbitrary: w) auto
lemma not\<^sub>n_size:
"size (not\<^sub>n \<phi>) = size \<phi>"
by (induction \<phi>) auto
subsubsection \<open>Subformulas\<close>
fun subfrmlsn :: "'a ltln \<Rightarrow> 'a ltln set"
where
"subfrmlsn (\<phi> and\<^sub>n \<psi>) = {\<phi> and\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn (\<phi> or\<^sub>n \<psi>) = {\<phi> or\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn (X\<^sub>n \<phi>) = {X\<^sub>n \<phi>} \<union> subfrmlsn \<phi>"
| "subfrmlsn (\<phi> U\<^sub>n \<psi>) = {\<phi> U\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn (\<phi> R\<^sub>n \<psi>) = {\<phi> R\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn (\<phi> W\<^sub>n \<psi>) = {\<phi> W\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn (\<phi> M\<^sub>n \<psi>) = {\<phi> M\<^sub>n \<psi>} \<union> subfrmlsn \<phi> \<union> subfrmlsn \<psi>"
| "subfrmlsn \<phi> = {\<phi>}"
lemma subfrmlsn_id[simp]:
"\<phi> \<in> subfrmlsn \<phi>"
by (induction \<phi>) auto
lemma subfrmlsn_finite:
"finite (subfrmlsn \<phi>)"
by (induction \<phi>) auto
lemma subfrmlsn_card:
"card (subfrmlsn \<phi>) \<le> size \<phi>"
by (induction \<phi>) (simp_all add: card_insert_if subfrmlsn_finite, (meson add_le_mono card_Un_le dual_order.trans le_SucI)+)
lemma subfrmlsn_subset:
"\<psi> \<in> subfrmlsn \<phi> \<Longrightarrow> subfrmlsn \<psi> \<subseteq> subfrmlsn \<phi>"
by (induction \<phi>) auto
lemma subfrmlsn_size:
"\<psi> \<in> subfrmlsn \<phi> \<Longrightarrow> size \<psi> < size \<phi> \<or> \<psi> = \<phi>"
by (induction \<phi>) auto
abbreviation
"size_set S \<equiv> sum (\<lambda>x. 2 * size x + 1) S"
lemma size_set_diff:
"finite S \<Longrightarrow> S' \<subseteq> S \<Longrightarrow> size_set (S - S') = size_set S - size_set S'"
using sum_diff_nat finite_subset by metis
subsubsection \<open>Constant Folding\<close>
lemma U_consts[intro, simp]:
"w \<Turnstile>\<^sub>n \<phi> U\<^sub>n true\<^sub>n"
"\<not> (w \<Turnstile>\<^sub>n \<phi> U\<^sub>n false\<^sub>n)"
"(w \<Turnstile>\<^sub>n false\<^sub>n U\<^sub>n \<phi>) = (w \<Turnstile>\<^sub>n \<phi>)"
by force+
lemma R_consts[intro, simp]:
"w \<Turnstile>\<^sub>n \<phi> R\<^sub>n true\<^sub>n"
"\<not> (w \<Turnstile>\<^sub>n \<phi> R\<^sub>n false\<^sub>n)"
"(w \<Turnstile>\<^sub>n true\<^sub>n R\<^sub>n \<phi>) = (w \<Turnstile>\<^sub>n \<phi>)"
by force+
lemma W_consts[intro, simp]:
"w \<Turnstile>\<^sub>n true\<^sub>n W\<^sub>n \<phi>"
"w \<Turnstile>\<^sub>n \<phi> W\<^sub>n true\<^sub>n"
"(w \<Turnstile>\<^sub>n false\<^sub>n W\<^sub>n \<phi>) = (w \<Turnstile>\<^sub>n \<phi>)"
"(w \<Turnstile>\<^sub>n \<phi> W\<^sub>n false\<^sub>n) = (w \<Turnstile>\<^sub>n G\<^sub>n \<phi>)"
by force+
lemma M_consts[intro, simp]:
"\<not> (w \<Turnstile>\<^sub>n false\<^sub>n M\<^sub>n \<phi>)"
"\<not> (w \<Turnstile>\<^sub>n \<phi> M\<^sub>n false\<^sub>n)"
"(w \<Turnstile>\<^sub>n true\<^sub>n M\<^sub>n \<phi>) = (w \<Turnstile>\<^sub>n \<phi>)"
"(w \<Turnstile>\<^sub>n \<phi> M\<^sub>n true\<^sub>n) = (w \<Turnstile>\<^sub>n F\<^sub>n \<phi>)"
by force+
subsubsection \<open>Distributivity\<close>
lemma until_and_left_distrib:
"w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 and\<^sub>n \<phi>\<^sub>2) U\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 U\<^sub>n \<psi>) and\<^sub>n (\<phi>\<^sub>2 U\<^sub>n \<psi>)"
proof
assume "w \<Turnstile>\<^sub>n \<phi>\<^sub>1 U\<^sub>n \<psi> and\<^sub>n \<phi>\<^sub>2 U\<^sub>n \<psi>"
then obtain i1 i2 where "suffix i1 w \<Turnstile>\<^sub>n \<psi> \<and> (\<forall>j<i1. suffix j w \<Turnstile>\<^sub>n \<phi>\<^sub>1)" and "suffix i2 w \<Turnstile>\<^sub>n \<psi> \<and> (\<forall>j<i2. suffix j w \<Turnstile>\<^sub>n \<phi>\<^sub>2)"
by auto
then have "suffix (min i1 i2) w \<Turnstile>\<^sub>n \<psi> \<and> (\<forall>j<min i1 i2. suffix j w \<Turnstile>\<^sub>n \<phi>\<^sub>1 and\<^sub>n \<phi>\<^sub>2)"
by (simp add: min_def)
then show "w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 and\<^sub>n \<phi>\<^sub>2) U\<^sub>n \<psi>"
by force
qed auto
lemma until_or_right_distrib:
"w \<Turnstile>\<^sub>n \<phi> U\<^sub>n (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2) \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi> U\<^sub>n \<psi>\<^sub>1) or\<^sub>n (\<phi> U\<^sub>n \<psi>\<^sub>2)"
by auto
lemma release_and_right_distrib:
"w \<Turnstile>\<^sub>n \<phi> R\<^sub>n (\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2) \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi> R\<^sub>n \<psi>\<^sub>1) and\<^sub>n (\<phi> R\<^sub>n \<psi>\<^sub>2)"
by auto
lemma release_or_left_distrib:
"w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 or\<^sub>n \<phi>\<^sub>2) R\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 R\<^sub>n \<psi>) or\<^sub>n (\<phi>\<^sub>2 R\<^sub>n \<psi>)"
by (metis not\<^sub>n.simps(6) not\<^sub>n.simps(9) not\<^sub>n_semantics until_and_left_distrib)
lemma strong_release_and_right_distrib:
"w \<Turnstile>\<^sub>n \<phi> M\<^sub>n (\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2) \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi> M\<^sub>n \<psi>\<^sub>1) and\<^sub>n (\<phi> M\<^sub>n \<psi>\<^sub>2)"
proof
assume "w \<Turnstile>\<^sub>n (\<phi> M\<^sub>n \<psi>\<^sub>1) and\<^sub>n (\<phi> M\<^sub>n \<psi>\<^sub>2)"
then obtain i1 i2 where "suffix i1 w \<Turnstile>\<^sub>n \<phi> \<and> (\<forall>j\<le>i1. suffix j w \<Turnstile>\<^sub>n \<psi>\<^sub>1)" and "suffix i2 w \<Turnstile>\<^sub>n \<phi> \<and> (\<forall>j\<le>i2. suffix j w \<Turnstile>\<^sub>n \<psi>\<^sub>2)"
by auto
then have "suffix (min i1 i2) w \<Turnstile>\<^sub>n \<phi> \<and> (\<forall>j\<le>min i1 i2. suffix j w \<Turnstile>\<^sub>n \<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2)"
by (simp add: min_def)
then show "w \<Turnstile>\<^sub>n \<phi> M\<^sub>n (\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2)"
by force
qed auto
lemma strong_release_or_left_distrib:
"w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 or\<^sub>n \<phi>\<^sub>2) M\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 M\<^sub>n \<psi>) or\<^sub>n (\<phi>\<^sub>2 M\<^sub>n \<psi>)"
by auto
lemma weak_until_and_left_distrib:
"w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 and\<^sub>n \<phi>\<^sub>2) W\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi>\<^sub>1 W\<^sub>n \<psi>) and\<^sub>n (\<phi>\<^sub>2 W\<^sub>n \<psi>)"
by auto
lemma weak_until_or_right_distrib:
"w \<Turnstile>\<^sub>n \<phi> W\<^sub>n (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2) \<longleftrightarrow> w \<Turnstile>\<^sub>n (\<phi> W\<^sub>n \<psi>\<^sub>1) or\<^sub>n (\<phi> W\<^sub>n \<psi>\<^sub>2)"
by (metis not\<^sub>n.simps(10) not\<^sub>n.simps(6) not\<^sub>n_semantics strong_release_and_right_distrib)
lemma next_until_distrib:
"w \<Turnstile>\<^sub>n X\<^sub>n (\<phi> U\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n (X\<^sub>n \<phi>) U\<^sub>n (X\<^sub>n \<psi>)"
by auto
lemma next_release_distrib:
"w \<Turnstile>\<^sub>n X\<^sub>n (\<phi> R\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n (X\<^sub>n \<phi>) R\<^sub>n (X\<^sub>n \<psi>)"
by auto
lemma next_weak_until_distrib:
"w \<Turnstile>\<^sub>n X\<^sub>n (\<phi> W\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n (X\<^sub>n \<phi>) W\<^sub>n (X\<^sub>n \<psi>)"
by auto
lemma next_strong_release_distrib:
"w \<Turnstile>\<^sub>n X\<^sub>n (\<phi> M\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n (X\<^sub>n \<phi>) M\<^sub>n (X\<^sub>n \<psi>)"
by auto
subsubsection \<open>Nested operators\<close>
lemma finally_until[simp]:
"w \<Turnstile>\<^sub>n F\<^sub>n (\<phi> U\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n F\<^sub>n \<psi>"
by auto force
lemma globally_release[simp]:
"w \<Turnstile>\<^sub>n G\<^sub>n (\<phi> R\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
by auto force
lemma globally_weak_until[simp]:
"w \<Turnstile>\<^sub>n G\<^sub>n (\<phi> W\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n G\<^sub>n (\<phi> or\<^sub>n \<psi>)"
by auto force
lemma finally_strong_release[simp]:
"w \<Turnstile>\<^sub>n F\<^sub>n (\<phi> M\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n F\<^sub>n (\<phi> and\<^sub>n \<psi>)"
by auto force
subsubsection \<open>Weak and strong operators\<close>
lemma ltln_weak_strong:
"w \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (G\<^sub>n \<phi>) or\<^sub>n (\<phi> U\<^sub>n \<psi>)"
"w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (G\<^sub>n \<psi>) or\<^sub>n (\<phi> M\<^sub>n \<psi>)"
proof auto
fix i
assume "\<forall>i. suffix i w \<Turnstile>\<^sub>n \<phi> \<or> (\<exists>j\<le>i. suffix j w \<Turnstile>\<^sub>n \<psi>)"
and "\<forall>i. suffix i w \<Turnstile>\<^sub>n \<psi> \<longrightarrow> (\<exists>j<i. \<not> suffix j w \<Turnstile>\<^sub>n \<phi>)"
then show "suffix i w \<Turnstile>\<^sub>n \<phi>"
by (induction i rule: less_induct) force
next
fix i k
assume "\<forall>j\<le>i. \<not> suffix j w \<Turnstile>\<^sub>n \<psi>"
and "suffix k w \<Turnstile>\<^sub>n \<psi>"
and "\<forall>j<k. suffix j w \<Turnstile>\<^sub>n \<phi>"
then show "suffix i w \<Turnstile>\<^sub>n \<phi>"
by (cases "i < k") simp_all
next
fix i
assume "\<forall>i. suffix i w \<Turnstile>\<^sub>n \<psi> \<or> (\<exists>j<i. suffix j w \<Turnstile>\<^sub>n \<phi>)"
and "\<forall>i. suffix i w \<Turnstile>\<^sub>n \<phi> \<longrightarrow> (\<exists>j\<le>i. \<not> suffix j w \<Turnstile>\<^sub>n \<psi>)"
then show "suffix i w \<Turnstile>\<^sub>n \<psi>"
by (induction i rule: less_induct) force
next
fix i k
assume "\<forall>j<i. \<not> suffix j w \<Turnstile>\<^sub>n \<phi>"
and "suffix k w \<Turnstile>\<^sub>n \<phi>"
and "\<forall>j\<le>k. suffix j w \<Turnstile>\<^sub>n \<psi>"
then show "suffix i w \<Turnstile>\<^sub>n \<psi>"
by (cases "i \<le> k") simp_all
qed
lemma ltln_strong_weak:
"w \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (F\<^sub>n \<psi>) and\<^sub>n (\<phi> W\<^sub>n \<psi>)"
"w \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (F\<^sub>n \<phi>) and\<^sub>n (\<phi> R\<^sub>n \<psi>)"
by (metis ltln_weak_strong not\<^sub>n.simps(1,5,8-11) not\<^sub>n_semantics)+
lemma ltln_strong_to_weak:
"w \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi>"
"w \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi>"
using ltln_weak_strong by simp_all blast+
lemma ltln_weak_to_strong:
"\<lbrakk>w \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi>; \<not> w \<Turnstile>\<^sub>n G\<^sub>n \<phi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi>"
"\<lbrakk>w \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi>; w \<Turnstile>\<^sub>n F\<^sub>n \<psi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi>"
"\<lbrakk>w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi>; \<not> w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi>"
"\<lbrakk>w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi>; w \<Turnstile>\<^sub>n F\<^sub>n \<phi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi>"
unfolding ltln_weak_strong[of w \<phi> \<psi>] by auto
lemma ltln_StrongRelease_to_Until:
"w \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<psi> U\<^sub>n (\<phi> and\<^sub>n \<psi>)"
using order.order_iff_strict by auto
lemma ltln_Release_to_WeakUntil:
"w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<psi> W\<^sub>n (\<phi> and\<^sub>n \<psi>)"
by (meson ltln_StrongRelease_to_Until ltln_weak_strong semantics_ltln.simps(6))
lemma ltln_WeakUntil_to_Release:
"w \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<psi> R\<^sub>n (\<phi> or\<^sub>n \<psi>)"
by (metis ltln_StrongRelease_to_Until not\<^sub>n.simps(6,9,10) not\<^sub>n_semantics)
lemma ltln_Until_to_StrongRelease:
"w \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<psi> M\<^sub>n (\<phi> or\<^sub>n \<psi>)"
by (metis ltln_Release_to_WeakUntil not\<^sub>n.simps(6,8,11) not\<^sub>n_semantics)
subsubsection \<open>GF and FG semantics\<close>
lemma GF_suffix:
"suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)"
by auto (metis ab_semigroup_add_class.add_ac(1) add.left_commute)
lemma FG_suffix:
"suffix i w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>) \<longleftrightarrow> w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>)"
by (auto simp: algebra_simps) (metis add.commute add.left_commute)
lemma GF_Inf_many:
"w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>) \<longleftrightarrow> (\<exists>\<^sub>\<infinity> i. suffix i w \<Turnstile>\<^sub>n \<phi>)"
unfolding INFM_nat_le
by simp (blast dest: le_Suc_ex intro: le_add1)
lemma FG_Alm_all:
"w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>) \<longleftrightarrow> (\<forall>\<^sub>\<infinity> i. suffix i w \<Turnstile>\<^sub>n \<phi>)"
unfolding MOST_nat_le
by simp (blast dest: le_Suc_ex intro: le_add1)
(* TODO: move to Infinite_Set.thy ?? *)
lemma MOST_nat_add:
"(\<forall>\<^sub>\<infinity>i::nat. P i) \<longleftrightarrow> (\<forall>\<^sub>\<infinity>i. P (i + j))"
by (simp add: cofinite_eq_sequentially)
lemma INFM_nat_add:
"(\<exists>\<^sub>\<infinity>i::nat. P i) \<longleftrightarrow> (\<exists>\<^sub>\<infinity>i. P (i + j))"
using MOST_nat_add not_MOST not_INFM by blast
lemma FG_suffix_G:
"w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>) \<Longrightarrow> \<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>"
proof -
assume "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>)"
then have "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n (G\<^sub>n \<phi>))"
by (meson globally_release semantics_ltln.simps(8))
then show "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>"
unfolding FG_Alm_all .
qed
lemma Alm_all_GF_F:
"\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>) \<longleftrightarrow> suffix i w \<Turnstile>\<^sub>n F\<^sub>n \<psi>"
unfolding MOST_nat
proof standard+
fix i :: nat
assume "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)"
then show "suffix i w \<Turnstile>\<^sub>n F\<^sub>n \<psi>"
unfolding GF_Inf_many INFM_nat by fastforce
next
fix i :: nat
assume suffix: "suffix i w \<Turnstile>\<^sub>n F\<^sub>n \<psi>"
assume max: "i > Max {i. suffix i w \<Turnstile>\<^sub>n \<psi>}"
with suffix obtain j where "j \<ge> i" and j_suffix: "suffix j w \<Turnstile>\<^sub>n \<psi>"
by simp (blast intro: le_add1)
with max have j_max: "j > Max {i. suffix i w \<Turnstile>\<^sub>n \<psi>}"
by fastforce
show "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)"
proof (cases "w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)")
case False
then have "\<not> (\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>)"
unfolding GF_Inf_many by simp
then have "finite {i. suffix i w \<Turnstile>\<^sub>n \<psi>}"
by (simp add: INFM_iff_infinite)
then have "\<forall>i>Max {i. suffix i w \<Turnstile>\<^sub>n \<psi>}. \<not> suffix i w \<Turnstile>\<^sub>n \<psi>"
using Max_ge not_le by auto
then show ?thesis
using j_suffix j_max by blast
qed force
qed
lemma Alm_all_FG_G:
"\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>) \<longleftrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
unfolding MOST_nat
proof standard+
fix i :: nat
assume "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
then show "suffix i w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>)"
unfolding FG_Alm_all INFM_nat by fastforce
next
fix i :: nat
assume suffix: "suffix i w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>)"
assume max: "i > Max {i. \<not> suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>}"
with suffix have "\<forall>\<^sub>\<infinity>j. suffix (i + j) w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
using FG_suffix_G[of "suffix i w"] suffix_suffix
by fastforce
then have "\<not> (\<exists>\<^sub>\<infinity>j. \<not> suffix j w \<Turnstile>\<^sub>n G\<^sub>n \<psi>)"
using MOST_nat_add[of "\<lambda>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>" i]
by (simp add: algebra_simps)
then have "finite {i. \<not> suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>}"
by (simp add: INFM_iff_infinite)
with max show "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
using Max_ge leD by blast
qed
subsubsection \<open>Expansion\<close>
lemma ltln_expand_Until:
"\<xi> \<Turnstile>\<^sub>n \<phi> U\<^sub>n \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>n \<psi> or\<^sub>n (\<phi> and\<^sub>n (X\<^sub>n (\<phi> U\<^sub>n \<psi>))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain i where "suffix i \<xi> \<Turnstile>\<^sub>n \<psi>"
and "\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>n \<phi>"
by auto
thus ?rhs
by (cases i) auto
next
assume ?rhs
show ?lhs
proof (cases "\<xi> \<Turnstile>\<^sub>n \<psi>")
case False
then have "\<xi> \<Turnstile>\<^sub>n \<phi>" and "\<xi> \<Turnstile>\<^sub>n X\<^sub>n (\<phi> U\<^sub>n \<psi>)"
- using `?rhs` by auto
+ using \<open>?rhs\<close> by auto
thus ?lhs
using less_Suc_eq_0_disj suffix_singleton_suffix by force
qed force
qed
lemma ltln_expand_Release:
"\<xi> \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>n \<psi> and\<^sub>n (\<phi> or\<^sub>n (X\<^sub>n (\<phi> R\<^sub>n \<psi>))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
thus ?rhs
using less_Suc_eq_0_disj by force
next
assume ?rhs
{
fix i
assume "\<not> suffix i \<xi> \<Turnstile>\<^sub>n \<psi>"
then have "\<exists>j<i. suffix j \<xi> \<Turnstile>\<^sub>n \<phi>"
- using `?rhs` by (cases i) force+
+ using \<open>?rhs\<close> by (cases i) force+
}
thus ?lhs
by auto
qed
lemma ltln_expand_WeakUntil:
"\<xi> \<Turnstile>\<^sub>n \<phi> W\<^sub>n \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>n \<psi> or\<^sub>n (\<phi> and\<^sub>n (X\<^sub>n (\<phi> W\<^sub>n \<psi>))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
thus ?rhs
by (metis ltln_expand_Release ltln_expand_Until ltln_weak_strong(1) semantics_ltln.simps(2,5,6,7))
next
assume ?rhs
{
fix i
assume "\<not> suffix i \<xi> \<Turnstile>\<^sub>n \<phi>"
then have "\<exists>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>n \<psi>"
- using `?rhs` by (cases i) force+
+ using \<open>?rhs\<close> by (cases i) force+
}
thus ?lhs
by auto
qed
lemma ltln_expand_StrongRelease:
"\<xi> \<Turnstile>\<^sub>n \<phi> M\<^sub>n \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>n \<psi> and\<^sub>n (\<phi> or\<^sub>n (X\<^sub>n (\<phi> M\<^sub>n \<psi>))))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain i where "suffix i \<xi> \<Turnstile>\<^sub>n \<phi>"
and "\<forall>j\<le>i. suffix j \<xi> \<Turnstile>\<^sub>n \<psi>"
by auto
thus ?rhs
by (cases i) auto
next
assume ?rhs
show ?lhs
proof (cases "\<xi> \<Turnstile>\<^sub>n \<phi>")
case True
thus ?lhs
- using `?rhs` ltln_expand_WeakUntil by fastforce
+ using \<open>?rhs\<close> ltln_expand_WeakUntil by fastforce
next
case False
thus ?lhs
- by (metis `?rhs` ltln_expand_WeakUntil not\<^sub>n.simps(5,6,7,11) not\<^sub>n_semantics)
+ by (metis \<open>?rhs\<close> ltln_expand_WeakUntil not\<^sub>n.simps(5,6,7,11) not\<^sub>n_semantics)
qed
qed
lemma ltln_Release_alterdef:
"w \<Turnstile>\<^sub>n \<phi> R\<^sub>n \<psi> \<longleftrightarrow> w \<Turnstile>\<^sub>n (G\<^sub>n \<psi>) or\<^sub>n (\<psi> U\<^sub>n (\<phi> and\<^sub>n \<psi>))"
proof (cases "\<exists>i. \<not>suffix i w \<Turnstile>\<^sub>n \<psi>")
case True
define i where "i \<equiv> Least (\<lambda>i. \<not>suffix i w \<Turnstile>\<^sub>n \<psi>)"
have "\<And>j. j < i \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<psi>" and "\<not> suffix i w \<Turnstile>\<^sub>n \<psi>"
using True LeastI not_less_Least unfolding i_def by fast+
hence *: "\<forall>i. suffix i w \<Turnstile>\<^sub>n \<psi> \<or> (\<exists>j<i. suffix j w \<Turnstile>\<^sub>n \<phi>) \<Longrightarrow> (\<exists>i. (suffix i w \<Turnstile>\<^sub>n \<psi> \<and> suffix i w \<Turnstile>\<^sub>n \<phi>) \<and> (\<forall>j<i. suffix j w \<Turnstile>\<^sub>n \<psi>))"
by fastforce
hence "\<exists>i. (suffix i w \<Turnstile>\<^sub>n \<psi> \<and> suffix i w \<Turnstile>\<^sub>n \<phi>) \<and> (\<forall>j<i. suffix j w \<Turnstile>\<^sub>n \<psi>) \<Longrightarrow> (\<forall>i. suffix i w \<Turnstile>\<^sub>n \<psi> \<or> (\<exists>j<i. suffix j w \<Turnstile>\<^sub>n \<phi>))"
using linorder_cases by blast
thus ?thesis
using True * by auto
qed auto
subsection \<open>LTL in restricted Negation Normal Form\<close>
text \<open>Some algorithms do not handle the operators W and M,
hence we also provide a datatype without these two operators.\<close>
subsubsection \<open>Syntax\<close>
datatype (atoms_ltlr: 'a) ltlr =
True_ltlr ("true\<^sub>r")
| False_ltlr ("false\<^sub>r")
| Prop_ltlr 'a ("prop\<^sub>r'(_')")
| Nprop_ltlr 'a ("nprop\<^sub>r'(_')")
| And_ltlr "'a ltlr" "'a ltlr" ("_ and\<^sub>r _" [82,82] 81)
| Or_ltlr "'a ltlr" "'a ltlr" ("_ or\<^sub>r _" [84,84] 83)
| Next_ltlr "'a ltlr" ("X\<^sub>r _" [88] 87)
| Until_ltlr "'a ltlr" "'a ltlr" ("_ U\<^sub>r _" [84,84] 83)
| Release_ltlr "'a ltlr" "'a ltlr" ("_ R\<^sub>r _" [84,84] 83)
subsubsection \<open>Semantics\<close>
primrec semantics_ltlr :: "['a set word, 'a ltlr] \<Rightarrow> bool" ("_ \<Turnstile>\<^sub>r _" [80,80] 80)
where
"\<xi> \<Turnstile>\<^sub>r true\<^sub>r = True"
| "\<xi> \<Turnstile>\<^sub>r false\<^sub>r = False"
| "\<xi> \<Turnstile>\<^sub>r prop\<^sub>r(q) = (q \<in> \<xi> 0)"
| "\<xi> \<Turnstile>\<^sub>r nprop\<^sub>r(q) = (q \<notin> \<xi> 0)"
| "\<xi> \<Turnstile>\<^sub>r \<phi> and\<^sub>r \<psi> = (\<xi> \<Turnstile>\<^sub>r \<phi> \<and> \<xi> \<Turnstile>\<^sub>r \<psi>)"
| "\<xi> \<Turnstile>\<^sub>r \<phi> or\<^sub>r \<psi> = (\<xi> \<Turnstile>\<^sub>r \<phi> \<or> \<xi> \<Turnstile>\<^sub>r \<psi>)"
| "\<xi> \<Turnstile>\<^sub>r X\<^sub>r \<phi> = (suffix 1 \<xi> \<Turnstile>\<^sub>r \<phi>)"
| "\<xi> \<Turnstile>\<^sub>r \<phi> U\<^sub>r \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>\<^sub>r \<psi> \<and> (\<forall>j<i. suffix j \<xi> \<Turnstile>\<^sub>r \<phi>))"
| "\<xi> \<Turnstile>\<^sub>r \<phi> R\<^sub>r \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>\<^sub>r \<psi> \<or> (\<exists>j<i. suffix j \<xi> \<Turnstile>\<^sub>r \<phi>))"
subsubsection \<open>Conversion\<close>
fun ltln_to_ltlr :: "'a ltln \<Rightarrow> 'a ltlr"
where
"ltln_to_ltlr true\<^sub>n = true\<^sub>r"
| "ltln_to_ltlr false\<^sub>n = false\<^sub>r"
| "ltln_to_ltlr prop\<^sub>n(a) = prop\<^sub>r(a)"
| "ltln_to_ltlr nprop\<^sub>n(a) = nprop\<^sub>r(a)"
| "ltln_to_ltlr (\<phi> and\<^sub>n \<psi>) = (ltln_to_ltlr \<phi>) and\<^sub>r (ltln_to_ltlr \<psi>)"
| "ltln_to_ltlr (\<phi> or\<^sub>n \<psi>) = (ltln_to_ltlr \<phi>) or\<^sub>r (ltln_to_ltlr \<psi>)"
| "ltln_to_ltlr (X\<^sub>n \<phi>) = X\<^sub>r (ltln_to_ltlr \<phi>)"
| "ltln_to_ltlr (\<phi> U\<^sub>n \<psi>) = (ltln_to_ltlr \<phi>) U\<^sub>r (ltln_to_ltlr \<psi>)"
| "ltln_to_ltlr (\<phi> R\<^sub>n \<psi>) = (ltln_to_ltlr \<phi>) R\<^sub>r (ltln_to_ltlr \<psi>)"
| "ltln_to_ltlr (\<phi> W\<^sub>n \<psi>) = (ltln_to_ltlr \<psi>) R\<^sub>r ((ltln_to_ltlr \<phi>) or\<^sub>r (ltln_to_ltlr \<psi>))"
| "ltln_to_ltlr (\<phi> M\<^sub>n \<psi>) = (ltln_to_ltlr \<psi>) U\<^sub>r ((ltln_to_ltlr \<phi>) and\<^sub>r (ltln_to_ltlr \<psi>))"
fun ltlr_to_ltln :: "'a ltlr \<Rightarrow> 'a ltln"
where
"ltlr_to_ltln true\<^sub>r = true\<^sub>n"
| "ltlr_to_ltln false\<^sub>r = false\<^sub>n"
| "ltlr_to_ltln prop\<^sub>r(a) = prop\<^sub>n(a)"
| "ltlr_to_ltln nprop\<^sub>r(a) = nprop\<^sub>n(a)"
| "ltlr_to_ltln (\<phi> and\<^sub>r \<psi>) = (ltlr_to_ltln \<phi>) and\<^sub>n (ltlr_to_ltln \<psi>)"
| "ltlr_to_ltln (\<phi> or\<^sub>r \<psi>) = (ltlr_to_ltln \<phi>) or\<^sub>n (ltlr_to_ltln \<psi>)"
| "ltlr_to_ltln (X\<^sub>r \<phi>) = X\<^sub>n (ltlr_to_ltln \<phi>)"
| "ltlr_to_ltln (\<phi> U\<^sub>r \<psi>) = (ltlr_to_ltln \<phi>) U\<^sub>n (ltlr_to_ltln \<psi>)"
| "ltlr_to_ltln (\<phi> R\<^sub>r \<psi>) = (ltlr_to_ltln \<phi>) R\<^sub>n (ltlr_to_ltln \<psi>)"
lemma ltln_to_ltlr_semantics:
"w \<Turnstile>\<^sub>r ltln_to_ltlr \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>"
by (induction \<phi> arbitrary: w) (unfold ltln_WeakUntil_to_Release ltln_StrongRelease_to_Until, simp_all)
lemma ltlr_to_ltln_semantics:
"w \<Turnstile>\<^sub>n ltlr_to_ltln \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>r \<phi>"
by (induction \<phi> arbitrary: w) simp_all
subsubsection \<open>Negation\<close>
fun not\<^sub>r
where
"not\<^sub>r true\<^sub>r = false\<^sub>r"
| "not\<^sub>r false\<^sub>r = true\<^sub>r"
| "not\<^sub>r prop\<^sub>r(a) = nprop\<^sub>r(a)"
| "not\<^sub>r nprop\<^sub>r(a) = prop\<^sub>r(a)"
| "not\<^sub>r (\<phi> and\<^sub>r \<psi>) = (not\<^sub>r \<phi>) or\<^sub>r (not\<^sub>r \<psi>)"
| "not\<^sub>r (\<phi> or\<^sub>r \<psi>) = (not\<^sub>r \<phi>) and\<^sub>r (not\<^sub>r \<psi>)"
| "not\<^sub>r (X\<^sub>r \<phi>) = X\<^sub>r (not\<^sub>r \<phi>)"
| "not\<^sub>r (\<phi> U\<^sub>r \<psi>) = (not\<^sub>r \<phi>) R\<^sub>r (not\<^sub>r \<psi>)"
| "not\<^sub>r (\<phi> R\<^sub>r \<psi>) = (not\<^sub>r \<phi>) U\<^sub>r (not\<^sub>r \<psi>)"
lemma not\<^sub>r_semantics [simp]:
"w \<Turnstile>\<^sub>r not\<^sub>r \<phi> \<longleftrightarrow> \<not> w \<Turnstile>\<^sub>r \<phi>"
by (induction \<phi> arbitrary: w) auto
subsubsection \<open>Subformulas\<close>
fun subfrmlsr :: "'a ltlr \<Rightarrow> 'a ltlr set"
where
"subfrmlsr (\<phi> and\<^sub>r \<psi>) = {\<phi> and\<^sub>r \<psi>} \<union> subfrmlsr \<phi> \<union> subfrmlsr \<psi>"
| "subfrmlsr (\<phi> or\<^sub>r \<psi>) = {\<phi> or\<^sub>r \<psi>} \<union> subfrmlsr \<phi> \<union> subfrmlsr \<psi>"
| "subfrmlsr (\<phi> U\<^sub>r \<psi>) = {\<phi> U\<^sub>r \<psi>} \<union> subfrmlsr \<phi> \<union> subfrmlsr \<psi>"
| "subfrmlsr (\<phi> R\<^sub>r \<psi>) = {\<phi> R\<^sub>r \<psi>} \<union> subfrmlsr \<phi> \<union> subfrmlsr \<psi>"
| "subfrmlsr (X\<^sub>r \<phi>) = {X\<^sub>r \<phi>} \<union> subfrmlsr \<phi>"
| "subfrmlsr x = {x}"
lemma subfrmlsr_id[simp]:
"\<phi> \<in> subfrmlsr \<phi>"
by (induction \<phi>) auto
lemma subfrmlsr_finite:
"finite (subfrmlsr \<phi>)"
by (induction \<phi>) auto
lemma subfrmlsr_subset:
"\<psi> \<in> subfrmlsr \<phi> \<Longrightarrow> subfrmlsr \<psi> \<subseteq> subfrmlsr \<phi>"
by (induction \<phi>) auto
lemma subfrmlsr_size:
"\<psi> \<in> subfrmlsr \<phi> \<Longrightarrow> size \<psi> < size \<phi> \<or> \<psi> = \<phi>"
by (induction \<phi>) auto
subsubsection \<open>Expansion lemmas\<close>
lemma ltlr_expand_Until:
"\<xi> \<Turnstile>\<^sub>r \<phi> U\<^sub>r \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>r \<psi> or\<^sub>r (\<phi> and\<^sub>r (X\<^sub>r (\<phi> U\<^sub>r \<psi>))))"
by (metis ltln_expand_Until ltlr_to_ltln.simps(5-8) ltlr_to_ltln_semantics)
lemma ltlr_expand_Release:
"\<xi> \<Turnstile>\<^sub>r \<phi> R\<^sub>r \<psi> \<longleftrightarrow> (\<xi> \<Turnstile>\<^sub>r \<psi> and\<^sub>r (\<phi> or\<^sub>r (X\<^sub>r (\<phi> R\<^sub>r \<psi>))))"
by (metis ltln_expand_Release ltlr_to_ltln.simps(5-7,9) ltlr_to_ltln_semantics)
subsection \<open>Propositional LTL\<close>
text \<open>We define the syntax and semantics of propositional linear-time
temporal logic PLTL.
PLTL formulas are built from atomic formulas, propositional connectives,
and the temporal operators ``next'' and ``until''. The following data
type definition is parameterized by the type of states over which
formulas are evaluated.\<close>
subsubsection \<open>Syntax\<close>
datatype 'a pltl =
False_ltlp ("false\<^sub>p")
| Atom_ltlp "'a \<Rightarrow> bool" ("atom\<^sub>p'(_')")
| Implies_ltlp "'a pltl" "'a pltl" ("_ implies\<^sub>p _" [81,81] 80)
| Next_ltlp "'a pltl" ("X\<^sub>p _" [88] 87)
| Until_ltlp "'a pltl" "'a pltl" ("_ U\<^sub>p _" [84,84] 83)
\<comment> \<open>Further connectives of PLTL can be defined in terms of the existing syntax.\<close>
definition Not_ltlp ("not\<^sub>p _" [85] 85)
where
"not\<^sub>p \<phi> \<equiv> \<phi> implies\<^sub>p false\<^sub>p"
definition True_ltlp ("true\<^sub>p")
where
"true\<^sub>p \<equiv> not\<^sub>p false\<^sub>p"
definition Or_ltlp ("_ or\<^sub>p _" [81,81] 80)
where
"\<phi> or\<^sub>p \<psi> \<equiv> (not\<^sub>p \<phi>) implies\<^sub>p \<psi>"
definition And_ltlp ("_ and\<^sub>p _" [82,82] 81)
where
"\<phi> and\<^sub>p \<psi> \<equiv> not\<^sub>p ((not\<^sub>p \<phi>) or\<^sub>p (not\<^sub>p \<psi>))"
definition Eventually_ltlp ("F\<^sub>p _" [88] 87)
where
"F\<^sub>p \<phi> \<equiv> true\<^sub>p U\<^sub>p \<phi>"
definition Always_ltlp ("G\<^sub>p _" [88] 87)
where
"G\<^sub>p \<phi> \<equiv> not\<^sub>p (F\<^sub>p (not\<^sub>p \<phi>))"
definition Release_ltlp ("_ R\<^sub>p _" [84,84] 83)
where
"\<phi> R\<^sub>p \<psi> \<equiv> not\<^sub>p ((not\<^sub>p \<phi>) U\<^sub>p (not\<^sub>p \<psi>))"
definition WeakUntil_ltlp ("_ W\<^sub>p _" [84,84] 83)
where
"\<phi> W\<^sub>p \<psi> \<equiv> \<psi> R\<^sub>p (\<phi> or\<^sub>p \<psi>)"
definition StrongRelease_ltlp ("_ M\<^sub>p _" [84,84] 83)
where
"\<phi> M\<^sub>p \<psi> \<equiv> \<psi> U\<^sub>p (\<phi> and\<^sub>p \<psi>)"
subsubsection \<open>Semantics\<close>
fun semantics_pltl :: "['a word, 'a pltl] \<Rightarrow> bool" ("_ \<Turnstile>\<^sub>p _" [80,80] 80)
where
"w \<Turnstile>\<^sub>p false\<^sub>p = False"
| "w \<Turnstile>\<^sub>p atom\<^sub>p(p) = (p (w 0))"
| "w \<Turnstile>\<^sub>p \<phi> implies\<^sub>p \<psi> = (w \<Turnstile>\<^sub>p \<phi> \<longrightarrow> w \<Turnstile>\<^sub>p \<psi>)"
| "w \<Turnstile>\<^sub>p X\<^sub>p \<phi> = (suffix 1 w \<Turnstile>\<^sub>p \<phi>)"
| "w \<Turnstile>\<^sub>p \<phi> U\<^sub>p \<psi> = (\<exists>i. suffix i w \<Turnstile>\<^sub>p \<psi> \<and> (\<forall>j<i. suffix j w \<Turnstile>\<^sub>p \<phi>))"
lemma semantics_pltl_sugar [simp]:
"w \<Turnstile>\<^sub>p not\<^sub>p \<phi> = (\<not>w \<Turnstile>\<^sub>p \<phi>)"
"w \<Turnstile>\<^sub>p true\<^sub>p = True"
"w \<Turnstile>\<^sub>p \<phi> or\<^sub>p \<psi> = (w \<Turnstile>\<^sub>p \<phi> \<or> w \<Turnstile>\<^sub>p \<psi>)"
"w \<Turnstile>\<^sub>p \<phi> and\<^sub>p \<psi> = (w \<Turnstile>\<^sub>p \<phi> \<and> w \<Turnstile>\<^sub>p \<psi>)"
"w \<Turnstile>\<^sub>p F\<^sub>p \<phi> = (\<exists>i. suffix i w \<Turnstile>\<^sub>p \<phi>)"
"w \<Turnstile>\<^sub>p G\<^sub>p \<phi> = (\<forall>i. suffix i w \<Turnstile>\<^sub>p \<phi>)"
"w \<Turnstile>\<^sub>p \<phi> R\<^sub>p \<psi> = (\<forall>i. suffix i w \<Turnstile>\<^sub>p \<psi> \<or> (\<exists>j<i. suffix j w \<Turnstile>\<^sub>p \<phi>))"
"w \<Turnstile>\<^sub>p \<phi> W\<^sub>p \<psi> = (\<forall>i. suffix i w \<Turnstile>\<^sub>p \<phi> \<or> (\<exists>j\<le>i. suffix j w \<Turnstile>\<^sub>p \<psi>))"
"w \<Turnstile>\<^sub>p \<phi> M\<^sub>p \<psi> = (\<exists>i. suffix i w \<Turnstile>\<^sub>p \<phi> \<and> (\<forall>j\<le>i. suffix j w \<Turnstile>\<^sub>p \<psi>))"
by (auto simp: Not_ltlp_def True_ltlp_def Or_ltlp_def And_ltlp_def Eventually_ltlp_def Always_ltlp_def Release_ltlp_def WeakUntil_ltlp_def StrongRelease_ltlp_def) (insert le_neq_implies_less, blast)+
definition "language_ltlp \<phi> \<equiv> {\<xi>. \<xi> \<Turnstile>\<^sub>p \<phi>}"
subsubsection \<open>Conversion\<close>
fun ltlc_to_pltl :: "'a ltlc \<Rightarrow> 'a set pltl"
where
"ltlc_to_pltl true\<^sub>c = true\<^sub>p"
| "ltlc_to_pltl false\<^sub>c = false\<^sub>p"
| "ltlc_to_pltl (prop\<^sub>c(q)) = atom\<^sub>p((\<in>) q)"
| "ltlc_to_pltl (not\<^sub>c \<phi>) = not\<^sub>p (ltlc_to_pltl \<phi>)"
| "ltlc_to_pltl (\<phi> and\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) and\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (\<phi> or\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) or\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (\<phi> implies\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) implies\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (X\<^sub>c \<phi>) = X\<^sub>p (ltlc_to_pltl \<phi>)"
| "ltlc_to_pltl (F\<^sub>c \<phi>) = F\<^sub>p (ltlc_to_pltl \<phi>)"
| "ltlc_to_pltl (G\<^sub>c \<phi>) = G\<^sub>p (ltlc_to_pltl \<phi>)"
| "ltlc_to_pltl (\<phi> U\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) U\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (\<phi> R\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) R\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (\<phi> W\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) W\<^sub>p (ltlc_to_pltl \<psi>)"
| "ltlc_to_pltl (\<phi> M\<^sub>c \<psi>) = (ltlc_to_pltl \<phi>) M\<^sub>p (ltlc_to_pltl \<psi>)"
lemma ltlc_to_pltl_semantics [simp]:
"w \<Turnstile>\<^sub>p (ltlc_to_pltl \<phi>) \<longleftrightarrow> w \<Turnstile>\<^sub>c \<phi>"
by (induction \<phi> arbitrary: w) simp_all
subsubsection \<open>Atoms\<close>
fun atoms_pltl :: "'a pltl \<Rightarrow> ('a \<Rightarrow> bool) set"
where
"atoms_pltl false\<^sub>p = {}"
| "atoms_pltl atom\<^sub>p(p) = {p}"
| "atoms_pltl (\<phi> implies\<^sub>p \<psi>) = atoms_pltl \<phi> \<union> atoms_pltl \<psi>"
| "atoms_pltl (X\<^sub>p \<phi>) = atoms_pltl \<phi>"
| "atoms_pltl (\<phi> U\<^sub>p \<psi>) = atoms_pltl \<phi> \<union> atoms_pltl \<psi>"
lemma atoms_finite [iff]:
"finite (atoms_pltl \<phi>)"
by (induct \<phi>) auto
lemma atoms_pltl_sugar [simp]:
"atoms_pltl (not\<^sub>p \<phi>) = atoms_pltl \<phi>"
"atoms_pltl true\<^sub>p = {}"
"atoms_pltl (\<phi> or\<^sub>p \<psi>) = atoms_pltl \<phi> \<union> atoms_pltl \<psi>"
"atoms_pltl (\<phi> and\<^sub>p \<psi>) = atoms_pltl \<phi> \<union> atoms_pltl \<psi>"
"atoms_pltl (F\<^sub>p \<phi>) = atoms_pltl \<phi>"
"atoms_pltl (G\<^sub>p \<phi>) = atoms_pltl \<phi>"
by (auto simp: Not_ltlp_def True_ltlp_def Or_ltlp_def And_ltlp_def Eventually_ltlp_def Always_ltlp_def)
end
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/Advice.thy b/thys/LTL_Master_Theorem/Logical_Characterization/Advice.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/Advice.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/Advice.thy
@@ -1,1012 +1,1012 @@
(*
Author: Benedikt Seidl
Author: Salomon Sickert
License: BSD
*)
section \<open>Advice functions\<close>
theory Advice
imports
LTL.LTL LTL.Equivalence_Relations
Syntactic_Fragments_and_Stability After
begin
subsection \<open>The GF and FG Advice Functions\<close>
fun GF_advice :: "'a ltln \<Rightarrow> 'a ltln set \<Rightarrow> 'a ltln" ("_[_]\<^sub>\<nu>" [90,60] 89)
where
"(X\<^sub>n \<psi>)[X]\<^sub>\<nu> = X\<^sub>n (\<psi>[X]\<^sub>\<nu>)"
| "(\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (\<psi>\<^sub>1[X]\<^sub>\<nu>) and\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>)"
| "(\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (\<psi>\<^sub>1[X]\<^sub>\<nu>) or\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>)"
| "(\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (\<psi>\<^sub>1[X]\<^sub>\<nu>) W\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>)"
| "(\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (\<psi>\<^sub>1[X]\<^sub>\<nu>) R\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>)"
| "(\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (if (\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2) \<in> X then (\<psi>\<^sub>1[X]\<^sub>\<nu>) W\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>) else false\<^sub>n)"
| "(\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2)[X]\<^sub>\<nu> = (if (\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2) \<in> X then (\<psi>\<^sub>1[X]\<^sub>\<nu>) R\<^sub>n (\<psi>\<^sub>2[X]\<^sub>\<nu>) else false\<^sub>n)"
| "\<phi>[_]\<^sub>\<nu> = \<phi>"
fun FG_advice :: "'a ltln \<Rightarrow> 'a ltln set \<Rightarrow> 'a ltln" ("_[_]\<^sub>\<mu>" [90,60] 89)
where
"(X\<^sub>n \<psi>)[Y]\<^sub>\<mu> = X\<^sub>n (\<psi>[Y]\<^sub>\<mu>)"
| "(\<psi>\<^sub>1 and\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (\<psi>\<^sub>1[Y]\<^sub>\<mu>) and\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
| "(\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (\<psi>\<^sub>1[Y]\<^sub>\<mu>) or\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
| "(\<psi>\<^sub>1 U\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (\<psi>\<^sub>1[Y]\<^sub>\<mu>) U\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
| "(\<psi>\<^sub>1 M\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (\<psi>\<^sub>1[Y]\<^sub>\<mu>) M\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
| "(\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (if (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2) \<in> Y then true\<^sub>n else (\<psi>\<^sub>1[Y]\<^sub>\<mu>) U\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>))"
| "(\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2)[Y]\<^sub>\<mu> = (if (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2) \<in> Y then true\<^sub>n else (\<psi>\<^sub>1[Y]\<^sub>\<mu>) M\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>))"
| "\<phi>[_]\<^sub>\<mu> = \<phi>"
lemma GF_advice_\<nu>LTL:
"\<phi>[X]\<^sub>\<nu> \<in> \<nu>LTL"
"\<phi> \<in> \<nu>LTL \<Longrightarrow> \<phi>[X]\<^sub>\<nu> = \<phi>"
by (induction \<phi>) auto
lemma FG_advice_\<mu>LTL:
"\<phi>[X]\<^sub>\<mu> \<in> \<mu>LTL"
"\<phi> \<in> \<mu>LTL \<Longrightarrow> \<phi>[X]\<^sub>\<mu> = \<phi>"
by (induction \<phi>) auto
lemma GF_advice_subfrmlsn:
"subfrmlsn (\<phi>[X]\<^sub>\<nu>) \<subseteq> {\<psi>[X]\<^sub>\<nu> | \<psi>. \<psi> \<in> subfrmlsn \<phi>}"
by (induction \<phi>) force+
lemma FG_advice_subfrmlsn:
"subfrmlsn (\<phi>[Y]\<^sub>\<mu>) \<subseteq> {\<psi>[Y]\<^sub>\<mu> | \<psi>. \<psi> \<in> subfrmlsn \<phi>}"
by (induction \<phi>) force+
lemma GF_advice_subfrmlsn_card:
"card (subfrmlsn (\<phi>[X]\<^sub>\<nu>)) \<le> card (subfrmlsn \<phi>)"
proof -
have "card (subfrmlsn (\<phi>[X]\<^sub>\<nu>)) \<le> card {\<psi>[X]\<^sub>\<nu> | \<psi>. \<psi> \<in> subfrmlsn \<phi>}"
by (simp add: subfrmlsn_finite GF_advice_subfrmlsn card_mono)
also have "\<dots> \<le> card (subfrmlsn \<phi>)"
by (metis Collect_mem_eq card_image_le image_Collect subfrmlsn_finite)
finally show ?thesis .
qed
lemma FG_advice_subfrmlsn_card:
"card (subfrmlsn (\<phi>[Y]\<^sub>\<mu>)) \<le> card (subfrmlsn \<phi>)"
proof -
have "card (subfrmlsn (\<phi>[Y]\<^sub>\<mu>)) \<le> card {\<psi>[Y]\<^sub>\<mu> | \<psi>. \<psi> \<in> subfrmlsn \<phi>}"
by (simp add: subfrmlsn_finite FG_advice_subfrmlsn card_mono)
also have "\<dots> \<le> card (subfrmlsn \<phi>)"
by (metis Collect_mem_eq card_image_le image_Collect subfrmlsn_finite)
finally show ?thesis .
qed
lemma GF_advice_monotone:
"X \<subseteq> Y \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<nu>"
proof (induction \<phi> arbitrary: w)
case (Until_ltln \<phi> \<psi>)
then show ?case
by (cases "\<phi> U\<^sub>n \<psi> \<in> X") (simp_all, blast)
next
case (Release_ltln \<phi> \<psi>)
then show ?case by (simp, blast)
next
case (WeakUntil_ltln \<phi> \<psi>)
then show ?case by (simp, blast)
next
case (StrongRelease_ltln \<phi> \<psi>)
then show ?case
by (cases "\<phi> M\<^sub>n \<psi> \<in> X") (simp_all, blast)
qed auto
lemma FG_advice_monotone:
"X \<subseteq> Y \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<mu> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>"
proof (induction \<phi> arbitrary: w)
case (Until_ltln \<phi> \<psi>)
then show ?case by (simp, blast)
next
case (Release_ltln \<phi> \<psi>)
then show ?case
by (cases "\<phi> R\<^sub>n \<psi> \<in> X") (auto, blast)
next
case (WeakUntil_ltln \<phi> \<psi>)
then show ?case
by (cases "\<phi> W\<^sub>n \<psi> \<in> X") (auto, blast)
next
case (StrongRelease_ltln \<phi> \<psi>)
then show ?case by (simp, blast)
qed auto
lemma GF_advice_ite_simps[simp]:
"(if P then true\<^sub>n else false\<^sub>n)[X]\<^sub>\<nu> = (if P then true\<^sub>n else false\<^sub>n)"
"(if P then false\<^sub>n else true\<^sub>n)[X]\<^sub>\<nu> = (if P then false\<^sub>n else true\<^sub>n)"
by simp_all
lemma FG_advice_ite_simps[simp]:
"(if P then true\<^sub>n else false\<^sub>n)[Y]\<^sub>\<mu> = (if P then true\<^sub>n else false\<^sub>n)"
"(if P then false\<^sub>n else true\<^sub>n)[Y]\<^sub>\<mu> = (if P then false\<^sub>n else true\<^sub>n)"
by simp_all
subsection \<open>Advice Functions on Nested Propositions\<close>
definition nested_prop_atoms\<^sub>\<nu> :: "'a ltln \<Rightarrow> 'a ltln set \<Rightarrow> 'a ltln set"
where
"nested_prop_atoms\<^sub>\<nu> \<phi> X = {\<psi>[X]\<^sub>\<nu> | \<psi>. \<psi> \<in> nested_prop_atoms \<phi>}"
definition nested_prop_atoms\<^sub>\<mu> :: "'a ltln \<Rightarrow> 'a ltln set \<Rightarrow> 'a ltln set"
where
"nested_prop_atoms\<^sub>\<mu> \<phi> X = {\<psi>[X]\<^sub>\<mu> | \<psi>. \<psi> \<in> nested_prop_atoms \<phi>}"
lemma nested_prop_atoms\<^sub>\<nu>_finite:
"finite (nested_prop_atoms\<^sub>\<nu> \<phi> X)"
by (simp add: nested_prop_atoms\<^sub>\<nu>_def nested_prop_atoms_finite)
lemma nested_prop_atoms\<^sub>\<mu>_finite:
"finite (nested_prop_atoms\<^sub>\<mu> \<phi> X)"
by (simp add: nested_prop_atoms\<^sub>\<mu>_def nested_prop_atoms_finite)
lemma nested_prop_atoms\<^sub>\<nu>_card:
"card (nested_prop_atoms\<^sub>\<nu> \<phi> X) \<le> card (nested_prop_atoms \<phi>)"
unfolding nested_prop_atoms\<^sub>\<nu>_def
by (metis Collect_mem_eq card_image_le image_Collect nested_prop_atoms_finite)
lemma nested_prop_atoms\<^sub>\<mu>_card:
"card (nested_prop_atoms\<^sub>\<mu> \<phi> X) \<le> card (nested_prop_atoms \<phi>)"
unfolding nested_prop_atoms\<^sub>\<mu>_def
by (metis Collect_mem_eq card_image_le image_Collect nested_prop_atoms_finite)
lemma GF_advice_nested_prop_atoms\<^sub>\<nu>:
"nested_prop_atoms (\<phi>[X]\<^sub>\<nu>) \<subseteq> nested_prop_atoms\<^sub>\<nu> \<phi> X"
by (induction \<phi>) (unfold nested_prop_atoms\<^sub>\<nu>_def, force+)
lemma FG_advice_nested_prop_atoms\<^sub>\<mu>:
"nested_prop_atoms (\<phi>[Y]\<^sub>\<mu>) \<subseteq> nested_prop_atoms\<^sub>\<mu> \<phi> Y"
by (induction \<phi>) (unfold nested_prop_atoms\<^sub>\<mu>_def, force+)
lemma nested_prop_atoms\<^sub>\<nu>_subset:
"nested_prop_atoms \<phi> \<subseteq> nested_prop_atoms \<psi> \<Longrightarrow> nested_prop_atoms\<^sub>\<nu> \<phi> X \<subseteq> nested_prop_atoms\<^sub>\<nu> \<psi> X"
unfolding nested_prop_atoms\<^sub>\<nu>_def by blast
lemma nested_prop_atoms\<^sub>\<mu>_subset:
"nested_prop_atoms \<phi> \<subseteq> nested_prop_atoms \<psi> \<Longrightarrow> nested_prop_atoms\<^sub>\<mu> \<phi> Y \<subseteq> nested_prop_atoms\<^sub>\<mu> \<psi> Y"
unfolding nested_prop_atoms\<^sub>\<mu>_def by blast
lemma GF_advice_nested_prop_atoms_card:
"card (nested_prop_atoms (\<phi>[X]\<^sub>\<nu>)) \<le> card (nested_prop_atoms \<phi>)"
proof -
have "card (nested_prop_atoms (\<phi>[X]\<^sub>\<nu>)) \<le> card (nested_prop_atoms\<^sub>\<nu> \<phi> X)"
by (simp add: nested_prop_atoms\<^sub>\<nu>_finite GF_advice_nested_prop_atoms\<^sub>\<nu> card_mono)
then show ?thesis
using nested_prop_atoms\<^sub>\<nu>_card le_trans by blast
qed
lemma FG_advice_nested_prop_atoms_card:
"card (nested_prop_atoms (\<phi>[Y]\<^sub>\<mu>)) \<le> card (nested_prop_atoms \<phi>)"
proof -
have "card (nested_prop_atoms (\<phi>[Y]\<^sub>\<mu>)) \<le> card (nested_prop_atoms\<^sub>\<mu> \<phi> Y)"
by (simp add: nested_prop_atoms\<^sub>\<mu>_finite FG_advice_nested_prop_atoms\<^sub>\<mu> card_mono)
then show ?thesis
using nested_prop_atoms\<^sub>\<mu>_card le_trans by blast
qed
subsection \<open>Intersecting the Advice Set\<close>
lemma GF_advice_inter:
"X \<inter> subformulas\<^sub>\<mu> \<phi> \<subseteq> S \<Longrightarrow> \<phi>[X \<inter> S]\<^sub>\<nu> = \<phi>[X]\<^sub>\<nu>"
by (induction \<phi>) auto
lemma GF_advice_inter_subformulas:
"\<phi>[X \<inter> subformulas\<^sub>\<mu> \<phi>]\<^sub>\<nu> = \<phi>[X]\<^sub>\<nu>"
using GF_advice_inter by blast
lemma GF_advice_minus_subformulas:
"\<psi> \<notin> subformulas\<^sub>\<mu> \<phi> \<Longrightarrow> \<phi>[X - {\<psi>}]\<^sub>\<nu> = \<phi>[X]\<^sub>\<nu>"
proof -
assume "\<psi> \<notin> subformulas\<^sub>\<mu> \<phi>"
then have "subformulas\<^sub>\<mu> \<phi> \<inter> X \<subseteq> X - {\<psi>}"
by blast
then show "\<phi>[X - {\<psi>}]\<^sub>\<nu> = \<phi>[X]\<^sub>\<nu>"
by (metis GF_advice_inter Diff_subset Int_absorb1 inf.commute)
qed
lemma GF_advice_minus_size:
"\<lbrakk>size \<phi> \<le> size \<psi>; \<phi> \<noteq> \<psi>\<rbrakk> \<Longrightarrow> \<phi>[X - {\<psi>}]\<^sub>\<nu> = \<phi>[X]\<^sub>\<nu>"
using subfrmlsn_size subformulas\<^sub>\<mu>_subfrmlsn GF_advice_minus_subformulas
by fastforce
lemma FG_advice_inter:
"Y \<inter> subformulas\<^sub>\<nu> \<phi> \<subseteq> S \<Longrightarrow> \<phi>[Y \<inter> S]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
by (induction \<phi>) auto
lemma FG_advice_inter_subformulas:
"\<phi>[Y \<inter> subformulas\<^sub>\<nu> \<phi>]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
using FG_advice_inter by blast
lemma FG_advice_minus_subformulas:
"\<psi> \<notin> subformulas\<^sub>\<nu> \<phi> \<Longrightarrow> \<phi>[Y - {\<psi>}]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
proof -
assume "\<psi> \<notin> subformulas\<^sub>\<nu> \<phi>"
then have "subformulas\<^sub>\<nu> \<phi> \<inter> Y \<subseteq> Y - {\<psi>}"
by blast
then show "\<phi>[Y - {\<psi>}]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
by (metis FG_advice_inter Diff_subset Int_absorb1 inf.commute)
qed
lemma FG_advice_minus_size:
"\<lbrakk>size \<phi> \<le> size \<psi>; \<phi> \<noteq> \<psi>\<rbrakk> \<Longrightarrow> \<phi>[Y - {\<psi>}]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
using subfrmlsn_size subformulas\<^sub>\<nu>_subfrmlsn FG_advice_minus_subformulas
by fastforce
lemma FG_advice_insert:
"\<lbrakk>\<psi> \<notin> Y; size \<phi> < size \<psi>\<rbrakk> \<Longrightarrow> \<phi>[insert \<psi> Y]\<^sub>\<mu> = \<phi>[Y]\<^sub>\<mu>"
by (metis FG_advice_minus_size Diff_insert_absorb less_imp_le neq_iff)
subsection \<open>Correctness GF-advice function\<close>
lemma GF_advice_a1:
"\<lbrakk>\<F> \<phi> w \<subseteq> X; w \<Turnstile>\<^sub>n \<phi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu>"
proof (induction \<phi> arbitrary: w)
case (Next_ltln \<phi>)
then show ?case
using \<F>_suffix by simp blast
next
case (Until_ltln \<phi>1 \<phi>2)
have "\<F> (\<phi>1 W\<^sub>n \<phi>2) w \<subseteq> \<F> (\<phi>1 U\<^sub>n \<phi>2) w"
by fastforce
then have "\<F> (\<phi>1 W\<^sub>n \<phi>2) w \<subseteq> X" and "w \<Turnstile>\<^sub>n \<phi>1 W\<^sub>n \<phi>2"
using Until_ltln.prems ltln_strong_to_weak by blast+
then have "w \<Turnstile>\<^sub>n \<phi>1[X]\<^sub>\<nu> W\<^sub>n \<phi>2[X]\<^sub>\<nu>"
using Until_ltln.IH
by simp (meson \<F>_suffix subset_trans sup.boundedE)
moreover
have "w \<Turnstile>\<^sub>n \<phi>1 U\<^sub>n \<phi>2"
using Until_ltln.prems by simp
then have "\<phi>1 U\<^sub>n \<phi>2 \<in> \<F> (\<phi>1 U\<^sub>n \<phi>2) w"
by force
then have "\<phi>1 U\<^sub>n \<phi>2 \<in> X"
using Until_ltln.prems by fast
ultimately show ?case
by auto
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
by simp (meson \<F>_suffix subset_trans sup.boundedE)
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
by simp (meson \<F>_suffix subset_trans sup.boundedE)
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
have "\<F> (\<phi>1 R\<^sub>n \<phi>2) w \<subseteq> \<F> (\<phi>1 M\<^sub>n \<phi>2) w"
by fastforce
then have "\<F> (\<phi>1 R\<^sub>n \<phi>2) w \<subseteq> X" and "w \<Turnstile>\<^sub>n \<phi>1 R\<^sub>n \<phi>2"
using StrongRelease_ltln.prems ltln_strong_to_weak by blast+
then have "w \<Turnstile>\<^sub>n \<phi>1[X]\<^sub>\<nu> R\<^sub>n \<phi>2[X]\<^sub>\<nu>"
using StrongRelease_ltln.IH
by simp (meson \<F>_suffix subset_trans sup.boundedE)
moreover
have "w \<Turnstile>\<^sub>n \<phi>1 M\<^sub>n \<phi>2"
using StrongRelease_ltln.prems by simp
then have "\<phi>1 M\<^sub>n \<phi>2 \<in> \<F> (\<phi>1 M\<^sub>n \<phi>2) w"
by force
then have "\<phi>1 M\<^sub>n \<phi>2 \<in> X"
using StrongRelease_ltln.prems by fast
ultimately show ?case
by auto
qed auto
lemma GF_advice_a2_helper:
"\<lbrakk>\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>); w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
proof (induction \<phi> arbitrary: w)
case (Next_ltln \<phi>)
then show ?case
unfolding GF_advice.simps semantics_ltln.simps(7)
using GF_suffix by blast
next
case (Until_ltln \<phi>1 \<phi>2)
then have "\<phi>1 U\<^sub>n \<phi>2 \<in> X"
using ccontr[of "\<phi>1 U\<^sub>n \<phi>2 \<in> X"] by force
then have "w \<Turnstile>\<^sub>n F\<^sub>n \<phi>2"
using Until_ltln.prems by fastforce
moreover
have "w \<Turnstile>\<^sub>n (\<phi>1 U\<^sub>n \<phi>2)[X]\<^sub>\<nu>"
using Until_ltln.prems by simp
then have "w \<Turnstile>\<^sub>n (\<phi>1[X]\<^sub>\<nu>) W\<^sub>n (\<phi>2[X]\<^sub>\<nu>)"
- unfolding GF_advice.simps using `\<phi>1 U\<^sub>n \<phi>2 \<in> X` by simp
+ unfolding GF_advice.simps using \<open>\<phi>1 U\<^sub>n \<phi>2 \<in> X\<close> by simp
then have "w \<Turnstile>\<^sub>n \<phi>1 W\<^sub>n \<phi>2"
unfolding GF_advice.simps semantics_ltln.simps(10)
by (metis GF_suffix Until_ltln.IH Until_ltln.prems(1))
ultimately show ?case
using ltln_weak_to_strong by blast
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
unfolding GF_advice.simps semantics_ltln.simps(9)
by (metis GF_suffix Release_ltln.IH Release_ltln.prems(1))
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
unfolding GF_advice.simps semantics_ltln.simps(10)
by (metis GF_suffix)
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then have "\<phi>1 M\<^sub>n \<phi>2 \<in> X"
using ccontr[of "\<phi>1 M\<^sub>n \<phi>2 \<in> X"] by force
then have "w \<Turnstile>\<^sub>n F\<^sub>n \<phi>1"
using StrongRelease_ltln.prems by fastforce
moreover
have "w \<Turnstile>\<^sub>n (\<phi>1 M\<^sub>n \<phi>2)[X]\<^sub>\<nu>"
using StrongRelease_ltln.prems by simp
then have "w \<Turnstile>\<^sub>n (\<phi>1[X]\<^sub>\<nu>) R\<^sub>n (\<phi>2[X]\<^sub>\<nu>)"
- unfolding GF_advice.simps using `\<phi>1 M\<^sub>n \<phi>2 \<in> X` by simp
+ unfolding GF_advice.simps using \<open>\<phi>1 M\<^sub>n \<phi>2 \<in> X\<close> by simp
then have "w \<Turnstile>\<^sub>n \<phi>1 R\<^sub>n \<phi>2"
unfolding GF_advice.simps semantics_ltln.simps(9)
by (metis GF_suffix StrongRelease_ltln.IH StrongRelease_ltln.prems(1))
ultimately show ?case
using ltln_weak_to_strong by blast
qed auto
lemma GF_advice_a2:
"\<lbrakk>X \<subseteq> \<G>\<F> \<phi> w; w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
by (metis GF_advice_a2_helper \<G>\<F>_elim subset_eq)
lemma GF_advice_a3:
"\<lbrakk>X = \<F> \<phi> w; X = \<G>\<F> \<phi> w\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu>"
using GF_advice_a1 GF_advice_a2 by fastforce
subsection \<open>Correctness FG-advice function\<close>
lemma FG_advice_b1:
"\<lbrakk>\<F>\<G> \<phi> w \<subseteq> Y; w \<Turnstile>\<^sub>n \<phi>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>"
proof (induction \<phi> arbitrary: w)
case (Next_ltln \<phi>)
then show ?case
using \<F>\<G>_suffix by simp blast
next
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
by simp (metis \<F>\<G>_suffix)
next
case (Release_ltln \<phi>1 \<phi>2)
show ?case
proof (cases "\<phi>1 R\<^sub>n \<phi>2 \<in> Y")
case False
then have "\<phi>1 R\<^sub>n \<phi>2 \<notin> \<F>\<G> (\<phi>1 R\<^sub>n \<phi>2) w"
using Release_ltln.prems by blast
then have "\<not> w \<Turnstile>\<^sub>n G\<^sub>n \<phi>2"
by fastforce
then have "w \<Turnstile>\<^sub>n \<phi>1 M\<^sub>n \<phi>2"
using Release_ltln.prems ltln_weak_to_strong by blast
moreover
have "\<F>\<G> (\<phi>1 M\<^sub>n \<phi>2) w \<subseteq> \<F>\<G> (\<phi>1 R\<^sub>n \<phi>2) w"
by fastforce
then have "\<F>\<G> (\<phi>1 M\<^sub>n \<phi>2) w \<subseteq> Y"
using Release_ltln.prems by blast
ultimately show ?thesis
using Release_ltln.IH by simp (metis \<F>\<G>_suffix)
qed simp
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
show ?case
proof (cases "\<phi>1 W\<^sub>n \<phi>2 \<in> Y")
case False
then have "\<phi>1 W\<^sub>n \<phi>2 \<notin> \<F>\<G> (\<phi>1 W\<^sub>n \<phi>2) w"
using WeakUntil_ltln.prems by blast
then have "\<not> w \<Turnstile>\<^sub>n G\<^sub>n \<phi>1"
by fastforce
then have "w \<Turnstile>\<^sub>n \<phi>1 U\<^sub>n \<phi>2"
using WeakUntil_ltln.prems ltln_weak_to_strong by blast
moreover
have "\<F>\<G> (\<phi>1 U\<^sub>n \<phi>2) w \<subseteq> \<F>\<G> (\<phi>1 W\<^sub>n \<phi>2) w"
by fastforce
then have "\<F>\<G> (\<phi>1 U\<^sub>n \<phi>2) w \<subseteq> Y"
using WeakUntil_ltln.prems by blast
ultimately show ?thesis
using WeakUntil_ltln.IH by simp (metis \<F>\<G>_suffix)
qed simp
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then show ?case
by simp (metis \<F>\<G>_suffix)
qed auto
lemma FG_advice_b2_helper:
"\<lbrakk>\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n G\<^sub>n \<psi>; w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
proof (induction \<phi> arbitrary: w)
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
by simp (metis (no_types, lifting) suffix_suffix)
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
proof (cases "\<phi>1 R\<^sub>n \<phi>2 \<in> Y")
case True
then show ?thesis
using Release_ltln.prems by force
next
case False
then have "w \<Turnstile>\<^sub>n (\<phi>1[Y]\<^sub>\<mu>) M\<^sub>n (\<phi>2[Y]\<^sub>\<mu>)"
using Release_ltln.prems by simp
then have "w \<Turnstile>\<^sub>n \<phi>1 M\<^sub>n \<phi>2"
using Release_ltln
by simp (metis (no_types, lifting) suffix_suffix)
then show ?thesis
using ltln_strong_to_weak by fast
qed
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
proof (cases "\<phi>1 W\<^sub>n \<phi>2 \<in> Y")
case True
then show ?thesis
using WeakUntil_ltln.prems by force
next
case False
then have "w \<Turnstile>\<^sub>n (\<phi>1[Y]\<^sub>\<mu>) U\<^sub>n (\<phi>2[Y]\<^sub>\<mu>)"
using WeakUntil_ltln.prems by simp
then have "w \<Turnstile>\<^sub>n \<phi>1 U\<^sub>n \<phi>2"
using WeakUntil_ltln
by simp (metis (no_types, lifting) suffix_suffix)
then show ?thesis
using ltln_strong_to_weak by fast
qed
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then show ?case
by simp (metis (no_types, lifting) suffix_suffix)
qed auto
lemma FG_advice_b2:
"\<lbrakk>Y \<subseteq> \<G> \<phi> w; w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
by (metis FG_advice_b2_helper \<G>_elim subset_eq)
lemma FG_advice_b3:
"\<lbrakk>Y = \<F>\<G> \<phi> w; Y = \<G> \<phi> w\<rbrakk> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>"
using FG_advice_b1 FG_advice_b2 by fastforce
subsection \<open>Advice Functions and the ``after'' Function\<close>
lemma GF_advice_af_letter:
"(x ## w) \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> w \<Turnstile>\<^sub>n (af_letter \<phi> x)[X]\<^sub>\<nu>"
proof (induction \<phi>)
case (Until_ltln \<phi>1 \<phi>2)
then have "w \<Turnstile>\<^sub>n af_letter ((\<phi>1 U\<^sub>n \<phi>2)[X]\<^sub>\<nu>) x"
using af_letter_build by blast
then show ?case
using Until_ltln.IH af_letter_build by fastforce
next
case (Release_ltln \<phi>1 \<phi>2)
then have "w \<Turnstile>\<^sub>n af_letter ((\<phi>1 R\<^sub>n \<phi>2)[X]\<^sub>\<nu>) x"
using af_letter_build by blast
then show ?case
using Release_ltln.IH af_letter_build by auto
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then have "w \<Turnstile>\<^sub>n af_letter ((\<phi>1 W\<^sub>n \<phi>2)[X]\<^sub>\<nu>) x"
using af_letter_build by blast
then show ?case
using WeakUntil_ltln.IH af_letter_build by auto
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then have "w \<Turnstile>\<^sub>n af_letter ((\<phi>1 M\<^sub>n \<phi>2)[X]\<^sub>\<nu>) x"
using af_letter_build by blast
then show ?case
using StrongRelease_ltln.IH af_letter_build by force
qed auto
lemma FG_advice_af_letter:
"w \<Turnstile>\<^sub>n (af_letter \<phi> x)[Y]\<^sub>\<mu> \<Longrightarrow> (x ## w) \<Turnstile>\<^sub>n \<phi>[Y]\<^sub>\<mu>"
proof (induction \<phi>)
case (Prop_ltln a)
then show ?case
using semantics_ltln.simps(3) by fastforce
next
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
unfolding af_letter.simps FG_advice.simps semantics_ltln.simps(5,6)
using af_letter_build apply (cases "w \<Turnstile>\<^sub>n af_letter \<phi>2 x[Y]\<^sub>\<mu>") apply force
by (metis af_letter.simps(8) semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
apply (cases "\<phi>1 R\<^sub>n \<phi>2 \<in> Y")
apply simp
unfolding af_letter.simps FG_advice.simps semantics_ltln.simps(5,6)
using af_letter_build apply (cases "w \<Turnstile>\<^sub>n af_letter \<phi>1 x[Y]\<^sub>\<mu>") apply force
by (metis (full_types) af_letter.simps(11) semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
apply (cases "\<phi>1 W\<^sub>n \<phi>2 \<in> Y")
apply simp
unfolding af_letter.simps FG_advice.simps semantics_ltln.simps(5,6)
using af_letter_build apply (cases "w \<Turnstile>\<^sub>n af_letter \<phi>2 x[Y]\<^sub>\<mu>") apply force
by (metis (full_types) af_letter.simps(8) semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then show ?case
unfolding af_letter.simps FG_advice.simps semantics_ltln.simps(5,6)
using af_letter_build apply (cases "w \<Turnstile>\<^sub>n af_letter \<phi>1 x[Y]\<^sub>\<mu>") apply force
by (metis af_letter.simps(11) semantics_ltln.simps(5) semantics_ltln.simps(6))
qed auto
lemma GF_advice_af:
"(w \<frown> w') \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> w' \<Turnstile>\<^sub>n (af \<phi> w)[X]\<^sub>\<nu>"
by (induction w arbitrary: \<phi>) (simp, insert GF_advice_af_letter, fastforce)
lemma FG_advice_af:
"w' \<Turnstile>\<^sub>n (af \<phi> w)[X]\<^sub>\<mu> \<Longrightarrow> (w \<frown> w') \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<mu>"
by (induction w arbitrary: \<phi>) (simp, insert FG_advice_af_letter, fastforce)
lemma GF_advice_af_2:
"w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> suffix i w \<Turnstile>\<^sub>n (af \<phi> (prefix i w))[X]\<^sub>\<nu>"
using GF_advice_af by force
lemma FG_advice_af_2:
"suffix i w \<Turnstile>\<^sub>n (af \<phi> (prefix i w))[X]\<^sub>\<mu> \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<mu>"
using FG_advice_af by force
(* TODO move to Omega_Words_Fun.thy ?? *)
lemma prefix_suffix_subsequence: "prefix i (suffix j w) = (w [j \<rightarrow> i + j])"
by (simp add: add.commute)
text \<open>We show this generic lemma to prove the following theorems:\<close>
lemma GF_advice_sync:
fixes index :: "nat \<Rightarrow> nat"
fixes formula :: "nat \<Rightarrow> 'a ltln"
assumes "\<And>i. i < n \<Longrightarrow> \<exists>j. suffix ((index i) + j) w \<Turnstile>\<^sub>n af (formula i) (w [index i \<rightarrow> (index i) + j])[X]\<^sub>\<nu>"
shows "\<exists>k. (\<forall>i < n. k \<ge> index i \<and> suffix k w \<Turnstile>\<^sub>n af (formula i) (w [index i \<rightarrow> k])[X]\<^sub>\<nu>)"
using assms
proof (induction n)
case (Suc n)
obtain k1 where leq1: "\<And>i. i < n \<Longrightarrow> k1 \<ge> index i"
and suffix1: "\<And>i. i < n \<Longrightarrow> suffix k1 w \<Turnstile>\<^sub>n af (formula i) (w [(index i) \<rightarrow> k1])[X]\<^sub>\<nu>"
using Suc less_SucI by blast
obtain k2 where leq2: "k2 \<ge> index n"
and suffix2: "suffix k2 w \<Turnstile>\<^sub>n af (formula n) (w [index n \<rightarrow> k2])[X]\<^sub>\<nu>"
using le_add1 Suc.prems by blast
define k where "k \<equiv> k1 + k2"
have "\<And>i. i < Suc n \<Longrightarrow> k \<ge> index i"
unfolding k_def by (metis leq1 leq2 less_SucE trans_le_add1 trans_le_add2)
moreover
{
have "\<And>i. i < n \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af (formula i) (w [(index i) \<rightarrow> k])[X]\<^sub>\<nu>"
unfolding k_def
by (metis GF_advice_af_2[OF suffix1, unfolded suffix_suffix prefix_suffix_subsequence] af_subsequence_append leq1 add.commute le_add1)
moreover
have "suffix k w \<Turnstile>\<^sub>n af (formula n) (w [index n \<rightarrow> k])[X]\<^sub>\<nu>"
unfolding k_def
by (metis GF_advice_af_2[OF suffix2, unfolded suffix_suffix prefix_suffix_subsequence] af_subsequence_append leq2 add.commute le_add1)
ultimately
have "\<And>i. i \<le> n \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af (formula i) (w [(index i) \<rightarrow> k])[X]\<^sub>\<nu>"
using nat_less_le by blast
}
ultimately
show ?case
by (meson less_Suc_eq_le)
qed simp
lemma GF_advice_sync_and:
assumes "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
assumes "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<psi> (prefix i w)[X]\<^sub>\<nu>"
shows "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu> \<and> suffix i w \<Turnstile>\<^sub>n af \<psi> (prefix i w)[X]\<^sub>\<nu>"
proof -
let ?formula = "\<lambda>i :: nat. (if (i = 0) then \<phi> else \<psi>)"
have assms: "\<And>i. i < 2 \<Longrightarrow> \<exists>j. suffix j w \<Turnstile>\<^sub>n af (?formula i) (w [0 \<rightarrow> j])[X]\<^sub>\<nu>"
using assms by simp
obtain k where k_def: "\<And>i :: nat. i < 2 \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af (if i = 0 then \<phi> else \<psi>) (prefix k w)[X]\<^sub>\<nu>"
using GF_advice_sync[of "2" "\<lambda>i. 0" w ?formula, simplified, OF assms, simplified] by blast
show ?thesis
using k_def[of 0] k_def[of 1] by auto
qed
lemma GF_advice_sync_less:
assumes "\<And>i. i < n \<Longrightarrow> \<exists>j. suffix (i + j) w \<Turnstile>\<^sub>n af \<phi> (w [i \<rightarrow> j + i])[X]\<^sub>\<nu>"
assumes "\<exists>j. suffix (n + j) w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> j + n])[X]\<^sub>\<nu>"
shows "\<exists>k \<ge> n. (\<forall>j < n. suffix k w \<Turnstile>\<^sub>n af \<phi> (w [j \<rightarrow> k])[X]\<^sub>\<nu>) \<and> suffix k w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> k])[X]\<^sub>\<nu>"
proof -
let ?index = "\<lambda>i. min i n"
let ?formula = "\<lambda>i. if (i < n) then \<phi> else \<psi>"
{
fix i
assume "i < Suc n"
then have min_def: "min i n = i"
by simp
have "\<exists>j. suffix ((?index i) + j) w \<Turnstile>\<^sub>n af (?formula i) (w [?index i \<rightarrow> (?index i) + j])[X]\<^sub>\<nu>"
unfolding min_def
by (cases "i < n")
(metis (full_types) assms(1) add.commute, metis (full_types) assms(2) \<open>i < Suc n\<close> add.commute less_SucE)
}
then obtain k where leq: "(\<And>i. i < Suc n \<Longrightarrow> min i n \<le> k)"
and suffix: "\<And>i. i < Suc n \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af (if i < n then \<phi> else \<psi>) (w [min i n \<rightarrow> k])[X]\<^sub>\<nu>"
using GF_advice_sync[of "Suc n" ?index w ?formula X] by metis
have "\<forall>j < n. suffix k w \<Turnstile>\<^sub>n af \<phi> (w [j \<rightarrow> k])[X]\<^sub>\<nu>"
using suffix by (metis (full_types) less_SucI min.strict_order_iff)
moreover
have "suffix k w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> k])[X]\<^sub>\<nu>"
using suffix[of n, simplified] by blast
moreover
have "k \<ge> n"
using leq by presburger
ultimately
show ?thesis
by auto
qed
lemma GF_advice_sync_lesseq:
assumes "\<And>i. i \<le> n \<Longrightarrow> \<exists>j. suffix (i + j) w \<Turnstile>\<^sub>n af \<phi> (w [i \<rightarrow> j + i])[X]\<^sub>\<nu>"
assumes "\<exists>j. suffix (n + j) w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> j + n])[X]\<^sub>\<nu>"
shows "\<exists>k \<ge> n. (\<forall>j \<le> n. suffix k w \<Turnstile>\<^sub>n af \<phi> (w [j \<rightarrow> k])[X]\<^sub>\<nu>) \<and> suffix k w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> k])[X]\<^sub>\<nu>"
proof -
let ?index = "\<lambda>i. min i n"
let ?formula = "\<lambda>i. if (i \<le> n) then \<phi> else \<psi>"
{
fix i
assume "i < Suc (Suc n)"
hence "\<exists>j. suffix ((?index i) + j) w \<Turnstile>\<^sub>n af (?formula i) (w [?index i \<rightarrow> (?index i) + j])[X]\<^sub>\<nu>"
proof (cases "i < Suc n")
case True
then have min_def: "min i n = i"
by simp
show ?thesis
unfolding min_def by (metis (full_types) assms(1) Suc_leI Suc_le_mono True add.commute)
next
case False
then have i_def: "i = Suc n"
using \<open>i < Suc (Suc n)\<close> less_antisym by blast
have min_def: "min i n = n"
unfolding i_def by simp
show ?thesis
using assms(2) False
by (simp add: min_def add.commute)
qed
}
then obtain k where leq: "(\<And>i. i \<le> Suc n \<Longrightarrow> min i n \<le> k)"
and suffix: "\<And>i :: nat. i < Suc (Suc n) \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af (if i \<le> n then \<phi> else \<psi>) (w [min i n \<rightarrow> k])[X]\<^sub>\<nu>"
using GF_advice_sync[of "Suc (Suc n)" ?index w ?formula X]
by (metis (no_types, hide_lams) less_Suc_eq min_le_iff_disj)
have "\<forall>j \<le> n. suffix k w \<Turnstile>\<^sub>n af \<phi> (w [j \<rightarrow> k])[X]\<^sub>\<nu>"
using suffix by (metis (full_types) le_SucI less_Suc_eq_le min.orderE)
moreover
have "suffix k w \<Turnstile>\<^sub>n af \<psi> (w [n \<rightarrow> k])[X]\<^sub>\<nu>"
using suffix[of "Suc n", simplified] by linarith
moreover
have "k \<ge> n"
using leq by presburger
ultimately
show ?thesis
by auto
qed
lemma af_subsequence_U_GF_advice:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [i \<rightarrow> n]))[X]\<^sub>\<nu>)"
assumes "\<And>j. j < i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<phi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> U\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
then have A: "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [0 \<rightarrow> n]))[X]\<^sub>\<nu>)"
by blast
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF A, of 1] by simp
then show ?case
unfolding GF_advice.simps af_subsequence_U semantics_ltln.simps by blast
next
case (Suc i)
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<phi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[OF zero_less_Suc, THEN GF_advice_af_2, unfolded suffix_suffix, of 1]
by simp
moreover
have B: "(Suc (n - 1)) = n"
using Suc by simp
note Suc.IH[of "n - 1" "suffix 1 w", unfolded suffix_suffix] Suc.prems
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> U\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono Suc_mono plus_1_eq_Suc subsequence_shift)
ultimately
show ?case
unfolding af_subsequence_U semantics_ltln.simps GF_advice.simps by blast
qed
lemma af_subsequence_M_GF_advice:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n ((af \<phi> (w [i \<rightarrow> n]))[X]\<^sub>\<nu>)"
assumes "\<And>j. j \<le> i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> M\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
then have A: "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [0 \<rightarrow> n]))[X]\<^sub>\<nu>)"
by blast
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF A, of 1] by simp
moreover
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<phi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF "0.prems"(2), of 1, unfolded suffix_suffix] by auto
ultimately
show ?case
unfolding af_subsequence_M GF_advice.simps semantics_ltln.simps by blast
next
case (Suc i)
have "suffix 1 (suffix n w) \<Turnstile>\<^sub>n af (af \<psi> (prefix n w)) [suffix n w 0][X]\<^sub>\<nu>"
by (metis (no_types) GF_advice_af_2 Suc.prems(3) plus_1_eq_Suc subsequence_singleton suffix_0 suffix_suffix zero_le)
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[THEN GF_advice_af_2, unfolded suffix_suffix, of 1] by simp
moreover
have B: "(Suc (n - 1)) = n"
using Suc by simp
note Suc.IH[of _ "suffix 1 w", unfolded subsequence_shift suffix_suffix]
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> M\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono plus_1_eq_Suc Suc.prems)
ultimately
show ?case
unfolding af_subsequence_M semantics_ltln.simps GF_advice.simps by blast
qed
lemma af_subsequence_R_GF_advice:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n ((af \<phi> (w [i \<rightarrow> n]))[X]\<^sub>\<nu>)"
assumes "\<And>j. j \<le> i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> R\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
then have A: "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [0 \<rightarrow> n]))[X]\<^sub>\<nu>)"
by blast
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF A, of 1] by simp
moreover
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<phi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF "0.prems"(2), of 1, unfolded suffix_suffix] by auto
ultimately
show ?case
unfolding af_subsequence_R GF_advice.simps semantics_ltln.simps by blast
next
case (Suc i)
have "suffix 1 (suffix n w) \<Turnstile>\<^sub>n af (af \<psi> (prefix n w)) [suffix n w 0][X]\<^sub>\<nu>"
by (metis (no_types) GF_advice_af_2 Suc.prems(3) plus_1_eq_Suc subsequence_singleton suffix_0 suffix_suffix zero_le)
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[THEN GF_advice_af_2, unfolded suffix_suffix, of 1] by simp
moreover
have B: "(Suc (n - 1)) = n"
using Suc by simp
note Suc.IH[of _ "suffix 1 w", unfolded subsequence_shift suffix_suffix]
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> R\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono plus_1_eq_Suc Suc.prems)
ultimately
show ?case
unfolding af_subsequence_R semantics_ltln.simps GF_advice.simps by blast
qed
lemma af_subsequence_W_GF_advice:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [i \<rightarrow> n]))[X]\<^sub>\<nu>)"
assumes "\<And>j. j < i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<phi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> W\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
then have A: "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [0 \<rightarrow> n]))[X]\<^sub>\<nu>)"
by blast
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF A, of 1] by simp
then show ?case
unfolding af_subsequence_W GF_advice.simps semantics_ltln.simps by blast
next
case (Suc i)
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<phi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[OF zero_less_Suc, THEN GF_advice_af_2, unfolded suffix_suffix, of 1]
by simp
moreover
have B: "(Suc (n - 1)) = n"
using Suc by simp
note Suc.IH[of "n - 1" "suffix 1 w", unfolded suffix_suffix] Suc.prems
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> W\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono Suc_mono plus_1_eq_Suc subsequence_shift)
ultimately
show ?case
unfolding af_subsequence_W unfolding semantics_ltln.simps GF_advice.simps by simp
qed
lemma af_subsequence_R_GF_advice_connect:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n af (\<phi> R\<^sub>n \<psi>) (w [i \<rightarrow> n])[X]\<^sub>\<nu>"
assumes "\<And>j. j \<le> i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> R\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
then have A: "suffix n w \<Turnstile>\<^sub>n ((af \<psi> (w [0 \<rightarrow> n]))[X]\<^sub>\<nu>)"
by blast
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF A, of 1] by simp
moreover
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> R\<^sub>n \<psi>) (w [0 \<rightarrow> Suc n]))[X]\<^sub>\<nu>"
using GF_advice_af_2[OF "0.prems"(2), of 1, unfolded suffix_suffix] by simp
ultimately
show ?case
unfolding af_subsequence_R GF_advice.simps semantics_ltln.simps by blast
next
case (Suc i)
have "suffix 1 (suffix n w) \<Turnstile>\<^sub>n af (af \<psi> (prefix n w)) [suffix n w 0][X]\<^sub>\<nu>"
by (metis (no_types) GF_advice_af_2 Suc.prems(3) plus_1_eq_Suc subsequence_singleton suffix_0 suffix_suffix zero_le)
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<psi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[THEN GF_advice_af_2, unfolded suffix_suffix, of 1] by simp
moreover
have B: "(Suc (n - 1)) = n"
using Suc by simp
note Suc.IH[of _ "suffix 1 w", unfolded subsequence_shift suffix_suffix]
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> R\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono plus_1_eq_Suc Suc.prems)
ultimately
show ?case
unfolding af_subsequence_R semantics_ltln.simps GF_advice.simps by blast
qed
lemma af_subsequence_W_GF_advice_connect:
assumes "i \<le> n"
assumes "suffix n w \<Turnstile>\<^sub>n af (\<phi> W\<^sub>n \<psi>) (w [i \<rightarrow> n])[X]\<^sub>\<nu>"
assumes "\<And>j. j < i \<Longrightarrow> suffix n w \<Turnstile>\<^sub>n ((af \<phi> (w [j \<rightarrow> n]))[X]\<^sub>\<nu>)"
shows "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> W\<^sub>n \<psi>) (prefix (Suc n) w))[X]\<^sub>\<nu>"
using assms
proof (induction i arbitrary: w n)
case 0
have "suffix (Suc n) w \<Turnstile>\<^sub>n af_letter (af (\<phi> W\<^sub>n \<psi>) (prefix n w)) (w n)[X]\<^sub>\<nu>"
by (simp add: "0.prems"(2) GF_advice_af_letter)
moreover
have "prefix (Suc n) w = prefix n w @ [w n]"
using subseq_to_Suc by blast
ultimately show ?case
by (metis (no_types) foldl.simps(1) foldl.simps(2) foldl_append)
next
case (Suc i)
have "suffix (Suc n) w \<Turnstile>\<^sub>n (af \<phi> (prefix (Suc n) w))[X]\<^sub>\<nu>"
using Suc.prems(3)[OF zero_less_Suc, THEN GF_advice_af_2, unfolded suffix_suffix, of 1] by simp
moreover
have "n > 0" and B: "(Suc (n - 1)) = n"
using Suc by simp+
note Suc.IH[of "n - 1" "suffix 1 w", unfolded suffix_suffix] Suc.prems
then have "suffix (Suc n) w \<Turnstile>\<^sub>n (af (\<phi> W\<^sub>n \<psi>) (w [1 \<rightarrow> (Suc n)]))[X]\<^sub>\<nu>"
by (metis B One_nat_def Suc_le_mono Suc_mono plus_1_eq_Suc subsequence_shift)
ultimately
show ?case
unfolding af_subsequence_W unfolding semantics_ltln.simps GF_advice.simps by simp
qed
subsection \<open>Advice Functions and Propositional Entailment\<close>
lemma GF_advice_prop_entailment:
"\<A> \<Turnstile>\<^sub>P \<phi>[X]\<^sub>\<nu> \<Longrightarrow> {\<psi>. \<psi>[X]\<^sub>\<nu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi>"
"false\<^sub>n \<notin> \<A> \<Longrightarrow> {\<psi>. \<psi>[X]\<^sub>\<nu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>[X]\<^sub>\<nu>"
by (induction \<phi>) (auto, meson, meson)
lemma GF_advice_iff_prop_entailment:
"false\<^sub>n \<notin> \<A> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>[X]\<^sub>\<nu> \<longleftrightarrow> {\<psi>. \<psi>[X]\<^sub>\<nu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi>"
by (metis GF_advice_prop_entailment)
lemma FG_advice_prop_entailment:
"true\<^sub>n \<in> \<A> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>[Y]\<^sub>\<mu> \<Longrightarrow> {\<psi>. \<psi>[Y]\<^sub>\<mu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi>"
"{\<psi>. \<psi>[Y]\<^sub>\<mu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>[Y]\<^sub>\<mu>"
by (induction \<phi>) auto
lemma FG_advice_iff_prop_entailment:
"true\<^sub>n \<in> \<A> \<Longrightarrow> \<A> \<Turnstile>\<^sub>P \<phi>[X]\<^sub>\<mu> \<longleftrightarrow> {\<psi>. \<psi>[X]\<^sub>\<mu> \<in> \<A>} \<Turnstile>\<^sub>P \<phi>"
by (metis FG_advice_prop_entailment)
lemma GF_advice_subst:
"\<phi>[X]\<^sub>\<nu> = subst \<phi> (\<lambda>\<psi>. Some (\<psi>[X]\<^sub>\<nu>))"
by (induction \<phi>) auto
lemma FG_advice_subst:
"\<phi>[X]\<^sub>\<mu> = subst \<phi> (\<lambda>\<psi>. Some (\<psi>[X]\<^sub>\<mu>))"
by (induction \<phi>) auto
lemma GF_advice_prop_congruent:
"\<phi> \<longrightarrow>\<^sub>P \<psi> \<Longrightarrow> \<phi>[X]\<^sub>\<nu> \<longrightarrow>\<^sub>P \<psi>[X]\<^sub>\<nu>"
"\<phi> \<sim>\<^sub>P \<psi> \<Longrightarrow> \<phi>[X]\<^sub>\<nu> \<sim>\<^sub>P \<psi>[X]\<^sub>\<nu>"
by (metis GF_advice_subst subst_respects_ltl_prop_entailment)+
lemma FG_advice_prop_congruent:
"\<phi> \<longrightarrow>\<^sub>P \<psi> \<Longrightarrow> \<phi>[X]\<^sub>\<mu> \<longrightarrow>\<^sub>P \<psi>[X]\<^sub>\<mu>"
"\<phi> \<sim>\<^sub>P \<psi> \<Longrightarrow> \<phi>[X]\<^sub>\<mu> \<sim>\<^sub>P \<psi>[X]\<^sub>\<mu>"
by (metis FG_advice_subst subst_respects_ltl_prop_entailment)+
subsection \<open>GF-advice with Equivalence Relations\<close>
locale GF_advice_congruent = ltl_equivalence +
fixes
normalise :: "'a ltln \<Rightarrow> 'a ltln"
assumes
normalise_eq: "\<phi> \<sim> normalise \<phi>"
assumes
normalise_monotonic: "w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> w \<Turnstile>\<^sub>n (normalise \<phi>)[X]\<^sub>\<nu>"
assumes
normalise_eventually_equivalent:
"w \<Turnstile>\<^sub>n (normalise \<phi>)[X]\<^sub>\<nu> \<Longrightarrow> (\<exists>i. suffix i w \<Turnstile>\<^sub>n (af \<phi> (prefix i w))[X]\<^sub>\<nu>)"
assumes
GF_advice_congruent: "\<phi> \<sim> \<psi> \<Longrightarrow> (normalise \<phi>)[X]\<^sub>\<nu> \<sim> (normalise \<psi>)[X]\<^sub>\<nu>"
begin
lemma normalise_language_equivalent[simp]:
"w \<Turnstile>\<^sub>n normalise \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>"
using normalise_eq ltl_lang_equiv_def eq_implies_lang by blast
end
interpretation prop_GF_advice_compatible: GF_advice_congruent "(\<sim>\<^sub>P)" "id"
by unfold_locales (simp add: GF_advice_af GF_advice_prop_congruent(2))+
end
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/After.thy b/thys/LTL_Master_Theorem/Logical_Characterization/After.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/After.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/After.thy
@@ -1,719 +1,719 @@
(*
Author: Benedikt Seidl
License: BSD
*)
section \<open>The ``after''-Function\<close>
theory After
imports
LTL.LTL LTL.Equivalence_Relations Syntactic_Fragments_and_Stability
begin
subsection \<open>Definition of af\<close>
primrec af_letter :: "'a ltln \<Rightarrow> 'a set \<Rightarrow> 'a ltln"
where
"af_letter true\<^sub>n \<nu> = true\<^sub>n"
| "af_letter false\<^sub>n \<nu> = false\<^sub>n"
| "af_letter prop\<^sub>n(a) \<nu> = (if a \<in> \<nu> then true\<^sub>n else false\<^sub>n)"
| "af_letter nprop\<^sub>n(a) \<nu> = (if a \<notin> \<nu> then true\<^sub>n else false\<^sub>n)"
| "af_letter (\<phi> and\<^sub>n \<psi>) \<nu> = (af_letter \<phi> \<nu>) and\<^sub>n (af_letter \<psi> \<nu>)"
| "af_letter (\<phi> or\<^sub>n \<psi>) \<nu> = (af_letter \<phi> \<nu>) or\<^sub>n (af_letter \<psi> \<nu>)"
| "af_letter (X\<^sub>n \<phi>) \<nu> = \<phi>"
| "af_letter (\<phi> U\<^sub>n \<psi>) \<nu> = (af_letter \<psi> \<nu>) or\<^sub>n ((af_letter \<phi> \<nu>) and\<^sub>n (\<phi> U\<^sub>n \<psi>))"
| "af_letter (\<phi> R\<^sub>n \<psi>) \<nu> = (af_letter \<psi> \<nu>) and\<^sub>n ((af_letter \<phi> \<nu>) or\<^sub>n (\<phi> R\<^sub>n \<psi>))"
| "af_letter (\<phi> W\<^sub>n \<psi>) \<nu> = (af_letter \<psi> \<nu>) or\<^sub>n ((af_letter \<phi> \<nu>) and\<^sub>n (\<phi> W\<^sub>n \<psi>))"
| "af_letter (\<phi> M\<^sub>n \<psi>) \<nu> = (af_letter \<psi> \<nu>) and\<^sub>n ((af_letter \<phi> \<nu>) or\<^sub>n (\<phi> M\<^sub>n \<psi>))"
abbreviation af :: "'a ltln \<Rightarrow> 'a set list \<Rightarrow> 'a ltln"
where
"af \<phi> w \<equiv> foldl af_letter \<phi> w"
lemma af_decompose:
"af (\<phi> and\<^sub>n \<psi>) w = (af \<phi> w) and\<^sub>n (af \<psi> w)"
"af (\<phi> or\<^sub>n \<psi>) w = (af \<phi> w) or\<^sub>n (af \<psi> w)"
by (induction w rule: rev_induct) simp_all
lemma af_simps[simp]:
"af true\<^sub>n w = true\<^sub>n"
"af false\<^sub>n w = false\<^sub>n"
"af (X\<^sub>n \<phi>) (x # xs) = af \<phi> xs"
by (induction w) simp_all
lemma af_ite_simps[simp]:
"af (if P then true\<^sub>n else false\<^sub>n) w = (if P then true\<^sub>n else false\<^sub>n)"
"af (if P then false\<^sub>n else true\<^sub>n) w = (if P then false\<^sub>n else true\<^sub>n)"
by simp_all
lemma af_subsequence_append:
"i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> af (af \<phi> (w [i \<rightarrow> j])) (w [j \<rightarrow> k]) = af \<phi> (w [i \<rightarrow> k])"
by (metis foldl_append le_Suc_ex map_append subsequence_def upt_add_eq_append)
lemma af_subsequence_U:
"af (\<phi> U\<^sub>n \<psi>) (w [0 \<rightarrow> Suc n]) = (af \<psi> (w [0 \<rightarrow> Suc n])) or\<^sub>n ((af \<phi> (w [0 \<rightarrow> Suc n])) and\<^sub>n af (\<phi> U\<^sub>n \<psi>) (w [1 \<rightarrow> Suc n]))"
by (induction n) fastforce+
lemma af_subsequence_U':
"af (\<phi> U\<^sub>n \<psi>) (a # xs) = (af \<psi> (a # xs)) or\<^sub>n ((af \<phi> (a # xs)) and\<^sub>n af (\<phi> U\<^sub>n \<psi>) xs)"
by (simp add: af_decompose)
lemma af_subsequence_R:
"af (\<phi> R\<^sub>n \<psi>) (w [0 \<rightarrow> Suc n]) = (af \<psi> (w [0 \<rightarrow> Suc n])) and\<^sub>n ((af \<phi> (w [0 \<rightarrow> Suc n])) or\<^sub>n af (\<phi> R\<^sub>n \<psi>) (w [1 \<rightarrow> Suc n]))"
by (induction n) fastforce+
lemma af_subsequence_R':
"af (\<phi> R\<^sub>n \<psi>) (a # xs) = (af \<psi> (a # xs)) and\<^sub>n ((af \<phi> (a # xs)) or\<^sub>n af (\<phi> R\<^sub>n \<psi>) xs)"
by (simp add: af_decompose)
lemma af_subsequence_W:
"af (\<phi> W\<^sub>n \<psi>) (w [0 \<rightarrow> Suc n]) = (af \<psi> (w [0 \<rightarrow> Suc n])) or\<^sub>n ((af \<phi> (w [0 \<rightarrow> Suc n])) and\<^sub>n af (\<phi> W\<^sub>n \<psi>) (w [1 \<rightarrow> Suc n]))"
by (induction n) fastforce+
lemma af_subsequence_W':
"af (\<phi> W\<^sub>n \<psi>) (a # xs) = (af \<psi> (a # xs)) or\<^sub>n ((af \<phi> (a # xs)) and\<^sub>n af (\<phi> W\<^sub>n \<psi>) xs)"
by (simp add: af_decompose)
lemma af_subsequence_M:
"af (\<phi> M\<^sub>n \<psi>) (w [0 \<rightarrow> Suc n]) = (af \<psi> (w [0 \<rightarrow> Suc n])) and\<^sub>n ((af \<phi> (w [0 \<rightarrow> Suc n])) or\<^sub>n af (\<phi> M\<^sub>n \<psi>) (w [1 \<rightarrow> Suc n]))"
by (induction n) fastforce+
lemma af_subsequence_M':
"af (\<phi> M\<^sub>n \<psi>) (a # xs) = (af \<psi> (a # xs)) and\<^sub>n ((af \<phi> (a # xs)) or\<^sub>n af (\<phi> M\<^sub>n \<psi>) xs)"
by (simp add: af_decompose)
lemma suffix_build[simp]:
"suffix (Suc n) (x ## xs) = suffix n xs"
by fastforce
lemma af_letter_build:
"(x ## w) \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n af_letter \<phi> x"
proof (induction \<phi> arbitrary: x w)
case (Until_ltln \<phi> \<psi>)
then show ?case
unfolding ltln_expand_Until by force
next
case (Release_ltln \<phi> \<psi>)
then show ?case
unfolding ltln_expand_Release by force
next
case (WeakUntil_ltln \<phi> \<psi>)
then show ?case
unfolding ltln_expand_WeakUntil by force
next
case (StrongRelease_ltln \<phi> \<psi>)
then show ?case
unfolding ltln_expand_StrongRelease by force
qed simp+
lemma af_ltl_continuation:
"(w \<frown> w') \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> w' \<Turnstile>\<^sub>n af \<phi> w"
proof (induction w arbitrary: \<phi> w')
case (Cons x xs)
then show ?case
using af_letter_build by fastforce
qed simp
subsection \<open>Range of the after function\<close>
lemma af_letter_atoms:
"atoms_ltln (af_letter \<phi> \<nu>) \<subseteq> atoms_ltln \<phi>"
by (induction \<phi>) auto
lemma af_atoms:
"atoms_ltln (af \<phi> w) \<subseteq> atoms_ltln \<phi>"
by (induction w rule: rev_induct) (simp, insert af_letter_atoms, fastforce)
lemma af_letter_nested_prop_atoms:
"nested_prop_atoms (af_letter \<phi> \<nu>) \<subseteq> nested_prop_atoms \<phi>"
by (induction \<phi>) auto
lemma af_nested_prop_atoms:
"nested_prop_atoms (af \<phi> w) \<subseteq> nested_prop_atoms \<phi>"
by (induction w rule: rev_induct) (auto, insert af_letter_nested_prop_atoms, blast)
lemma af_letter_range:
"range (af_letter \<phi>) \<subseteq> {\<psi>. nested_prop_atoms \<psi> \<subseteq> nested_prop_atoms \<phi>}"
using af_letter_nested_prop_atoms by blast
lemma af_range:
"range (af \<phi>) \<subseteq> {\<psi>. nested_prop_atoms \<psi> \<subseteq> nested_prop_atoms \<phi>}"
using af_nested_prop_atoms by blast
subsection \<open>Subformulas of the after function\<close>
lemma af_letter_subformulas\<^sub>\<mu>:
"subformulas\<^sub>\<mu> (af_letter \<phi> \<xi>) = subformulas\<^sub>\<mu> \<phi>"
by (induction \<phi>) auto
lemma af_subformulas\<^sub>\<mu>:
"subformulas\<^sub>\<mu> (af \<phi> w) = subformulas\<^sub>\<mu> \<phi>"
using af_letter_subformulas\<^sub>\<mu>
by (induction w arbitrary: \<phi> rule: rev_induct) (simp, fastforce)
lemma af_letter_subformulas\<^sub>\<nu>:
"subformulas\<^sub>\<nu> (af_letter \<phi> \<xi>) = subformulas\<^sub>\<nu> \<phi>"
by (induction \<phi>) auto
lemma af_subformulas\<^sub>\<nu>:
"subformulas\<^sub>\<nu> (af \<phi> w) = subformulas\<^sub>\<nu> \<phi>"
using af_letter_subformulas\<^sub>\<nu>
by (induction w arbitrary: \<phi> rule: rev_induct) (simp, fastforce)
subsection \<open>Stability and the after function\<close>
lemma \<G>\<F>_af:
"\<G>\<F> (af \<phi> (prefix i w)) (suffix i w) = \<G>\<F> \<phi> (suffix i w)"
unfolding \<G>\<F>_semantics' af_subformulas\<^sub>\<mu> by blast
lemma \<F>_af:
"\<F> (af \<phi> (prefix i w)) (suffix i w) = \<F> \<phi> (suffix i w)"
unfolding \<F>_semantics' af_subformulas\<^sub>\<mu> by blast
lemma \<F>\<G>_af:
"\<F>\<G> (af \<phi> (prefix i w)) (suffix i w) = \<F>\<G> \<phi> (suffix i w)"
unfolding \<F>\<G>_semantics' af_subformulas\<^sub>\<nu> by blast
lemma \<G>_af:
"\<G> (af \<phi> (prefix i w)) (suffix i w) = \<G> \<phi> (suffix i w)"
unfolding \<G>_semantics' af_subformulas\<^sub>\<nu> by blast
subsection \<open>Congruence\<close>
lemma af_letter_lang_congruent:
"\<phi> \<sim>\<^sub>L \<psi> \<Longrightarrow> af_letter \<phi> \<nu> \<sim>\<^sub>L af_letter \<psi> \<nu>"
unfolding ltl_lang_equiv_def
using af_letter_build by blast
lemma af_lang_congruent:
"\<phi> \<sim>\<^sub>L \<psi> \<Longrightarrow> af \<phi> w \<sim>\<^sub>L af \<psi> w"
unfolding ltl_lang_equiv_def using af_ltl_continuation
by (induction \<phi>) blast+
lemma af_letter_subst:
"af_letter \<phi> \<nu> = subst \<phi> (\<lambda>\<psi>. Some (af_letter \<psi> \<nu>))"
by (induction \<phi>) auto
lemma af_letter_prop_congruent:
"\<phi> \<longrightarrow>\<^sub>P \<psi> \<Longrightarrow> af_letter \<phi> \<nu> \<longrightarrow>\<^sub>P af_letter \<psi> \<nu>"
"\<phi> \<sim>\<^sub>P \<psi> \<Longrightarrow> af_letter \<phi> \<nu> \<sim>\<^sub>P af_letter \<psi> \<nu>"
by (metis af_letter_subst subst_respects_ltl_prop_entailment)+
lemma af_prop_congruent:
"\<phi> \<longrightarrow>\<^sub>P \<psi> \<Longrightarrow> af \<phi> w \<longrightarrow>\<^sub>P af \<psi> w"
"\<phi> \<sim>\<^sub>P \<psi> \<Longrightarrow> af \<phi> w \<sim>\<^sub>P af \<psi> w"
by (induction w arbitrary: \<phi> \<psi>) (insert af_letter_prop_congruent, fastforce+)
lemma af_letter_const_congruent:
"\<phi> \<sim>\<^sub>C \<psi> \<Longrightarrow> af_letter \<phi> \<nu> \<sim>\<^sub>C af_letter \<psi> \<nu>"
by (metis af_letter_subst subst_respects_ltl_const_entailment)
lemma af_const_congruent:
"\<phi> \<sim>\<^sub>C \<psi> \<Longrightarrow> af \<phi> w \<sim>\<^sub>C af \<psi> w"
by (induction w arbitrary: \<phi> \<psi>) (insert af_letter_const_congruent, fastforce+)
lemma af_letter_one_step_back:
"{x. \<A> \<Turnstile>\<^sub>P af_letter x \<sigma>} \<Turnstile>\<^sub>P \<phi> \<longleftrightarrow> \<A> \<Turnstile>\<^sub>P af_letter \<phi> \<sigma>"
by (induction \<phi>) simp_all
subsection \<open>Implications\<close>
lemma af_F_prefix_prop:
"af (F\<^sub>n \<phi>) w \<longrightarrow>\<^sub>P af (F\<^sub>n \<phi>) (w' @ w)"
by (induction w') (simp add: ltl_prop_implies_def af_decompose(1,2))+
lemma af_G_prefix_prop:
"af (G\<^sub>n \<phi>) (w' @ w) \<longrightarrow>\<^sub>P af (G\<^sub>n \<phi>) w"
by (induction w') (simp add: ltl_prop_implies_def af_decompose(1,2))+
lemma af_F_prefix_lang:
"w \<Turnstile>\<^sub>n af (F\<^sub>n \<phi>) ys \<Longrightarrow> w \<Turnstile>\<^sub>n af (F\<^sub>n \<phi>) (xs @ ys)"
using af_F_prefix_prop ltl_prop_implication_implies_ltl_implication by blast
lemma af_G_prefix_lang:
"w \<Turnstile>\<^sub>n af (G\<^sub>n \<phi>) (xs @ ys) \<Longrightarrow> w \<Turnstile>\<^sub>n af (G\<^sub>n \<phi>) ys"
using af_G_prefix_prop ltl_prop_implication_implies_ltl_implication by blast
lemma af_F_prefix_const_equiv_true:
"af (F\<^sub>n \<phi>) w \<sim>\<^sub>C true\<^sub>n \<Longrightarrow> af (F\<^sub>n \<phi>) (w' @ w) \<sim>\<^sub>C true\<^sub>n"
using af_F_prefix_prop ltl_const_equiv_implies_prop_equiv(1) ltl_prop_equiv_true_implies_true by blast
lemma af_G_prefix_const_equiv_false:
"af (G\<^sub>n \<phi>) w \<sim>\<^sub>C false\<^sub>n \<Longrightarrow> af (G\<^sub>n \<phi>) (w' @ w) \<sim>\<^sub>C false\<^sub>n"
using af_G_prefix_prop ltl_const_equiv_implies_prop_equiv(2) ltl_prop_equiv_false_implied_by_false by blast
lemma af_F_prefix_lang_equiv_true:
"af (F\<^sub>n \<phi>) w \<sim>\<^sub>L true\<^sub>n \<Longrightarrow> af (F\<^sub>n \<phi>) (w' @ w) \<sim>\<^sub>L true\<^sub>n"
unfolding ltl_lang_equiv_def
using af_F_prefix_lang by fastforce
lemma af_G_prefix_lang_equiv_false:
"af (G\<^sub>n \<phi>) w \<sim>\<^sub>L false\<^sub>n \<Longrightarrow> af (G\<^sub>n \<phi>) (w' @ w) \<sim>\<^sub>L false\<^sub>n"
unfolding ltl_lang_equiv_def
using af_G_prefix_lang by fastforce
locale af_congruent = ltl_equivalence +
assumes
af_letter_congruent: "\<phi> \<sim> \<psi> \<Longrightarrow> af_letter \<phi> \<nu> \<sim> af_letter \<psi> \<nu>"
begin
lemma af_congruentness:
"\<phi> \<sim> \<psi> \<Longrightarrow> af \<phi> xs \<sim> af \<psi> xs"
by (induction xs arbitrary: \<phi> \<psi>) (insert af_letter_congruent, fastforce+)
lemma af_append_congruent:
"af \<phi> w \<sim> af \<psi> w \<Longrightarrow> af \<phi> (w @ w') \<sim> af \<psi> (w @ w')"
by (simp add: af_congruentness)
lemma af_append_true:
"af \<phi> w \<sim> true\<^sub>n \<Longrightarrow> af \<phi> (w @ w') \<sim> true\<^sub>n"
using af_congruentness by fastforce
lemma af_append_false:
"af \<phi> w \<sim> false\<^sub>n \<Longrightarrow> af \<phi> (w @ w') \<sim> false\<^sub>n"
using af_congruentness by fastforce
lemma prefix_append_subsequence:
"i \<le> j \<Longrightarrow> (prefix i w) @ (w [i \<rightarrow> j]) = prefix j w"
by (metis le_add_diff_inverse subsequence_append)
lemma af_prefix_congruent:
"i \<le> j \<Longrightarrow> af \<phi> (prefix i w) \<sim> af \<psi> (prefix i w) \<Longrightarrow> af \<phi> (prefix j w) \<sim> af \<psi> (prefix j w)"
by (metis af_congruentness foldl_append prefix_append_subsequence)+
lemma af_prefix_true:
"i \<le> j \<Longrightarrow> af \<phi> (prefix i w) \<sim> true\<^sub>n \<Longrightarrow> af \<phi> (prefix j w) \<sim> true\<^sub>n"
by (metis af_append_true prefix_append_subsequence)
lemma af_prefix_false:
"i \<le> j \<Longrightarrow> af \<phi> (prefix i w) \<sim> false\<^sub>n \<Longrightarrow> af \<phi> (prefix j w) \<sim> false\<^sub>n"
by (metis af_append_false prefix_append_subsequence)
end
interpretation lang_af_congruent: af_congruent "(\<sim>\<^sub>L)"
by unfold_locales (rule af_letter_lang_congruent)
interpretation prop_af_congruent: af_congruent "(\<sim>\<^sub>P)"
by unfold_locales (rule af_letter_prop_congruent)
interpretation const_af_congruent: af_congruent "(\<sim>\<^sub>C)"
by unfold_locales (rule af_letter_const_congruent)
subsection \<open>After in @{term \<mu>LTL} and @{term \<nu>LTL}\<close>
lemma valid_prefix_implies_ltl:
"af \<phi> (prefix i w) \<sim>\<^sub>L true\<^sub>n \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
proof -
assume "af \<phi> (prefix i w) \<sim>\<^sub>L true\<^sub>n"
then have "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)"
unfolding ltl_lang_equiv_def using semantics_ltln.simps(1) by blast
then show "w \<Turnstile>\<^sub>n \<phi>"
using af_ltl_continuation by force
qed
lemma ltl_implies_satisfiable_prefix:
"w \<Turnstile>\<^sub>n \<phi> \<Longrightarrow> \<not> (af \<phi> (prefix i w) \<sim>\<^sub>L false\<^sub>n)"
proof -
assume "w \<Turnstile>\<^sub>n \<phi>"
then have "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)"
using af_ltl_continuation by fastforce
then show "\<not> (af \<phi> (prefix i w) \<sim>\<^sub>L false\<^sub>n)"
unfolding ltl_lang_equiv_def using semantics_ltln.simps(2) by blast
qed
lemma \<mu>LTL_implies_valid_prefix:
"\<phi> \<in> \<mu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<Longrightarrow> \<exists>i. af \<phi> (prefix i w) \<sim>\<^sub>C true\<^sub>n"
proof (induction \<phi> arbitrary: w)
case True_ltln
then show ?case
using ltl_const_equiv_equivp equivp_reflp by fastforce
next
case (Prop_ltln x)
then show ?case
by (metis af_letter.simps(3) foldl_Cons foldl_Nil ltl_const_equiv_equivp equivp_reflp semantics_ltln.simps(3) subsequence_singleton)
next
case (Nprop_ltln x)
then show ?case
by (metis af_letter.simps(4) foldl_Cons foldl_Nil ltl_const_equiv_equivp equivp_reflp semantics_ltln.simps(4) subsequence_singleton)
next
case (And_ltln \<phi>1 \<phi>2)
then obtain i1 i2 where "af \<phi>1 (prefix i1 w) \<sim>\<^sub>C true\<^sub>n" and "af \<phi>2 (prefix i2 w) \<sim>\<^sub>C true\<^sub>n"
by fastforce
then have "af \<phi>1 (prefix (i1 + i2) w) \<sim>\<^sub>C true\<^sub>n" and "af \<phi>2 (prefix (i2 + i1) w) \<sim>\<^sub>C true\<^sub>n"
using const_af_congruent.af_prefix_true le_add1 by blast+
then have "af (\<phi>1 and\<^sub>n \<phi>2) (prefix (i1 + i2) w) \<sim>\<^sub>C true\<^sub>n"
unfolding af_decompose by (simp add: add.commute)
then show ?case
by blast
next
case (Or_ltln \<phi>1 \<phi>2)
then obtain i where "af \<phi>1 (prefix i w) \<sim>\<^sub>C true\<^sub>n \<or> af \<phi>2 (prefix i w) \<sim>\<^sub>C true\<^sub>n"
by auto
then show ?case
unfolding af_decompose by auto
next
case (Next_ltln \<phi>)
then obtain i where "af \<phi> (prefix i (suffix 1 w)) \<sim>\<^sub>C true\<^sub>n"
by fastforce
then show ?case
by (metis (no_types, lifting) One_nat_def add.right_neutral af_simps(3) foldl_Nil foldl_append subsequence_append subsequence_shift subsequence_singleton)
next
case (Until_ltln \<phi>1 \<phi>2)
then obtain k where "suffix k w \<Turnstile>\<^sub>n \<phi>2" and "\<forall>j<k. suffix j w \<Turnstile>\<^sub>n \<phi>1"
by fastforce
then show ?case
proof (induction k arbitrary: w)
case 0
then obtain i where "af \<phi>2 (prefix i w) \<sim>\<^sub>C true\<^sub>n"
using Until_ltln by fastforce
then have "af \<phi>2 (prefix (Suc i) w) \<sim>\<^sub>C true\<^sub>n"
using const_af_congruent.af_prefix_true le_SucI by blast
then have "af (\<phi>1 U\<^sub>n \<phi>2) (prefix (Suc i) w) \<sim>\<^sub>C true\<^sub>n"
unfolding af_subsequence_U by simp
then show ?case
by blast
next
case (Suc k)
then have "suffix k (suffix 1 w) \<Turnstile>\<^sub>n \<phi>2" and "\<forall>j<k. suffix j (suffix 1 w) \<Turnstile>\<^sub>n \<phi>1"
by (simp add: Suc.prems)+
then obtain i where i_def: "af (\<phi>1 U\<^sub>n \<phi>2) (prefix i (suffix 1 w)) \<sim>\<^sub>C true\<^sub>n"
using Suc.IH by blast
obtain j where "af \<phi>1 (prefix j w) \<sim>\<^sub>C true\<^sub>n"
using Until_ltln Suc by fastforce
then have "af \<phi>1 (prefix (Suc (j + i)) w) \<sim>\<^sub>C true\<^sub>n"
using const_af_congruent.af_prefix_true le_SucI le_add1 by blast
moreover
from i_def have "af (\<phi>1 U\<^sub>n \<phi>2) (w [1 \<rightarrow> Suc (j + i)]) \<sim>\<^sub>C true\<^sub>n"
by (metis One_nat_def const_af_congruent.af_prefix_true le_add2 plus_1_eq_Suc subsequence_shift)
ultimately
have "af (\<phi>1 U\<^sub>n \<phi>2) (prefix (Suc (j + i)) w) \<sim>\<^sub>C true\<^sub>n"
unfolding af_subsequence_U by simp
then show ?case
by blast
qed
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then obtain k where "suffix k w \<Turnstile>\<^sub>n \<phi>1" and "\<forall>j\<le>k. suffix j w \<Turnstile>\<^sub>n \<phi>2"
by fastforce
then show ?case
proof (induction k arbitrary: w)
case 0
then obtain i1 i2 where "af \<phi>1 (prefix i1 w) \<sim>\<^sub>C true\<^sub>n" and "af \<phi>2 (prefix i2 w) \<sim>\<^sub>C true\<^sub>n"
using StrongRelease_ltln by fastforce
then have "af \<phi>1 (prefix (Suc (i1 + i2)) w) \<sim>\<^sub>C true\<^sub>n" and "af \<phi>2 (prefix (Suc (i2 + i1)) w) \<sim>\<^sub>C true\<^sub>n"
using const_af_congruent.af_prefix_true le_SucI le_add1 by blast+
then have "af (\<phi>1 M\<^sub>n \<phi>2) (prefix (Suc (i1 + i2)) w) \<sim>\<^sub>C true\<^sub>n"
unfolding af_subsequence_M by (simp add: add.commute)
then show ?case
by blast
next
case (Suc k)
then have "suffix k (suffix 1 w) \<Turnstile>\<^sub>n \<phi>1" and "\<forall>j\<le>k. suffix j (suffix 1 w) \<Turnstile>\<^sub>n \<phi>2"
by (simp add: Suc.prems)+
then obtain i where i_def: "af (\<phi>1 M\<^sub>n \<phi>2) (prefix i (suffix 1 w)) \<sim>\<^sub>C true\<^sub>n"
using Suc.IH by blast
obtain j where "af \<phi>2 (prefix j w) \<sim>\<^sub>C true\<^sub>n"
using StrongRelease_ltln Suc by fastforce
then have "af \<phi>2 (prefix (Suc (j + i)) w) \<sim>\<^sub>C true\<^sub>n"
using const_af_congruent.af_prefix_true le_SucI le_add1 by blast
moreover
from i_def have "af (\<phi>1 M\<^sub>n \<phi>2) (w [1 \<rightarrow> Suc (j + i)]) \<sim>\<^sub>C true\<^sub>n"
by (metis One_nat_def const_af_congruent.af_prefix_true le_add2 plus_1_eq_Suc subsequence_shift)
ultimately
have "af (\<phi>1 M\<^sub>n \<phi>2) (prefix (Suc (j + i)) w) \<sim>\<^sub>C true\<^sub>n"
unfolding af_subsequence_M by simp
then show ?case
by blast
qed
qed force+
lemma satisfiable_prefix_implies_\<nu>LTL:
"\<phi> \<in> \<nu>LTL \<Longrightarrow> \<nexists>i. af \<phi> (prefix i w) \<sim>\<^sub>C false\<^sub>n \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
proof (erule contrapos_np, induction \<phi> arbitrary: w)
case False_ltln
then show ?case
using ltl_const_equiv_equivp equivp_reflp by fastforce
next
case (Prop_ltln x)
then show ?case
by (metis af_letter.simps(3) foldl_Cons foldl_Nil ltl_const_equiv_equivp equivp_reflp semantics_ltln.simps(3) subsequence_singleton)
next
case (Nprop_ltln x)
then show ?case
by (metis af_letter.simps(4) foldl_Cons foldl_Nil ltl_const_equiv_equivp equivp_reflp semantics_ltln.simps(4) subsequence_singleton)
next
case (And_ltln \<phi>1 \<phi>2)
then obtain i where "af \<phi>1 (prefix i w) \<sim>\<^sub>C false\<^sub>n \<or> af \<phi>2 (prefix i w) \<sim>\<^sub>C false\<^sub>n"
by auto
then show ?case
unfolding af_decompose by auto
next
case (Or_ltln \<phi>1 \<phi>2)
then obtain i1 i2 where "af \<phi>1 (prefix i1 w) \<sim>\<^sub>C false\<^sub>n" and "af \<phi>2 (prefix i2 w) \<sim>\<^sub>C false\<^sub>n"
by fastforce
then have "af \<phi>1 (prefix (i1 + i2) w) \<sim>\<^sub>C false\<^sub>n" and "af \<phi>2 (prefix (i2 + i1) w) \<sim>\<^sub>C false\<^sub>n"
using const_af_congruent.af_prefix_false le_add1 by blast+
then have "af (\<phi>1 or\<^sub>n \<phi>2) (prefix (i1 + i2) w) \<sim>\<^sub>C false\<^sub>n"
unfolding af_decompose by (simp add: add.commute)
then show ?case
by blast
next
case (Next_ltln \<phi>)
then obtain i where "af \<phi> (prefix i (suffix 1 w)) \<sim>\<^sub>C false\<^sub>n"
by fastforce
then show ?case
by (metis (no_types, lifting) One_nat_def add.right_neutral af_simps(3) foldl_Nil foldl_append subsequence_append subsequence_shift subsequence_singleton)
next
case (Release_ltln \<phi>1 \<phi>2)
then obtain k where "\<not> suffix k w \<Turnstile>\<^sub>n \<phi>2" and "\<forall>j<k. \<not> suffix j w \<Turnstile>\<^sub>n \<phi>1"
by fastforce
then show ?case
proof (induction k arbitrary: w)
case 0
then obtain i where "af \<phi>2 (prefix i w) \<sim>\<^sub>C false\<^sub>n"
using Release_ltln by fastforce
then have "af \<phi>2 (prefix (Suc i) w) \<sim>\<^sub>C false\<^sub>n"
using const_af_congruent.af_prefix_false le_SucI by blast
then have "af (\<phi>1 R\<^sub>n \<phi>2) (prefix (Suc i) w) \<sim>\<^sub>C false\<^sub>n"
unfolding af_subsequence_R by simp
then show ?case
by blast
next
case (Suc k)
then have "\<not> suffix k (suffix 1 w) \<Turnstile>\<^sub>n \<phi>2" and "\<forall>j<k. \<not> suffix j (suffix 1 w) \<Turnstile>\<^sub>n \<phi>1"
by (simp add: Suc.prems)+
then obtain i where i_def: "af (\<phi>1 R\<^sub>n \<phi>2) (prefix i (suffix 1 w)) \<sim>\<^sub>C false\<^sub>n"
using Suc.IH by blast
obtain j where "af \<phi>1 (prefix j w) \<sim>\<^sub>C false\<^sub>n"
using Release_ltln Suc by fastforce
then have "af \<phi>1 (prefix (Suc (j + i)) w) \<sim>\<^sub>C false\<^sub>n"
using const_af_congruent.af_prefix_false le_SucI le_add1 by blast
moreover
from i_def have "af (\<phi>1 R\<^sub>n \<phi>2) (w [1 \<rightarrow> Suc (j + i)]) \<sim>\<^sub>C false\<^sub>n"
by (metis One_nat_def const_af_congruent.af_prefix_false le_add2 plus_1_eq_Suc subsequence_shift)
ultimately
have "af (\<phi>1 R\<^sub>n \<phi>2) (prefix (Suc (j + i)) w) \<sim>\<^sub>C false\<^sub>n"
unfolding af_subsequence_R by auto
then show ?case
by blast
qed
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then obtain k where "\<not> suffix k w \<Turnstile>\<^sub>n \<phi>1" and "\<forall>j\<le>k. \<not> suffix j w \<Turnstile>\<^sub>n \<phi>2"
by fastforce
then show ?case
proof (induction k arbitrary: w)
case 0
then obtain i1 i2 where "af \<phi>1 (prefix i1 w) \<sim>\<^sub>C false\<^sub>n" and "af \<phi>2 (prefix i2 w) \<sim>\<^sub>C false\<^sub>n"
using WeakUntil_ltln by fastforce
then have "af \<phi>1 (prefix (Suc (i1 + i2)) w) \<sim>\<^sub>C false\<^sub>n" and "af \<phi>2 (prefix (Suc (i2 + i1)) w) \<sim>\<^sub>C false\<^sub>n"
using const_af_congruent.af_prefix_false le_SucI le_add1 by blast+
then have "af (\<phi>1 W\<^sub>n \<phi>2) (prefix (Suc (i1 + i2)) w) \<sim>\<^sub>C false\<^sub>n"
unfolding af_subsequence_W by (simp add: add.commute)
then show ?case
by blast
next
case (Suc k)
then have "\<not> suffix k (suffix 1 w) \<Turnstile>\<^sub>n \<phi>1" and "\<forall>j\<le>k. \<not> suffix j (suffix 1 w) \<Turnstile>\<^sub>n \<phi>2"
by (simp add: Suc.prems)+
then obtain i where i_def: "af (\<phi>1 W\<^sub>n \<phi>2) (prefix i (suffix 1 w)) \<sim>\<^sub>C false\<^sub>n"
using Suc.IH by blast
obtain j where "af \<phi>2 (prefix j w) \<sim>\<^sub>C false\<^sub>n"
using WeakUntil_ltln Suc by fastforce
then have "af \<phi>2 (prefix (Suc (j + i)) w) \<sim>\<^sub>C false\<^sub>n"
using const_af_congruent.af_prefix_false le_SucI le_add1 by blast
moreover
from i_def have "af (\<phi>1 W\<^sub>n \<phi>2) (w [1 \<rightarrow> Suc (j + i)]) \<sim>\<^sub>C false\<^sub>n"
by (metis One_nat_def const_af_congruent.af_prefix_false le_add2 plus_1_eq_Suc subsequence_shift)
ultimately
have "af (\<phi>1 W\<^sub>n \<phi>2) (prefix (Suc (j + i)) w) \<sim>\<^sub>C false\<^sub>n"
unfolding af_subsequence_W by simp
then show ?case
by blast
qed
qed force+
context ltl_equivalence
begin
lemma valid_prefix_implies_ltl:
"af \<phi> (prefix i w) \<sim> true\<^sub>n \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
using valid_prefix_implies_ltl eq_implies_lang by blast
lemma ltl_implies_satisfiable_prefix:
"w \<Turnstile>\<^sub>n \<phi> \<Longrightarrow> \<not> (af \<phi> (prefix i w) \<sim> false\<^sub>n)"
using ltl_implies_satisfiable_prefix eq_implies_lang by blast
lemma \<mu>LTL_implies_valid_prefix:
"\<phi> \<in> \<mu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<Longrightarrow> \<exists>i. af \<phi> (prefix i w) \<sim> true\<^sub>n"
using \<mu>LTL_implies_valid_prefix const_implies_eq by blast
lemma satisfiable_prefix_implies_\<nu>LTL:
"\<phi> \<in> \<nu>LTL \<Longrightarrow> \<nexists>i. af \<phi> (prefix i w) \<sim> false\<^sub>n \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi>"
using satisfiable_prefix_implies_\<nu>LTL const_implies_eq by blast
lemma af_\<mu>LTL:
"\<phi> \<in> \<mu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> (\<exists>i. af \<phi> (prefix i w) \<sim> true\<^sub>n)"
using valid_prefix_implies_ltl \<mu>LTL_implies_valid_prefix by blast
lemma af_\<nu>LTL:
"\<phi> \<in> \<nu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow> (\<forall>i. \<not> (af \<phi> (prefix i w) \<sim> false\<^sub>n))"
using ltl_implies_satisfiable_prefix satisfiable_prefix_implies_\<nu>LTL by blast
lemma af_\<mu>LTL_GF:
"\<phi> \<in> \<mu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>) \<longleftrightarrow> (\<forall>i. \<exists>j. af (F\<^sub>n \<phi>) (w[i \<rightarrow> j]) \<sim> true\<^sub>n)"
proof -
assume "\<phi> \<in> \<mu>LTL"
then have "F\<^sub>n \<phi> \<in> \<mu>LTL"
by simp
have "w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>) \<longleftrightarrow> (\<forall>i. suffix i w \<Turnstile>\<^sub>n F\<^sub>n \<phi>)"
by simp
also have "\<dots> \<longleftrightarrow> (\<forall>i. \<exists>j. af (F\<^sub>n \<phi>) (prefix j (suffix i w)) \<sim> true\<^sub>n)"
- using af_\<mu>LTL[OF `F\<^sub>n \<phi> \<in> \<mu>LTL`] by blast
+ using af_\<mu>LTL[OF \<open>F\<^sub>n \<phi> \<in> \<mu>LTL\<close>] by blast
also have "\<dots> \<longleftrightarrow> (\<forall>i. \<exists>j. af (F\<^sub>n \<phi>) (prefix (j - i) (suffix i w)) \<sim> true\<^sub>n)"
by (metis diff_add_inverse)
also have "\<dots> \<longleftrightarrow> (\<forall>i. \<exists>j. af (F\<^sub>n \<phi>) (w[i \<rightarrow> j]) \<sim> true\<^sub>n)"
unfolding subsequence_prefix_suffix ..
finally show ?thesis
by blast
qed
lemma af_\<nu>LTL_FG:
"\<phi> \<in> \<nu>LTL \<Longrightarrow> w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>) \<longleftrightarrow> (\<exists>i. \<forall>j. \<not> (af (G\<^sub>n \<phi>) (w[i \<rightarrow> j]) \<sim> false\<^sub>n))"
proof -
assume "\<phi> \<in> \<nu>LTL"
then have "G\<^sub>n \<phi> \<in> \<nu>LTL"
by simp
have "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>) \<longleftrightarrow> (\<exists>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>)"
by force
also have "\<dots> \<longleftrightarrow> (\<exists>i. \<forall>j. \<not> (af (G\<^sub>n \<phi>) (prefix j (suffix i w)) \<sim> false\<^sub>n))"
- using af_\<nu>LTL[OF `G\<^sub>n \<phi> \<in> \<nu>LTL`] by blast
+ using af_\<nu>LTL[OF \<open>G\<^sub>n \<phi> \<in> \<nu>LTL\<close>] by blast
also have "\<dots> \<longleftrightarrow> (\<exists>i. \<forall>j. \<not> (af (G\<^sub>n \<phi>) (prefix (j - i) (suffix i w)) \<sim> false\<^sub>n))"
by (metis diff_add_inverse)
also have "\<dots> \<longleftrightarrow> (\<exists>i. \<forall>j. \<not> (af (G\<^sub>n \<phi>) (w[i \<rightarrow> j]) \<sim> false\<^sub>n))"
unfolding subsequence_prefix_suffix ..
finally show ?thesis
by blast
qed
end
text \<open>Bring Propositional Equivalence into scope\<close>
interpretation af_congruent "(\<sim>\<^sub>P)"
by unfold_locales (rule af_letter_prop_congruent)
end
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/Asymmetric_Master_Theorem.thy b/thys/LTL_Master_Theorem/Logical_Characterization/Asymmetric_Master_Theorem.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/Asymmetric_Master_Theorem.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/Asymmetric_Master_Theorem.thy
@@ -1,278 +1,278 @@
(*
Author: Benedikt Seidl
License: BSD
*)
section \<open>Asymmetric Variant of the Master Theorem\<close>
theory Asymmetric_Master_Theorem
imports
Advice After
begin
text \<open>This variant of the Master Theorem fixes only a subset @{term Y}
of @{term \<nu>LTL} subformulas and all conditions depend on the
index @{term i}. While this does not lead to a simple DRA construction,
but can be used to build NBAs and LDBAs.\<close>
lemma FG_advice_b1_helper:
"\<psi> \<in> subfrmlsn \<phi> \<Longrightarrow> suffix i w \<Turnstile>\<^sub>n \<psi> \<Longrightarrow> suffix i w \<Turnstile>\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>"
proof -
assume "\<psi> \<in> subfrmlsn \<phi>"
then have "\<F>\<G> \<psi> (suffix i w) \<subseteq> \<F>\<G> \<phi> w"
using \<F>\<G>_suffix subformulas\<^sub>\<nu>_subset unfolding \<F>\<G>_semantics' by fast
moreover
assume "suffix i w \<Turnstile>\<^sub>n \<psi>"
ultimately show "suffix i w \<Turnstile>\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>"
using FG_advice_b1 by blast
qed
lemma FG_advice_b2_helper:
"S \<subseteq> \<G> \<phi> (suffix i w) \<Longrightarrow> i \<le> j \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<psi>[S]\<^sub>\<mu> \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<psi>"
proof -
fix i j
assume "S \<subseteq> \<G> \<phi> (suffix i w)" and "i \<le> j" and "suffix j w \<Turnstile>\<^sub>n \<psi>[S]\<^sub>\<mu>"
then have "suffix j w \<Turnstile>\<^sub>n \<psi>[S \<inter> subformulas\<^sub>\<nu> \<psi>]\<^sub>\<mu>"
using FG_advice_inter_subformulas by metis
moreover
have "S \<inter> subformulas\<^sub>\<nu> \<psi> \<subseteq> \<G> \<psi> (suffix i w)"
- using `S \<subseteq> \<G> \<phi> (suffix i w)` unfolding \<G>_semantics' by blast
+ using \<open>S \<subseteq> \<G> \<phi> (suffix i w)\<close> unfolding \<G>_semantics' by blast
then have "S \<inter> subformulas\<^sub>\<nu> \<psi> \<subseteq> \<G> \<psi> (suffix j w)"
using \<G>_suffix \<open>i \<le> j\<close> inf.absorb_iff2 le_Suc_ex by fastforce
ultimately show "suffix j w \<Turnstile>\<^sub>n \<psi>"
using FG_advice_b2 by blast
qed
lemma Y_\<G>:
assumes
Y_\<nu>: "Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
Y_G_1: "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
and
Y_G_2: "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[Y]\<^sub>\<mu>)"
shows
"Y \<subseteq> \<G> \<phi> (suffix i w)"
proof -
\<comment> \<open>Custom induction rule with @{term size} as a partial order\<close>
note induct = finite_ranking_induct[where f = size]
have "finite Y"
using Y_\<nu> finite_subset subformulas\<^sub>\<nu>_finite by auto
then show ?thesis
using assms
proof (induction Y rule: induct)
case (insert \<psi> S)
show ?case
proof (cases "\<psi> \<notin> S")
assume "\<psi> \<notin> S"
- note FG_advice_insert = FG_advice_insert[OF `\<psi> \<notin> S`]
+ note FG_advice_insert = FG_advice_insert[OF \<open>\<psi> \<notin> S\<close>]
{
\<comment> \<open>Show @{term "S \<subseteq> \<G> \<phi> (suffix i w)"}\<close>
{
fix \<psi>\<^sub>1 \<psi>\<^sub>2
assume "\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> S"
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2[insert \<psi> S]\<^sub>\<mu>"
using insert(5) by blast
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2[S]\<^sub>\<mu>"
- using `\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> S` FG_advice_insert insert.hyps(2)
+ using \<open>\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> S\<close> FG_advice_insert insert.hyps(2)
by fastforce
}
moreover
{
fix \<psi>\<^sub>1 \<psi>\<^sub>2
assume "\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> S"
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[insert \<psi> S]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[insert \<psi> S]\<^sub>\<mu>)"
using insert(6) by blast
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[S]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[S]\<^sub>\<mu>)"
- using `\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> S` FG_advice_insert insert.hyps(2)
+ using \<open>\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> S\<close> FG_advice_insert insert.hyps(2)
by fastforce
}
ultimately
have "S \<subseteq> \<G> \<phi> (suffix i w)"
using insert.IH insert.prems(1) by blast
}
moreover
{
\<comment> \<open>Show @{term "\<psi> \<in> \<G> \<phi> (suffix i w)"}\<close>
have "\<psi> \<in> subformulas\<^sub>\<nu> \<phi>"
using insert.prems(1) by fast
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>"
using subformulas\<^sub>\<nu>_semantics
proof (cases \<psi>)
case (Release_ltln \<psi>\<^sub>1 \<psi>\<^sub>2)
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2[insert \<psi> S]\<^sub>\<mu>"
using insert.prems(2) by blast
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2[S]\<^sub>\<mu>"
using Release_ltln FG_advice_insert by simp
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2"
- using FG_advice_b2_helper[OF `S \<subseteq> \<G> \<phi> (suffix i w)`] by auto
+ using FG_advice_b2_helper[OF \<open>S \<subseteq> \<G> \<phi> (suffix i w)\<close>] by auto
then show ?thesis
using Release_ltln globally_release
by blast
next
case (WeakUntil_ltln \<psi>\<^sub>1 \<psi>\<^sub>2)
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[insert \<psi> S]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[insert \<psi> S]\<^sub>\<mu>)"
using insert.prems(3) by blast
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2)[S]\<^sub>\<mu>"
using WeakUntil_ltln FG_advice_insert by simp
then have "suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2)"
- using FG_advice_b2_helper[OF `S \<subseteq> \<G> \<phi> (suffix i w)`, of _ "\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2"]
+ using FG_advice_b2_helper[OF \<open>S \<subseteq> \<G> \<phi> (suffix i w)\<close>, of _ "\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2"]
by force
then show ?thesis
unfolding WeakUntil_ltln semantics_ltln.simps
by (metis order_refl suffix_suffix)
qed fast+
then have "\<psi> \<in> \<G> \<phi> (suffix i w)"
- unfolding \<G>_semantics using `\<psi> \<in> subformulas\<^sub>\<nu> \<phi>`
+ unfolding \<G>_semantics using \<open>\<psi> \<in> subformulas\<^sub>\<nu> \<phi>\<close>
by simp
}
ultimately show ?thesis
by blast
next
assume "\<not> \<psi> \<notin> S"
then have "insert \<psi> S = S"
by auto
then show ?thesis
using insert by simp
qed
qed simp
qed
theorem asymmetric_master_theorem_ltr:
assumes
"w \<Turnstile>\<^sub>n \<phi>"
obtains Y and i where
"Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
"suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[Y]\<^sub>\<mu>"
and
"\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
and
"\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[Y]\<^sub>\<mu>)"
proof
let ?Y = "\<F>\<G> \<phi> w"
show "?Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
by (rule \<F>\<G>_subformulas\<^sub>\<nu>)
next
let ?Y = "\<F>\<G> \<phi> w"
let ?i = "SOME i. ?Y = \<G> \<phi> (suffix i w)"
have "suffix ?i w \<Turnstile>\<^sub>n af \<phi> (prefix ?i w)"
using af_ltl_continuation \<open>w \<Turnstile>\<^sub>n \<phi>\<close> by fastforce
then show "suffix ?i w \<Turnstile>\<^sub>n af \<phi> (prefix ?i w)[?Y]\<^sub>\<mu>"
by (metis \<F>\<G>_suffix FG_advice_b1 \<F>\<G>_af order_refl)
next
let ?Y = "\<F>\<G> \<phi> w"
let ?i = "SOME i. ?Y = \<G> \<phi> (suffix i w)"
have "\<exists>i. ?Y = \<G> \<phi> (suffix i w)"
using suffix_\<nu>_stable \<F>\<G>_suffix unfolding \<nu>_stable_def MOST_nat
by fast
then have Y_G: "?Y = \<G> \<phi> (suffix ?i w)"
by (metis (mono_tags, lifting) someI_ex)
show "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> ?Y \<longrightarrow> suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2[?Y]\<^sub>\<mu>)"
proof safe
fix \<psi>\<^sub>1 \<psi>\<^sub>2
assume "\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> ?Y"
then have "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2)"
using Y_G \<G>_semantics' by blast
then have "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n \<psi>\<^sub>2"
by force
moreover
have "\<psi>\<^sub>2 \<in> subfrmlsn \<phi>"
- using \<F>\<G>_subfrmlsn `\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> ?Y` subfrmlsn_subset by force
+ using \<F>\<G>_subfrmlsn \<open>\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> ?Y\<close> subfrmlsn_subset by force
ultimately show "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2 [?Y]\<^sub>\<mu>)"
using FG_advice_b1_helper by fastforce
qed
next
let ?Y = "\<F>\<G> \<phi> w"
let ?i = "SOME i. ?Y = \<G> \<phi> (suffix i w)"
have "\<exists>i. ?Y = \<G> \<phi> (suffix i w)"
using suffix_\<nu>_stable \<F>\<G>_suffix unfolding \<nu>_stable_def MOST_nat
by fast
then have Y_G: "?Y = \<G> \<phi> (suffix ?i w)"
by (rule someI_ex)
show "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> ?Y \<longrightarrow> suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[?Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[?Y]\<^sub>\<mu>)"
proof safe
fix \<psi>\<^sub>1 \<psi>\<^sub>2
assume "\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> ?Y"
then have "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2)"
using Y_G \<G>_semantics' by blast
then have "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1 or\<^sub>n \<psi>\<^sub>2)"
by force
moreover
have "\<psi>\<^sub>1 \<in> subfrmlsn \<phi>" and "\<psi>\<^sub>2 \<in> subfrmlsn \<phi>"
- using \<F>\<G>_subfrmlsn `\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> ?Y` subfrmlsn_subset by force+
+ using \<F>\<G>_subfrmlsn \<open>\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> ?Y\<close> subfrmlsn_subset by force+
ultimately show "suffix ?i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[?Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[?Y]\<^sub>\<mu>)"
using FG_advice_b1_helper by fastforce
qed
qed
theorem asymmetric_master_theorem_rtl:
assumes
1: "Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
2: "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[Y]\<^sub>\<mu>"
and
3: "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>)"
and
4: "\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[Y]\<^sub>\<mu>)"
shows
"w \<Turnstile>\<^sub>n \<phi>"
proof -
have "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)"
by (metis assms Y_\<G> FG_advice_b2 \<G>_af)
then show "w \<Turnstile>\<^sub>n \<phi>"
using af_ltl_continuation by force
qed
theorem asymmetric_master_theorem:
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow>
(\<exists>i. \<exists>Y \<subseteq> subformulas\<^sub>\<nu> \<phi>.
suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[Y]\<^sub>\<mu>
\<and> (\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>2[Y]\<^sub>\<mu>))
\<and> (\<forall>\<psi>\<^sub>1 \<psi>\<^sub>2. \<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2 \<in> Y \<longrightarrow> suffix i w \<Turnstile>\<^sub>n G\<^sub>n (\<psi>\<^sub>1[Y]\<^sub>\<mu> or\<^sub>n \<psi>\<^sub>2[Y]\<^sub>\<mu>)))"
by (metis asymmetric_master_theorem_ltr asymmetric_master_theorem_rtl)
end
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/Extra_Equivalence_Relations.thy b/thys/LTL_Master_Theorem/Logical_Characterization/Extra_Equivalence_Relations.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/Extra_Equivalence_Relations.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/Extra_Equivalence_Relations.thy
@@ -1,247 +1,247 @@
(*
Author: Salomon Sickert
License: BSD
*)
section \<open>Additional Equivalence Relations\<close>
theory Extra_Equivalence_Relations
imports
LTL.LTL LTL.Equivalence_Relations After Advice
begin
subsection \<open>Propositional Equivalence with Implicit LTL Unfolding\<close>
fun Unf :: "'a ltln \<Rightarrow> 'a ltln"
where
"Unf (\<phi> U\<^sub>n \<psi>) = ((\<phi> U\<^sub>n \<psi>) and\<^sub>n Unf \<phi>) or\<^sub>n Unf \<psi>"
| "Unf (\<phi> W\<^sub>n \<psi>) = ((\<phi> W\<^sub>n \<psi>) and\<^sub>n Unf \<phi>) or\<^sub>n Unf \<psi>"
| "Unf (\<phi> M\<^sub>n \<psi>) = ((\<phi> M\<^sub>n \<psi>) or\<^sub>n Unf \<phi>) and\<^sub>n Unf \<psi>"
| "Unf (\<phi> R\<^sub>n \<psi>) = ((\<phi> R\<^sub>n \<psi>) or\<^sub>n Unf \<phi>) and\<^sub>n Unf \<psi>"
| "Unf (\<phi> and\<^sub>n \<psi>) = Unf \<phi> and\<^sub>n Unf \<psi>"
| "Unf (\<phi> or\<^sub>n \<psi>) = Unf \<phi> or\<^sub>n Unf \<psi>"
| "Unf \<phi> = \<phi>"
lemma Unf_sound:
"w \<Turnstile>\<^sub>n Unf \<phi> \<longleftrightarrow> w \<Turnstile>\<^sub>n \<phi>"
proof (induction \<phi> arbitrary: w)
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
by (simp, metis less_linear not_less0 suffix_0)
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
by (simp, metis less_linear not_less0 suffix_0)
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
by (simp, metis bot.extremum_unique bot_nat_def less_eq_nat.simps(1) suffix_0)
qed (simp_all, fastforce)
lemma Unf_lang_equiv:
"\<phi> \<sim>\<^sub>L Unf \<phi>"
by (simp add: Unf_sound ltl_lang_equiv_def)
lemma Unf_idem:
"Unf (Unf \<phi>) \<sim>\<^sub>P Unf \<phi>"
by (induction \<phi>) (auto simp: ltl_prop_equiv_def)
definition ltl_prop_unfold_equiv :: "'a ltln \<Rightarrow> 'a ltln \<Rightarrow> bool" (infix "\<sim>\<^sub>Q" 75)
where
"\<phi> \<sim>\<^sub>Q \<psi> \<equiv> (Unf \<phi>) \<sim>\<^sub>P (Unf \<psi>)"
lemma ltl_prop_unfold_equiv_equivp:
"equivp (\<sim>\<^sub>Q)"
by (metis ltl_prop_equiv_equivp ltl_prop_unfold_equiv_def equivpI equivp_def reflpI sympI transpI)
lemma unfolding_prop_unfold_idem:
"Unf \<phi> \<sim>\<^sub>Q \<phi>"
unfolding ltl_prop_unfold_equiv_def by (rule Unf_idem)
lemma unfolding_is_subst: "Unf \<phi> = subst \<phi> (\<lambda>\<psi>. Some (Unf \<psi>))"
by (induction \<phi>) auto
lemma ltl_prop_equiv_implies_ltl_prop_unfold_equiv:
"\<phi> \<sim>\<^sub>P \<psi> \<Longrightarrow> \<phi> \<sim>\<^sub>Q \<psi>"
by (metis ltl_prop_unfold_equiv_def unfolding_is_subst subst_respects_ltl_prop_entailment(2))
lemma ltl_prop_unfold_equiv_implies_ltl_lang_equiv:
"\<phi> \<sim>\<^sub>Q \<psi> \<Longrightarrow> \<phi> \<sim>\<^sub>L \<psi>"
by (metis ltl_prop_equiv_implies_ltl_lang_equiv ltl_lang_equiv_def Unf_sound ltl_prop_unfold_equiv_def)
lemma ltl_prop_unfold_equiv_gt_and_lt:
"(\<sim>\<^sub>C) \<le> (\<sim>\<^sub>Q)" "(\<sim>\<^sub>P) \<le> (\<sim>\<^sub>Q)" "(\<sim>\<^sub>Q) \<le> (\<sim>\<^sub>L)"
using ltl_prop_equiv_implies_ltl_prop_unfold_equiv ltl_prop_equivalence.ge_const_equiv ltl_prop_unfold_equiv_implies_ltl_lang_equiv
by blast+
quotient_type 'a ltln\<^sub>Q = "'a ltln" / "(\<sim>\<^sub>Q)"
by (rule ltl_prop_unfold_equiv_equivp)
instantiation ltln\<^sub>Q :: (type) equal
begin
lift_definition ltln\<^sub>Q_eq_test :: "'a ltln\<^sub>Q \<Rightarrow> 'a ltln\<^sub>Q \<Rightarrow> bool" is "\<lambda>x y. x \<sim>\<^sub>Q y"
by (metis ltln\<^sub>Q.abs_eq_iff)
definition
eq\<^sub>Q: "equal_class.equal \<equiv> ltln\<^sub>Q_eq_test"
instance
by (standard; simp add: eq\<^sub>Q ltln\<^sub>Q_eq_test.rep_eq, metis Quotient_ltln\<^sub>Q Quotient_rel_rep)
end
lemma af_letter_unfolding:
"af_letter (Unf \<phi>) \<nu> \<sim>\<^sub>P af_letter \<phi> \<nu>"
by (induction \<phi>) (simp_all add: ltl_prop_equiv_def, blast+)
lemma af_letter_prop_unfold_congruent:
assumes "\<phi> \<sim>\<^sub>Q \<psi>"
shows "af_letter \<phi> \<nu> \<sim>\<^sub>Q af_letter \<psi> \<nu>"
proof -
have "Unf \<phi> \<sim>\<^sub>P Unf \<psi>"
using assms by (simp add: ltl_prop_unfold_equiv_def ltl_prop_equiv_def)
then have "af_letter (Unf \<phi>) \<nu> \<sim>\<^sub>P af_letter (Unf \<psi>) \<nu>"
by (simp add: prop_af_congruent.af_letter_congruent)
then have "af_letter \<phi> \<nu> \<sim>\<^sub>P af_letter \<psi> \<nu>"
by (metis af_letter_unfolding ltl_prop_equivalence.eq_sym ltl_prop_equivalence.eq_trans)
then show "af_letter \<phi> \<nu> \<sim>\<^sub>Q af_letter \<psi> \<nu>"
by (rule ltl_prop_equiv_implies_ltl_prop_unfold_equiv)
qed
lemma GF_advice_prop_unfold_congruent:
assumes "\<phi> \<sim>\<^sub>Q \<psi>"
shows "(Unf \<phi>)[X]\<^sub>\<nu> \<sim>\<^sub>Q (Unf \<psi>)[X]\<^sub>\<nu>"
proof -
have "Unf \<phi> \<sim>\<^sub>P Unf \<psi>"
using assms
by (simp add: ltl_prop_unfold_equiv_def ltl_prop_equiv_def)
then have "(Unf \<phi>)[X]\<^sub>\<nu> \<sim>\<^sub>P (Unf \<psi>)[X]\<^sub>\<nu>"
by (simp add: GF_advice_prop_congruent(2))
then show "(Unf \<phi>)[X]\<^sub>\<nu> \<sim>\<^sub>Q (Unf \<psi>)[X]\<^sub>\<nu>"
by (simp add: ltl_prop_equiv_implies_ltl_prop_unfold_equiv)
qed
interpretation prop_unfold_equivalence: ltl_equivalence "(\<sim>\<^sub>Q)"
by unfold_locales (metis ltl_prop_unfold_equiv_equivp ltl_prop_unfold_equiv_gt_and_lt)+
interpretation af_congruent "(\<sim>\<^sub>Q)"
by unfold_locales (rule af_letter_prop_unfold_congruent)
lemma unfolding_monotonic:
"w \<Turnstile>\<^sub>n \<phi>[X]\<^sub>\<nu> \<Longrightarrow> w \<Turnstile>\<^sub>n (Unf \<phi>)[X]\<^sub>\<nu>"
proof (induction \<phi>)
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
by (cases "(\<phi>1 U\<^sub>n \<phi>2) \<in> X") force+
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
using ltln_expand_Release by auto
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
using ltln_expand_WeakUntil by auto
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then show ?case
by (cases "(\<phi>1 M\<^sub>n \<phi>2) \<in> X") force+
qed auto
lemma unfolding_next_step_equivalent:
"w \<Turnstile>\<^sub>n (Unf \<phi>)[X]\<^sub>\<nu> \<Longrightarrow> suffix 1 w \<Turnstile>\<^sub>n (af_letter \<phi> (w 0))[X]\<^sub>\<nu>"
proof (induction \<phi>)
case (Next_ltln \<phi>)
then show ?case
unfolding Unf.simps by (metis GF_advice_af_letter build_split)
next
case (Until_ltln \<phi>1 \<phi>2)
then show ?case
unfolding Unf.simps
by (metis GF_advice.simps(2) GF_advice.simps(3) GF_advice_af_letter af_letter.simps(8) build_split semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
unfolding Unf.simps
by (metis GF_advice.simps(2) GF_advice.simps(3) GF_advice_af_letter One_nat_def af_letter.simps(9) build_first semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
unfolding Unf.simps
by (metis GF_advice.simps(2) GF_advice.simps(3) GF_advice_af_letter af_letter.simps(10) build_split semantics_ltln.simps(5) semantics_ltln.simps(6))
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
then show ?case
unfolding Unf.simps
by (metis GF_advice.simps(2) GF_advice.simps(3) GF_advice_af_letter af_letter.simps(11) build_split semantics_ltln.simps(5) semantics_ltln.simps(6))
qed auto
lemma nested_prop_atoms_Unf:
"nested_prop_atoms (Unf \<phi>) \<subseteq> nested_prop_atoms \<phi>"
by (induction \<phi>) auto
(* TODO move somewhere?? *)
lemma refine_image:
assumes "\<And>x y. f x = f y \<longrightarrow> g x = g y"
assumes "finite (f ` X)"
shows "finite (g ` X)"
and "card (f ` X) \<ge> card (g ` X)"
proof -
obtain Y where "Y \<subseteq> X" and "finite Y" and Y_def: "f ` X = f ` Y"
using assms by (meson finite_subset_image subset_refl)
moreover
{
fix x
assume "x \<in> X"
then have "g x \<in> g ` Y"
by (metis (no_types, hide_lams) \<open>x \<in> X\<close> assms(1) Y_def image_iff)
}
then have "g ` X = g ` Y"
- using assms `Y \<subseteq> X` by blast
+ using assms \<open>Y \<subseteq> X\<close> by blast
ultimately
show "finite (g ` X)"
by simp
from \<open>finite Y\<close> have "card (f ` Y) \<ge> card (g ` Y)"
proof (induction Y rule: finite_induct)
case (insert x F)
then have 1: "finite (g ` F)" and 2: "finite (f ` F)"
by simp_all
have "f x \<in> f ` F \<Longrightarrow> g x \<in> g ` F"
using assms(1) by blast
then show ?case
using insert by (simp add: card_insert_if[OF 1] card_insert_if[OF 2])
qed simp
then show "card (f ` X) \<ge> card (g ` X)"
by (simp add: Y_def \<open>g ` X = g ` Y\<close>)
qed
lemma abs_ltln\<^sub>P_implies_abs_ltln\<^sub>Q:
"abs_ltln\<^sub>P \<phi> = abs_ltln\<^sub>P \<psi> \<longrightarrow> abs_ltln\<^sub>Q \<phi> = abs_ltln\<^sub>Q \<psi>"
by (simp add: ltl_prop_equiv_implies_ltl_prop_unfold_equiv ltln\<^sub>P.abs_eq_iff ltln\<^sub>Q.abs_eq_iff)
lemmas prop_unfold_equiv_helper = refine_image[of abs_ltln\<^sub>P abs_ltln\<^sub>Q, OF abs_ltln\<^sub>P_implies_abs_ltln\<^sub>Q]
lemma prop_unfold_equiv_finite:
"finite P \<Longrightarrow> finite {abs_ltln\<^sub>Q \<psi> |\<psi>. prop_atoms \<psi> \<subseteq> P}"
using prop_unfold_equiv_helper(1)[OF prop_equiv_finite[unfolded image_Collect[symmetric]]]
unfolding image_Collect[symmetric] .
lemma prop_unfold_equiv_card:
"finite P \<Longrightarrow> card {abs_ltln\<^sub>Q \<psi> |\<psi>. prop_atoms \<psi> \<subseteq> P} \<le> 2 ^ 2 ^ card P"
using prop_unfold_equiv_helper(2)[OF prop_equiv_finite[unfolded image_Collect[symmetric]]] prop_equiv_card
unfolding image_Collect[symmetric]
by fastforce
lemma Unf_eventually_equivalent:
"w \<Turnstile>\<^sub>n Unf \<phi>[X]\<^sub>\<nu> \<Longrightarrow> \<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
by (metis (full_types) One_nat_def foldl.simps(1) foldl.simps(2) subsequence_singleton unfolding_next_step_equivalent)
interpretation prop_unfold_GF_advice_compatible: GF_advice_congruent "(\<sim>\<^sub>Q)" Unf
by unfold_locales (simp_all add: unfolding_prop_unfold_idem prop_unfold_equivalence.eq_sym unfolding_monotonic Unf_eventually_equivalent GF_advice_prop_unfold_congruent)
end
\ No newline at end of file
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/Master_Theorem.thy b/thys/LTL_Master_Theorem/Logical_Characterization/Master_Theorem.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/Master_Theorem.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/Master_Theorem.thy
@@ -1,352 +1,352 @@
(*
Author: Benedikt Seidl
License: BSD
*)
section \<open>The Master Theorem\<close>
theory Master_Theorem
imports
Advice After
begin
subsection \<open>Checking @{term "X \<subseteq> \<G>\<F> \<phi> w"} and @{term "Y \<subseteq> \<F>\<G> \<phi> w"}\<close>
lemma X_\<G>\<F>_Y_\<F>\<G>:
assumes
X_\<mu>: "X \<subseteq> subformulas\<^sub>\<mu> \<phi>"
and
Y_\<nu>: "Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
X_GF: "\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
and
Y_FG: "\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
shows
"X \<subseteq> \<G>\<F> \<phi> w \<and> Y \<subseteq> \<F>\<G> \<phi> w"
proof -
\<comment> \<open>Custom induction rule with @{term size} as a partial order\<close>
note induct = finite_ranking_induct[where f = size]
have "finite (X \<union> Y)"
using subformulas\<^sub>\<mu>_finite subformulas\<^sub>\<nu>_finite X_\<mu> Y_\<nu> finite_subset
by blast+
then show ?thesis
using assms
proof (induction "X \<union> Y" arbitrary: X Y \<phi> rule: induct)
case (insert \<psi> S)
note IH = insert(3)
- note insert_S = `insert \<psi> S = X \<union> Y`
- note X_\<mu> = `X \<subseteq> subformulas\<^sub>\<mu> \<phi>`
- note Y_\<nu> = `Y \<subseteq> subformulas\<^sub>\<nu> \<phi>`
- note X_GF = `\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)`
- note Y_FG = `\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)`
+ note insert_S = \<open>insert \<psi> S = X \<union> Y\<close>
+ note X_\<mu> = \<open>X \<subseteq> subformulas\<^sub>\<mu> \<phi>\<close>
+ note Y_\<nu> = \<open>Y \<subseteq> subformulas\<^sub>\<nu> \<phi>\<close>
+ note X_GF = \<open>\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)\<close>
+ note Y_FG = \<open>\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)\<close>
from X_\<mu> Y_\<nu> have "X \<inter> Y = {}"
using subformulas\<^sub>\<mu>\<^sub>\<nu>_disjoint by fast
from insert_S X_\<mu> Y_\<nu> have "\<psi> \<in> subfrmlsn \<phi>"
using subformulas\<^sub>\<mu>_subfrmlsn subformulas\<^sub>\<nu>_subfrmlsn by blast
show ?case
proof (cases "\<psi> \<notin> S", cases "\<psi> \<in> X")
assume "\<psi> \<notin> S" and "\<psi> \<in> X"
{
\<comment> \<open>Show @{term "X - {\<psi>} \<subseteq> \<G>\<F> \<phi> w"} and @{term "Y \<subseteq> \<F>\<G> \<phi> w"}\<close>
then have "\<psi> \<notin> Y"
- using `X \<inter> Y = {}` by auto
+ using \<open>X \<inter> Y = {}\<close> by auto
then have "S = (X - {\<psi>}) \<union> Y"
- using insert_S `\<psi> \<notin> S` by fast
+ using insert_S \<open>\<psi> \<notin> S\<close> by fast
moreover
have "\<forall>\<psi>' \<in> Y. \<psi>'[X - {\<psi>}]\<^sub>\<nu> = \<psi>'[X]\<^sub>\<nu>"
- using GF_advice_minus_size insert(1,2,4) `\<psi> \<notin> Y` by fast
+ using GF_advice_minus_size insert(1,2,4) \<open>\<psi> \<notin> Y\<close> by fast
ultimately have "X - {\<psi>} \<subseteq> \<G>\<F> \<phi> w" and "Y \<subseteq> \<F>\<G> \<phi> w"
using IH[of "X - {\<psi>}" Y \<phi>] X_\<mu> Y_\<nu> X_GF Y_FG by auto
}
moreover
{
\<comment> \<open>Show @{term "\<psi> \<in> \<G>\<F> \<phi> w"}\<close>
have "w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
- using X_GF `\<psi> \<in> X` by simp
+ using X_GF \<open>\<psi> \<in> X\<close> by simp
then have "\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[Y]\<^sub>\<mu>"
unfolding GF_Inf_many by simp
moreover
from Y_\<nu> have "finite Y"
using subformulas\<^sub>\<nu>_finite finite_subset by auto
have "\<forall>\<phi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<phi>)"
- using `Y \<subseteq> \<F>\<G> \<phi> w` by (blast dest: \<F>\<G>_elim)
+ using \<open>Y \<subseteq> \<F>\<G> \<phi> w\<close> by (blast dest: \<F>\<G>_elim)
then have "\<forall>\<phi> \<in> Y. \<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>"
using FG_suffix_G by blast
then have "\<forall>\<^sub>\<infinity>i. \<forall>\<phi> \<in> Y. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>"
- using `finite Y` eventually_ball_finite by fast
+ using \<open>finite Y\<close> eventually_ball_finite by fast
ultimately
have "\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[Y]\<^sub>\<mu> \<and> (\<forall>\<phi> \<in> Y. suffix i w \<Turnstile>\<^sub>n G\<^sub>n \<phi>)"
using INFM_conjI by auto
then have "\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>"
by (elim frequently_elim1) (metis FG_advice_b2_helper)
then have "w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)"
unfolding GF_Inf_many by simp
then have "\<psi> \<in> \<G>\<F> \<phi> w"
- unfolding \<G>\<F>_semantics using `\<psi> \<in> X` X_\<mu> by auto
+ unfolding \<G>\<F>_semantics using \<open>\<psi> \<in> X\<close> X_\<mu> by auto
}
ultimately show ?thesis
by auto
next
assume "\<psi> \<notin> S" and "\<psi> \<notin> X"
then have "\<psi> \<in> Y"
using insert by fast
{
\<comment> \<open>Show @{term "X \<subseteq> \<G>\<F> \<phi> w"} and @{term "Y - {\<psi>} \<subseteq> \<F>\<G> \<phi> w"}\<close>
then have "S \<inter> X = X"
- using insert `\<psi> \<notin> X` by fast
+ using insert \<open>\<psi> \<notin> X\<close> by fast
then have "S = X \<union> (Y - {\<psi>})"
- using insert_S `\<psi> \<notin> S` by fast
+ using insert_S \<open>\<psi> \<notin> S\<close> by fast
moreover
have "\<forall>\<psi>' \<in> X. \<psi>'[Y - {\<psi>}]\<^sub>\<mu> = \<psi>'[Y]\<^sub>\<mu>"
- using FG_advice_minus_size insert(1,2,4) `\<psi> \<notin> X` by fast
+ using FG_advice_minus_size insert(1,2,4) \<open>\<psi> \<notin> X\<close> by fast
ultimately have "X \<subseteq> \<G>\<F> \<phi> w" and "Y - {\<psi>} \<subseteq> \<F>\<G> \<phi> w"
using IH[of X "Y - {\<psi>}" \<phi>] X_\<mu> Y_\<nu> X_GF Y_FG by auto
}
moreover
{
\<comment> \<open>Show @{term "\<psi> \<in> \<F>\<G> \<phi> w"}\<close>
have "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
- using Y_FG `\<psi> \<in> Y` by simp
+ using Y_FG \<open>\<psi> \<in> Y\<close> by simp
then have "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[X]\<^sub>\<nu>"
unfolding FG_Alm_all by simp
moreover
have "\<forall>\<phi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>)"
- using `X \<subseteq> \<G>\<F> \<phi> w` by (blast dest: \<G>\<F>_elim)
+ using \<open>X \<subseteq> \<G>\<F> \<phi> w\<close> by (blast dest: \<G>\<F>_elim)
then have "\<forall>\<^sub>\<infinity>i. \<forall>\<phi> \<in> X. suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>)"
by simp
ultimately
have "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[X]\<^sub>\<nu> \<and> (\<forall>\<phi> \<in> X. suffix i w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<phi>))"
using MOST_conjI by auto
then have "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>"
by (elim MOST_mono) (metis GF_advice_a2_helper)
then have "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>)"
unfolding FG_Alm_all by simp
then have "\<psi> \<in> \<F>\<G> \<phi> w"
- unfolding \<F>\<G>_semantics using `\<psi> \<in> Y` Y_\<nu> by auto
+ unfolding \<F>\<G>_semantics using \<open>\<psi> \<in> Y\<close> Y_\<nu> by auto
}
ultimately show ?thesis
by auto
next
assume "\<not> \<psi> \<notin> S"
then have "S = X \<union> Y"
using insert by fast
then show ?thesis
using insert by auto
qed
qed fast
qed
lemma \<G>\<F>_implies_GF:
"\<forall>\<psi> \<in> \<G>\<F> \<phi> w. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>)"
proof safe
fix \<psi>
assume "\<psi> \<in> \<G>\<F> \<phi> w"
then have "\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>"
using \<G>\<F>_elim GF_Inf_many by blast
moreover
have "\<psi> \<in> subfrmlsn \<phi>"
- using `\<psi> \<in> \<G>\<F> \<phi> w` \<G>\<F>_subfrmlsn by blast
+ using \<open>\<psi> \<in> \<G>\<F> \<phi> w\<close> \<G>\<F>_subfrmlsn by blast
then have "\<And>i w. \<F>\<G> \<psi> (suffix i w) \<subseteq> \<F>\<G> \<phi> w"
using \<F>\<G>_suffix \<F>\<G>_subset by blast
ultimately have "\<exists>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>"
by (elim frequently_elim1) (metis FG_advice_b1)
then show "w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>)"
unfolding GF_Inf_many by simp
qed
lemma \<F>\<G>_implies_FG:
"\<forall>\<psi> \<in> \<F>\<G> \<phi> w. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[\<G>\<F> \<phi> w]\<^sub>\<nu>)"
proof safe
fix \<psi>
assume "\<psi> \<in> \<F>\<G> \<phi> w"
then have "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>"
using \<F>\<G>_elim FG_Alm_all by blast
moreover
{
have "\<psi> \<in> subfrmlsn \<phi>"
- using `\<psi> \<in> \<F>\<G> \<phi> w` \<F>\<G>_subfrmlsn by blast
+ using \<open>\<psi> \<in> \<F>\<G> \<phi> w\<close> \<F>\<G>_subfrmlsn by blast
moreover have "\<forall>\<^sub>\<infinity>i. \<G>\<F> \<psi> (suffix i w) = \<F> \<psi> (suffix i w)"
using suffix_\<mu>_stable unfolding \<mu>_stable_def by blast
ultimately have "\<forall>\<^sub>\<infinity>i. \<F> \<psi> (suffix i w) \<subseteq> \<G>\<F> \<phi> w"
unfolding MOST_nat_le by (metis \<G>\<F>_subset \<G>\<F>_suffix)
}
ultimately have "\<forall>\<^sub>\<infinity>i. \<F> \<psi> (suffix i w) \<subseteq> \<G>\<F> \<phi> w \<and> suffix i w \<Turnstile>\<^sub>n \<psi>"
using eventually_conj by auto
then have "\<forall>\<^sub>\<infinity>i. suffix i w \<Turnstile>\<^sub>n \<psi>[\<G>\<F> \<phi> w]\<^sub>\<nu>"
using GF_advice_a1 by (elim eventually_mono) auto
then show "w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[\<G>\<F> \<phi> w]\<^sub>\<nu>)"
unfolding FG_Alm_all by simp
qed
subsection \<open>Putting the pieces together: The Master Theorem\<close>
theorem master_theorem_ltr:
assumes
"w \<Turnstile>\<^sub>n \<phi>"
obtains X and Y where
"X \<subseteq> subformulas\<^sub>\<mu> \<phi>"
and
"Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
"\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
and
"\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
and
"\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
proof
show "\<G>\<F> \<phi> w \<subseteq> subformulas\<^sub>\<mu> \<phi>"
by (rule \<G>\<F>_subformulas\<^sub>\<mu>)
next
show "\<F>\<G> \<phi> w \<subseteq> subformulas\<^sub>\<nu> \<phi>"
by (rule \<F>\<G>_subformulas\<^sub>\<nu>)
next
obtain i where "\<G>\<F> \<phi> (suffix i w) = \<F> \<phi> (suffix i w)"
using suffix_\<mu>_stable unfolding MOST_nat \<mu>_stable_def by fast
then have "\<F> (af \<phi> (prefix i w)) (suffix i w) \<subseteq> \<G>\<F> \<phi> w"
using \<G>\<F>_af \<F>_af \<G>\<F>_suffix by fast
moreover
have "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)"
- using af_ltl_continuation `w \<Turnstile>\<^sub>n \<phi>` by fastforce
+ using af_ltl_continuation \<open>w \<Turnstile>\<^sub>n \<phi>\<close> by fastforce
ultimately show "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[\<G>\<F> \<phi> w]\<^sub>\<nu>"
using GF_advice_a1 by blast
next
show "\<forall>\<psi> \<in> \<G>\<F> \<phi> w. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[\<F>\<G> \<phi> w]\<^sub>\<mu>)"
by (rule \<G>\<F>_implies_GF)
next
show "\<forall>\<psi> \<in> \<F>\<G> \<phi> w. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[\<G>\<F> \<phi> w]\<^sub>\<nu>)"
by (rule \<F>\<G>_implies_FG)
qed
theorem master_theorem_rtl:
assumes
"X \<subseteq> subformulas\<^sub>\<mu> \<phi>"
and
"Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and
1: "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
and
2: "\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
and
3: "\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
shows
"w \<Turnstile>\<^sub>n \<phi>"
proof -
from 1 obtain i where "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
by blast
moreover
from assms have "X \<subseteq> \<G>\<F> \<phi> w"
using X_\<G>\<F>_Y_\<F>\<G> by blast
then have "X \<subseteq> \<G>\<F> \<phi> (suffix i w)"
using \<G>\<F>_suffix by fast
ultimately have "suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)"
using GF_advice_a2 \<G>\<F>_af by metis
then show "w \<Turnstile>\<^sub>n \<phi>"
using af_ltl_continuation by force
qed
theorem master_theorem:
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow>
(\<exists>X \<subseteq> subformulas\<^sub>\<mu> \<phi>.
(\<exists>Y \<subseteq> subformulas\<^sub>\<nu> \<phi>.
(\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>)
\<and> (\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>))
\<and> (\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>))))"
by (metis master_theorem_ltr master_theorem_rtl)
subsection \<open>The Master Theorem on Languages\<close>
definition "L\<^sub>1 \<phi> X = {w. \<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>}"
definition "L\<^sub>2 X Y = {w. \<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)}"
definition "L\<^sub>3 X Y = {w. \<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)}"
corollary master_theorem_language:
"language_ltln \<phi> = \<Union> {L\<^sub>1 \<phi> X \<inter> L\<^sub>2 X Y \<inter> L\<^sub>3 X Y | X Y. X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<and> Y \<subseteq> subformulas\<^sub>\<nu> \<phi>}"
proof safe
fix w
assume "w \<in> language_ltln \<phi>"
then have "w \<Turnstile>\<^sub>n \<phi>"
unfolding language_ltln_def by simp
then obtain X Y where "X \<subseteq> subformulas\<^sub>\<mu> \<phi>" and "Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
and "\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
and "\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
using master_theorem_ltr by metis
then have "w \<in> L\<^sub>1 \<phi> X" and "w \<in> L\<^sub>2 X Y" and "w \<in> L\<^sub>3 X Y"
unfolding L\<^sub>1_def L\<^sub>2_def L\<^sub>3_def by simp+
then show "w \<in> \<Union> {L\<^sub>1 \<phi> X \<inter> L\<^sub>2 X Y \<inter> L\<^sub>3 X Y | X Y. X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<and> Y \<subseteq> subformulas\<^sub>\<nu> \<phi>}"
- using `X \<subseteq> subformulas\<^sub>\<mu> \<phi>` `Y \<subseteq> subformulas\<^sub>\<nu> \<phi>` by blast
+ using \<open>X \<subseteq> subformulas\<^sub>\<mu> \<phi>\<close> \<open>Y \<subseteq> subformulas\<^sub>\<nu> \<phi>\<close> by blast
next
fix w X Y
assume "X \<subseteq> subformulas\<^sub>\<mu> \<phi>" and "Y \<subseteq> subformulas\<^sub>\<nu> \<phi>"
and "w \<in> L\<^sub>1 \<phi> X" and "w \<in> L\<^sub>2 X Y" and "w \<in> L\<^sub>3 X Y"
then show "w \<in> language_ltln \<phi>"
unfolding language_ltln_def L\<^sub>1_def L\<^sub>2_def L\<^sub>3_def
using master_theorem_rtl by blast
qed
end
diff --git a/thys/LTL_Master_Theorem/Logical_Characterization/Restricted_Master_Theorem.thy b/thys/LTL_Master_Theorem/Logical_Characterization/Restricted_Master_Theorem.thy
--- a/thys/LTL_Master_Theorem/Logical_Characterization/Restricted_Master_Theorem.thy
+++ b/thys/LTL_Master_Theorem/Logical_Characterization/Restricted_Master_Theorem.thy
@@ -1,558 +1,558 @@
(*
Author: Salomon Sickert
Author: Benedikt Seidl
License: BSD
*)
section \<open>Master Theorem with Reduced Subformulas\<close>
theory Restricted_Master_Theorem
imports
Master_Theorem
begin
subsection \<open>Restricted Set of Subformulas\<close>
fun restricted_subformulas_inner :: "'a ltln \<Rightarrow> 'a ltln set"
where
"restricted_subformulas_inner (\<phi> and\<^sub>n \<psi>) = restricted_subformulas_inner \<phi> \<union> restricted_subformulas_inner \<psi>"
| "restricted_subformulas_inner (\<phi> or\<^sub>n \<psi>) = restricted_subformulas_inner \<phi> \<union> restricted_subformulas_inner \<psi>"
| "restricted_subformulas_inner (X\<^sub>n \<phi>) = restricted_subformulas_inner \<phi>"
| "restricted_subformulas_inner (\<phi> U\<^sub>n \<psi>) = subformulas\<^sub>\<nu> (\<phi> U\<^sub>n \<psi>) \<union> subformulas\<^sub>\<mu> (\<phi> U\<^sub>n \<psi>)"
| "restricted_subformulas_inner (\<phi> R\<^sub>n \<psi>) = restricted_subformulas_inner \<phi> \<union> restricted_subformulas_inner \<psi>"
| "restricted_subformulas_inner (\<phi> W\<^sub>n \<psi>) = restricted_subformulas_inner \<phi> \<union> restricted_subformulas_inner \<psi>"
| "restricted_subformulas_inner (\<phi> M\<^sub>n \<psi>) = subformulas\<^sub>\<nu> (\<phi> M\<^sub>n \<psi>) \<union> subformulas\<^sub>\<mu> (\<phi> M\<^sub>n \<psi>)"
| "restricted_subformulas_inner _ = {}"
fun restricted_subformulas :: "'a ltln \<Rightarrow> 'a ltln set"
where
"restricted_subformulas (\<phi> and\<^sub>n \<psi>) = restricted_subformulas \<phi> \<union> restricted_subformulas \<psi>"
| "restricted_subformulas (\<phi> or\<^sub>n \<psi>) = restricted_subformulas \<phi> \<union> restricted_subformulas \<psi>"
| "restricted_subformulas (X\<^sub>n \<phi>) = restricted_subformulas \<phi>"
| "restricted_subformulas (\<phi> U\<^sub>n \<psi>) = restricted_subformulas \<phi> \<union> restricted_subformulas \<psi>"
| "restricted_subformulas (\<phi> R\<^sub>n \<psi>) = restricted_subformulas \<phi> \<union> restricted_subformulas_inner \<psi>"
| "restricted_subformulas (\<phi> W\<^sub>n \<psi>) = restricted_subformulas_inner \<phi> \<union> restricted_subformulas \<psi>"
| "restricted_subformulas (\<phi> M\<^sub>n \<psi>) = restricted_subformulas \<phi> \<union> restricted_subformulas \<psi>"
| "restricted_subformulas _ = {}"
lemma GF_advice_restricted_subformulas_inner:
"restricted_subformulas_inner (\<phi>[X]\<^sub>\<nu>) = {}"
by (induction \<phi>) simp_all
lemma GF_advice_restricted_subformulas:
"restricted_subformulas (\<phi>[X]\<^sub>\<nu>) = {}"
by (induction \<phi>) (simp_all add: GF_advice_restricted_subformulas_inner)
lemma restricted_subformulas_inner_subset:
"restricted_subformulas_inner \<phi> \<subseteq> subformulas\<^sub>\<nu> \<phi> \<union> subformulas\<^sub>\<mu> \<phi>"
by (induction \<phi>) auto
lemma restricted_subformulas_subset':
"restricted_subformulas \<phi> \<subseteq> restricted_subformulas_inner \<phi>"
by (induction \<phi>) (insert restricted_subformulas_inner_subset, auto)
lemma restricted_subformulas_subset:
"restricted_subformulas \<phi> \<subseteq> subformulas\<^sub>\<nu> \<phi> \<union> subformulas\<^sub>\<mu> \<phi>"
using restricted_subformulas_inner_subset restricted_subformulas_subset' by auto
lemma restricted_subformulas_size:
"\<psi> \<in> restricted_subformulas \<phi> \<Longrightarrow> size \<psi> < size \<phi>"
proof -
have "\<And>\<phi>. restricted_subformulas_inner \<phi> \<subseteq> subfrmlsn \<phi>"
using restricted_subformulas_inner_subset subformulas\<^sub>\<mu>\<^sub>\<nu>_subfrmlsn by blast
then have inner: "\<And>\<psi> \<phi>. \<psi> \<in> restricted_subformulas_inner \<phi> \<Longrightarrow> size \<psi> \<le> size \<phi>"
using subfrmlsn_size dual_order.strict_implies_order
by blast
show "\<psi> \<in> restricted_subformulas \<phi> \<Longrightarrow> size \<psi> < size \<phi>"
by (induction \<phi> arbitrary: \<psi>) (fastforce dest: inner)+
qed
lemma restricted_subformulas_notin:
"\<phi> \<notin> restricted_subformulas \<phi>"
using restricted_subformulas_size by auto
lemma restricted_subformulas_superset:
"\<psi> \<in> restricted_subformulas \<phi> \<Longrightarrow> subformulas\<^sub>\<nu> \<psi> \<union> subformulas\<^sub>\<mu> \<psi> \<subseteq> restricted_subformulas \<phi>"
proof -
assume "\<psi> \<in> restricted_subformulas \<phi>"
then obtain \<chi> x where
"\<psi> \<in> restricted_subformulas_inner \<chi>" and "(x R\<^sub>n \<chi>) \<in> subformulas\<^sub>\<nu> \<phi> \<or> (\<chi> W\<^sub>n x) \<in> subformulas\<^sub>\<nu> \<phi> "
by (induction \<phi>) auto
moreover
have "\<And>\<psi>\<^sub>1 \<psi>\<^sub>2. (\<psi>\<^sub>1 R\<^sub>n \<psi>\<^sub>2) \<in> subformulas\<^sub>\<nu> \<phi> \<Longrightarrow> restricted_subformulas_inner \<psi>\<^sub>2 \<subseteq> restricted_subformulas \<phi>"
"\<And>\<psi>\<^sub>1 \<psi>\<^sub>2. (\<psi>\<^sub>1 W\<^sub>n \<psi>\<^sub>2) \<in> subformulas\<^sub>\<nu> \<phi> \<Longrightarrow> restricted_subformulas_inner \<psi>\<^sub>1 \<subseteq> restricted_subformulas \<phi>"
by (induction \<phi>) (simp_all; insert restricted_subformulas_subset', blast)+
moreover
have "subformulas\<^sub>\<nu> \<psi> \<union> subformulas\<^sub>\<mu> \<psi> \<subseteq> restricted_subformulas_inner \<chi>"
- using `\<psi> \<in> restricted_subformulas_inner \<chi>`
+ using \<open>\<psi> \<in> restricted_subformulas_inner \<chi>\<close>
proof (induction \<chi>)
case (Until_ltln \<chi>1 \<chi>2)
then show ?case
apply (cases "\<psi> = \<chi>1 U\<^sub>n \<chi>2")
apply auto[1]
apply simp
apply (cases "\<psi> \<in> subformulas\<^sub>\<nu> \<chi>1")
apply (meson le_supI1 le_supI2 subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subfrmlsn subformulas\<^sub>\<nu>_subset subset_eq subset_insertI2)
apply (cases "\<psi> \<in> subformulas\<^sub>\<nu> \<chi>2")
apply (meson le_supI1 le_supI2 subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subfrmlsn subformulas\<^sub>\<nu>_subset subset_eq subset_insertI2)
apply (cases "\<psi> \<in> subformulas\<^sub>\<mu> \<chi>1")
apply (metis (no_types, hide_lams) Un_insert_right subformulas\<^sub>\<mu>_subfrmlsn subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subset subsetD sup.coboundedI2 sup_commute)
apply simp
by (metis (no_types, hide_lams) Un_insert_right subformulas\<^sub>\<mu>_subfrmlsn subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subset subsetD sup.coboundedI2 sup_commute)
next
case (Release_ltln \<chi>1 \<chi>2)
then show ?case by simp blast
next
case (WeakUntil_ltln \<chi>1 \<chi>2)
then show ?case by simp blast
next
case (StrongRelease_ltln \<chi>1 \<chi>2)
then show ?case
apply (cases "\<psi> = \<chi>1 M\<^sub>n \<chi>2")
apply auto[1]
apply simp
apply (cases "\<psi> \<in> subformulas\<^sub>\<nu> \<chi>1")
apply (meson le_supI1 le_supI2 subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subfrmlsn subformulas\<^sub>\<nu>_subset subset_eq subset_insertI2)
apply (cases "\<psi> \<in> subformulas\<^sub>\<nu> \<chi>2")
apply (meson le_supI1 le_supI2 subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subfrmlsn subformulas\<^sub>\<nu>_subset subset_eq subset_insertI2)
apply (cases "\<psi> \<in> subformulas\<^sub>\<mu> \<chi>1")
apply (metis (no_types, hide_lams) Un_insert_right subformulas\<^sub>\<mu>_subfrmlsn subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subset subsetD sup.coboundedI2 sup_commute)
apply simp
by (metis (no_types, hide_lams) Un_insert_right subformulas\<^sub>\<mu>_subfrmlsn subformulas\<^sub>\<mu>_subset subformulas\<^sub>\<nu>_subset subsetD sup.coboundedI2 sup_commute)
qed auto
ultimately
show "subformulas\<^sub>\<nu> \<psi> \<union> subformulas\<^sub>\<mu> \<psi> \<subseteq> restricted_subformulas \<phi>"
by blast
qed
lemma restricted_subformulas_W_\<mu>:
"subformulas\<^sub>\<mu> \<phi> \<subseteq> restricted_subformulas (\<phi> W\<^sub>n \<psi>)"
by (induction \<phi>) auto
lemma restricted_subformulas_R_\<mu>:
"subformulas\<^sub>\<mu> \<psi> \<subseteq> restricted_subformulas (\<phi> R\<^sub>n \<psi>)"
by (induction \<psi>) auto
lemma restrict_af_letter:
"restricted_subformulas (af_letter \<phi> \<sigma>) = restricted_subformulas \<phi>"
proof (induction \<phi>)
case (Release_ltln \<phi>1 \<phi>2)
then show ?case
using restricted_subformulas_subset' by simp blast
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
then show ?case
using restricted_subformulas_subset' by simp blast
qed auto
lemma restrict_af:
"restricted_subformulas (af \<phi> w) = restricted_subformulas \<phi>"
by (induction w rule: rev_induct) (auto simp: restrict_af_letter)
subsection \<open>Restricted Master Theorem / Lemmas\<close>
lemma delay_2:
assumes "\<mu>_stable \<phi> w"
assumes "w \<Turnstile>\<^sub>n \<phi>"
shows "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>]\<^sub>\<nu>"
using assms
proof (induction \<phi> arbitrary: w)
case (Prop_ltln x)
then show ?case
by (metis GF_advice.simps(10) GF_advice_af prefix_suffix)
next
case (Nprop_ltln x)
then show ?case
by (metis GF_advice.simps(11) GF_advice_af prefix_suffix)
next
case (And_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 and\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have "?X1 \<subseteq> ?X" and "?X2 \<subseteq> ?X"
by auto
moreover
obtain i j where "suffix i w \<Turnstile>\<^sub>n af \<phi>1 (prefix i w)[?X1]\<^sub>\<nu>"
and "suffix j w \<Turnstile>\<^sub>n af \<phi>2 (prefix j w)[?X2]\<^sub>\<nu>"
using \<mu>_stable_subfrmlsn[OF \<open>\<mu>_stable (\<phi>1 and\<^sub>n \<phi>2) w\<close>] And_ltln by fastforce
ultimately
obtain k where "suffix k w \<Turnstile>\<^sub>n af \<phi>1 (prefix k w)[?X]\<^sub>\<nu>"
and "suffix k w \<Turnstile>\<^sub>n af \<phi>2 (prefix k w)[?X]\<^sub>\<nu>"
using GF_advice_sync_and GF_advice_monotone by blast
thus ?case
unfolding af_decompose semantics_ltln.simps(5) GF_advice.simps by blast
next
case (Or_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 and\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have "?X1 \<subseteq> ?X" and "?X2 \<subseteq> ?X"
by auto
moreover
obtain i j where "suffix i w \<Turnstile>\<^sub>n af \<phi>1 (prefix i w)[?X1]\<^sub>\<nu> \<or> suffix j w \<Turnstile>\<^sub>n af \<phi>2 (prefix j w)[?X2]\<^sub>\<nu>"
using \<mu>_stable_subfrmlsn[OF \<open>\<mu>_stable (\<phi>1 or\<^sub>n \<phi>2) w\<close>] Or_ltln by fastforce
ultimately
obtain k where "suffix k w \<Turnstile>\<^sub>n af \<phi>1 (prefix k w)[?X]\<^sub>\<nu> \<or> suffix k w \<Turnstile>\<^sub>n af \<phi>2 (prefix k w)[?X]\<^sub>\<nu>"
using GF_advice_monotone by blast
thus ?case
unfolding af_decompose semantics_ltln.simps(6) GF_advice.simps by auto
next
case (Next_ltln \<phi>)
then have stable: "\<mu>_stable \<phi> (suffix 1 w)"
and suffix: "suffix 1 w \<Turnstile>\<^sub>n \<phi>"
using \<F>_suffix \<G>\<F>_\<F>_subset \<G>\<F>_suffix
by (simp_all add: \<mu>_stable_def) fast
show ?case
by (metis (no_types, lifting) Next_ltln.IH[OF stable suffix, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix] One_nat_def add.commute af_simps(3) foldl_Nil foldl_append restricted_subformulas.simps(3) subsequence_append subsequence_singleton)
next
case (Until_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 U\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have stable_1: "\<And>i. \<mu>_stable \<phi>1 (suffix i w)"
and stable_2: "\<And>i. \<mu>_stable \<phi>2 (suffix i w)"
using \<mu>_stable_subfrmlsn[OF Until_ltln.prems(1)] by (simp add: \<mu>_stable_suffix)+
obtain i where "\<And>j. j < i \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<phi>1" and "suffix i w \<Turnstile>\<^sub>n \<phi>2"
using Until_ltln by auto
then have "\<And>j. j < i \<Longrightarrow> \<exists>k. suffix (j + k) w \<Turnstile>\<^sub>n af \<phi>1 (w [j \<rightarrow> k + j])[?X1]\<^sub>\<nu>"
and "\<exists>k. suffix (i + k) w \<Turnstile>\<^sub>n af \<phi>2 (w [i \<rightarrow> k + i])[?X2]\<^sub>\<nu>"
using Until_ltln.IH(1)[OF stable_1, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
using Until_ltln.IH(2)[OF stable_2, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
by blast+
moreover
have "?X1 \<subseteq> ?X"
and "?X2 \<subseteq> ?X"
by auto
ultimately
obtain k where "k \<ge> i"
and "suffix k w \<Turnstile>\<^sub>n af \<phi>2 (w [i \<rightarrow> k])[?X]\<^sub>\<nu>"
and "\<And>j. j < i \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af \<phi>1 (w [j \<rightarrow> k])[?X]\<^sub>\<nu>"
using GF_advice_sync_less[of i w \<phi>1 ?X \<phi>2] GF_advice_monotone[of _ ?X] by meson
hence "suffix (Suc k) w \<Turnstile>\<^sub>n af (\<phi>1 U\<^sub>n \<phi>2) (prefix (Suc k) w)[?X]\<^sub>\<nu>"
by (rule af_subsequence_U_GF_advice)
then show ?case
by blast
next
case (WeakUntil_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 W\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have stable_1: "\<And>i. \<mu>_stable \<phi>1 (suffix i w)"
and stable_2: "\<And>i. \<mu>_stable \<phi>2 (suffix i w)"
using \<mu>_stable_subfrmlsn[OF WeakUntil_ltln.prems(1)] by (simp add: \<mu>_stable_suffix)+
{
assume Until_ltln: "w \<Turnstile>\<^sub>n \<phi>1 U\<^sub>n \<phi>2"
then obtain i where "\<And>j. j < i \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<phi>1" and "suffix i w \<Turnstile>\<^sub>n \<phi>2"
by auto
then have "\<And>j. j < i \<Longrightarrow> \<exists>k. suffix (j + k) w \<Turnstile>\<^sub>n af \<phi>1 (w [j \<rightarrow> k + j])[?X1]\<^sub>\<nu>"
and "\<exists>k. suffix (i + k) w \<Turnstile>\<^sub>n af \<phi>2 (w [i \<rightarrow> k + i])[?X2]\<^sub>\<nu>"
using WeakUntil_ltln.IH(1)[OF stable_1, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
using WeakUntil_ltln.IH(2)[OF stable_2, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
by blast+
moreover
have "?X1 \<subseteq> ?X"
and "?X2 \<subseteq> ?X"
using restricted_subformulas_subset' by force+
ultimately
obtain k where "k \<ge> i"
and "suffix k w \<Turnstile>\<^sub>n af \<phi>2 (w [i \<rightarrow> k])[?X]\<^sub>\<nu>"
and "\<And>j. j < i \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af \<phi>1 (w [j \<rightarrow> k])[?X]\<^sub>\<nu>"
using GF_advice_sync_less[of i w \<phi>1 ?X \<phi>2] GF_advice_monotone[of _ ?X] by meson
hence "suffix (Suc k) w \<Turnstile>\<^sub>n af (\<phi>1 W\<^sub>n \<phi>2) (prefix (Suc k) w)[?X]\<^sub>\<nu>"
by (rule af_subsequence_W_GF_advice)
hence ?case
by blast
}
moreover
{
assume Globally_ltln: "w \<Turnstile>\<^sub>n G\<^sub>n \<phi>1"
{
fix j
have "suffix j w \<Turnstile>\<^sub>n \<phi>1"
using Globally_ltln by simp
then have "suffix j w \<Turnstile>\<^sub>n \<phi>1[{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)}]\<^sub>\<nu>"
by (metis stable_1 GF_advice_a1 \<G>\<F>_suffix \<mu>_stable_def \<G>\<F>_elim mem_Collect_eq subsetI)
then have "suffix j w \<Turnstile>\<^sub>n \<phi>1[?X]\<^sub>\<nu>"
by (metis GF_advice_inter restricted_subformulas_W_\<mu> le_infI2)
}
then have "w \<Turnstile>\<^sub>n (\<phi>1 W\<^sub>n \<phi>2)[?X]\<^sub>\<nu>"
by simp
then have ?case
using GF_advice_af_2 by blast
}
ultimately
show ?case
using WeakUntil_ltln.prems(2) ltln_weak_to_strong(1) by blast
next
case (Release_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 R\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have stable_1: "\<And>i. \<mu>_stable \<phi>1 (suffix i w)"
and stable_2: "\<And>i. \<mu>_stable \<phi>2 (suffix i w)"
using \<mu>_stable_subfrmlsn[OF Release_ltln.prems(1)] by (simp add: \<mu>_stable_suffix)+
{
assume Until_ltln: "w \<Turnstile>\<^sub>n \<phi>1 M\<^sub>n \<phi>2"
then obtain i where "\<And>j. j \<le> i \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<phi>2" and "suffix i w \<Turnstile>\<^sub>n \<phi>1"
by auto
then have "\<And>j. j \<le> i \<Longrightarrow> \<exists>k. suffix (j + k) w \<Turnstile>\<^sub>n af \<phi>2 (w [j \<rightarrow> k + j])[?X2]\<^sub>\<nu>"
and "\<exists>k. suffix (i + k) w \<Turnstile>\<^sub>n af \<phi>1 (w [i \<rightarrow> k + i])[?X1]\<^sub>\<nu>"
using Release_ltln.IH(1)[OF stable_1, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
using Release_ltln.IH(2)[OF stable_2, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
by blast+
moreover
have "?X1 \<subseteq> ?X"
and "?X2 \<subseteq> ?X"
using restricted_subformulas_subset' by force+
ultimately
obtain k where "k \<ge> i"
and "suffix k w \<Turnstile>\<^sub>n af \<phi>1 (w [i \<rightarrow> k])[?X]\<^sub>\<nu>"
and "\<And>j. j \<le> i \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af \<phi>2 (w [j \<rightarrow> k])[?X]\<^sub>\<nu>"
using GF_advice_sync_lesseq[of i w \<phi>2 ?X \<phi>1] GF_advice_monotone[of _ ?X] by meson
hence "suffix (Suc k) w \<Turnstile>\<^sub>n af (\<phi>1 R\<^sub>n \<phi>2) (prefix (Suc k) w)[?X]\<^sub>\<nu>"
by (rule af_subsequence_R_GF_advice)
hence ?case
by blast
}
moreover
{
assume Globally_ltln: "w \<Turnstile>\<^sub>n G\<^sub>n \<phi>2"
{
fix j
have "suffix j w \<Turnstile>\<^sub>n \<phi>2"
using Globally_ltln by simp
then have "suffix j w \<Turnstile>\<^sub>n \<phi>2[{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)}]\<^sub>\<nu>"
by (metis stable_2 GF_advice_a1 \<G>\<F>_suffix \<mu>_stable_def \<G>\<F>_elim mem_Collect_eq subsetI)
then have "suffix j w \<Turnstile>\<^sub>n \<phi>2[?X]\<^sub>\<nu>"
by (metis GF_advice_inter restricted_subformulas_R_\<mu> le_infI2)
}
then have "w \<Turnstile>\<^sub>n (\<phi>1 R\<^sub>n \<phi>2)[?X]\<^sub>\<nu>"
by simp
then have ?case
using GF_advice_af_2 by blast
}
ultimately
show ?case
using Release_ltln.prems(2) ltln_weak_to_strong(3) by blast
next
case (StrongRelease_ltln \<phi>1 \<phi>2)
let ?X = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas (\<phi>1 M\<^sub>n \<phi>2)"
let ?X1 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>1"
let ?X2 = "{\<psi>. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>)} \<inter> restricted_subformulas \<phi>2"
have stable_1: "\<And>i. \<mu>_stable \<phi>1 (suffix i w)"
and stable_2: "\<And>i. \<mu>_stable \<phi>2 (suffix i w)"
using \<mu>_stable_subfrmlsn[OF StrongRelease_ltln.prems(1)] by (simp add: \<mu>_stable_suffix)+
obtain i where "\<And>j. j \<le> i \<Longrightarrow> suffix j w \<Turnstile>\<^sub>n \<phi>2" and "suffix i w \<Turnstile>\<^sub>n \<phi>1"
using StrongRelease_ltln by auto
then have "\<And>j. j \<le> i \<Longrightarrow> \<exists>k. suffix (j + k) w \<Turnstile>\<^sub>n af \<phi>2 (w [j \<rightarrow> k + j])[?X2]\<^sub>\<nu>"
and "\<exists>k. suffix (i + k) w \<Turnstile>\<^sub>n af \<phi>1 (w [i \<rightarrow> k + i])[?X1]\<^sub>\<nu>"
using StrongRelease_ltln.IH(1)[OF stable_1, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
using StrongRelease_ltln.IH(2)[OF stable_2, unfolded suffix_suffix prefix_suffix_subsequence GF_suffix]
by blast+
moreover
have "?X1 \<subseteq> ?X"
and "?X2 \<subseteq> ?X"
by auto
ultimately
obtain k where "k \<ge> i"
and "suffix k w \<Turnstile>\<^sub>n af \<phi>1 (w [i \<rightarrow> k])[?X]\<^sub>\<nu>"
and "\<And>j. j \<le> i \<Longrightarrow> suffix k w \<Turnstile>\<^sub>n af \<phi>2 (w [j \<rightarrow> k])[?X]\<^sub>\<nu>"
using GF_advice_sync_lesseq[of i w \<phi>2 ?X \<phi>1] GF_advice_monotone[of _ ?X] by meson
hence "suffix (Suc k) w \<Turnstile>\<^sub>n af (\<phi>1 M\<^sub>n \<phi>2) (prefix (Suc k) w)[?X]\<^sub>\<nu>"
by (rule af_subsequence_M_GF_advice)
then show ?case
by blast
qed simp+
theorem master_theorem_restricted:
"w \<Turnstile>\<^sub>n \<phi> \<longleftrightarrow>
(\<exists>X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi>.
(\<exists>Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>.
(\<exists>i. (suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>)
\<and> (\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>))
\<and> (\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)))))"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
obtain i where "\<mu>_stable \<phi> (suffix i w)"
by (metis MOST_nat less_Suc_eq suffix_\<mu>_stable)
hence stable: "\<mu>_stable (af \<phi> (prefix i w)) (suffix i w)"
by (simp add: \<F>_af \<G>\<F>_af \<mu>_stable_def)
let ?\<phi>' = "af \<phi> (prefix i w)"
let ?X' = "\<G>\<F> \<phi> w \<inter> restricted_subformulas \<phi>"
let ?Y' = "\<F>\<G> \<phi> w \<inter> restricted_subformulas \<phi>"
have 1: "suffix i w \<Turnstile>\<^sub>n ?\<phi>'"
- using `?lhs` af_ltl_continuation by force
+ using \<open>?lhs\<close> af_ltl_continuation by force
have 2: "\<And>j. af (af \<phi> (prefix i w)) (prefix j (suffix i w)) = af \<phi> (prefix (i + j) w)"
by (simp add: subsequence_append)
have 3: "\<G>\<F> \<phi> w = \<G>\<F> \<phi> (suffix i w)"
using \<G>\<F>_af \<G>\<F>_suffix by blast
have "\<exists>j. suffix (i + j) w \<Turnstile>\<^sub>n af (?\<phi>') (prefix j (suffix i w))[?X']\<^sub>\<nu>"
using delay_2[OF stable 1] unfolding suffix_suffix 2 restrict_af 3 unfolding \<G>\<F>_semantics'
by (metis (no_types, lifting) GF_advice_inter_subformulas af_subformulas\<^sub>\<mu> inf_assoc inf_commute)
hence "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[?X']\<^sub>\<nu>"
using 2 by auto
moreover
{
fix \<psi>
have "\<And>X. \<psi> \<in> restricted_subformulas \<phi> \<Longrightarrow> \<psi>[X \<inter> restricted_subformulas \<phi>]\<^sub>\<mu> = \<psi>[X]\<^sub>\<mu>"
by (metis le_supE restricted_subformulas_superset FG_advice_inter inf.coboundedI2)
hence "\<psi> \<in> ?X' \<Longrightarrow> w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[?Y']\<^sub>\<mu>)"
using \<G>\<F>_implies_GF by force
}
moreover
{
fix \<psi>
have "\<And>X. \<psi> \<in> restricted_subformulas \<phi> \<Longrightarrow> \<psi>[X \<inter> restricted_subformulas \<phi>]\<^sub>\<nu> = \<psi>[X]\<^sub>\<nu>"
by (metis le_supE restricted_subformulas_superset GF_advice_inter inf.coboundedI2)
hence "\<psi> \<in> ?Y' \<Longrightarrow> w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[?X']\<^sub>\<nu>)"
using \<F>\<G>_implies_FG by force
}
moreover
have "?X' \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi>"
using \<G>\<F>_subformulas\<^sub>\<mu> by blast
moreover
have "?Y' \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>"
using \<F>\<G>_subformulas\<^sub>\<nu> by blast
ultimately show ?rhs
by meson
qed (insert master_theorem, fast)
corollary master_theorem_restricted_language:
"language_ltln \<phi> = \<Union> {L\<^sub>1 \<phi> X \<inter> L\<^sub>2 X Y \<inter> L\<^sub>3 X Y | X Y. X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi> \<and> Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>}"
proof safe
fix w
assume "w \<in> language_ltln \<phi>"
then have "w \<Turnstile>\<^sub>n \<phi>"
unfolding language_ltln_def by simp
then obtain X Y where
1: "X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi>"
and 2: "Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>"
and "\<exists>i. suffix i w \<Turnstile>\<^sub>n af \<phi> (prefix i w)[X]\<^sub>\<nu>"
and "\<forall>\<psi> \<in> X. w \<Turnstile>\<^sub>n G\<^sub>n (F\<^sub>n \<psi>[Y]\<^sub>\<mu>)"
and "\<forall>\<psi> \<in> Y. w \<Turnstile>\<^sub>n F\<^sub>n (G\<^sub>n \<psi>[X]\<^sub>\<nu>)"
using master_theorem_restricted by metis
then have "w \<in> L\<^sub>1 \<phi> X" and "w \<in> L\<^sub>2 X Y" and "w \<in> L\<^sub>3 X Y"
unfolding L\<^sub>1_def L\<^sub>2_def L\<^sub>3_def by simp+
then show "w \<in> \<Union> {L\<^sub>1 \<phi> X \<inter> L\<^sub>2 X Y \<inter> L\<^sub>3 X Y | X Y. X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi> \<and> Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>}"
using 1 2 by blast
next
fix w X Y
assume "X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi>" and "Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi>"
and "w \<in> L\<^sub>1 \<phi> X" and "w \<in> L\<^sub>2 X Y" and "w \<in> L\<^sub>3 X Y"
then show "w \<in> language_ltln \<phi>"
unfolding language_ltln_def L\<^sub>1_def L\<^sub>2_def L\<^sub>3_def
using master_theorem_restricted by blast
qed
subsection \<open>Definitions with Lists for Code Export\<close>
definition restricted_advice_sets :: "'a ltln \<Rightarrow> ('a ltln list \<times> 'a ltln list) list"
where
"restricted_advice_sets \<phi> = List.product (subseqs (List.filter (\<lambda>x. x \<in> restricted_subformulas \<phi>) (subformulas\<^sub>\<mu>_list \<phi>))) (subseqs (List.filter (\<lambda>x. x \<in> restricted_subformulas \<phi>) (subformulas\<^sub>\<nu>_list \<phi>)))"
lemma subseqs_subformulas\<^sub>\<mu>_restricted_list:
"X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi> \<longleftrightarrow> (\<exists>xs. X = set xs \<and> xs \<in> set (subseqs (List.filter (\<lambda>x. x \<in> restricted_subformulas \<phi>) (subformulas\<^sub>\<mu>_list \<phi>))))"
by (metis in_set_subseqs inf_commute inter_set_filter subformulas\<^sub>\<mu>_list_set subset_subseq)
lemma subseqs_subformulas\<^sub>\<nu>_restricted_list:
"Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi> \<longleftrightarrow> (\<exists>ys. Y = set ys \<and> ys \<in> set (subseqs (List.filter (\<lambda>x. x \<in> restricted_subformulas \<phi>) (subformulas\<^sub>\<nu>_list \<phi>))))"
by (metis in_set_subseqs inf_commute inter_set_filter subformulas\<^sub>\<nu>_list_set subset_subseq)
lemma restricted_advice_sets_subformulas:
"X \<subseteq> subformulas\<^sub>\<mu> \<phi> \<inter> restricted_subformulas \<phi> \<and> Y \<subseteq> subformulas\<^sub>\<nu> \<phi> \<inter> restricted_subformulas \<phi> \<longleftrightarrow> (\<exists>xs ys. X = set xs \<and> Y = set ys \<and> (xs, ys) \<in> set (restricted_advice_sets \<phi>))"
unfolding restricted_advice_sets_def set_product subseqs_subformulas\<^sub>\<mu>_restricted_list subseqs_subformulas\<^sub>\<nu>_restricted_list by blast
lemma restricted_advice_sets_not_empty:
"restricted_advice_sets \<phi> \<noteq> []"
unfolding restricted_advice_sets_def using subseqs_not_empty product_not_empty by blast
end
diff --git a/thys/Linear_Inequalities/Convex_Hull.thy b/thys/Linear_Inequalities/Convex_Hull.thy
--- a/thys/Linear_Inequalities/Convex_Hull.thy
+++ b/thys/Linear_Inequalities/Convex_Hull.thy
@@ -1,616 +1,616 @@
section \<open>Convex Hulls\<close>
text \<open>We define the notion of convex hull of a set or list of vectors and derive basic
properties thereof.\<close>
theory Convex_Hull
imports Cone
begin
context gram_schmidt
begin
definition "convex_lincomb c Vs b = (nonneg_lincomb c Vs b \<and> sum c Vs = 1)"
definition "convex_lincomb_list c Vs b = (nonneg_lincomb_list c Vs b \<and> sum c {0..<length Vs} = 1)"
definition "convex_hull Vs = {x. \<exists> Ws c. finite Ws \<and> Ws \<subseteq> Vs \<and> convex_lincomb c Ws x}"
lemma convex_hull_carrier[intro]: "Vs \<subseteq> carrier_vec n \<Longrightarrow> convex_hull Vs \<subseteq> carrier_vec n"
unfolding convex_hull_def convex_lincomb_def nonneg_lincomb_def by auto
lemma convex_hull_mono: "Vs \<subseteq> Ws \<Longrightarrow> convex_hull Vs \<subseteq> convex_hull Ws"
unfolding convex_hull_def by auto
lemma convex_lincomb_empty[simp]: "\<not> (convex_lincomb c {} x)"
unfolding convex_lincomb_def by simp
lemma set_in_convex_hull:
assumes "A \<subseteq> carrier_vec n"
shows "A \<subseteq> convex_hull A"
proof
fix a
assume "a \<in> A"
hence acarr: "a \<in> carrier_vec n" using assms by auto
hence "convex_lincomb (\<lambda> x. 1) {a} a " unfolding convex_lincomb_def
by (auto simp: nonneg_lincomb_def lincomb_def)
then show "a \<in> convex_hull A" using \<open>a \<in> A\<close> unfolding convex_hull_def by auto
qed
lemma convex_hull_empty[simp]:
"convex_hull {} = {}"
"A \<subseteq> carrier_vec n \<Longrightarrow> convex_hull A = {} \<longleftrightarrow> A = {}"
proof -
show "convex_hull {} = {}" unfolding convex_hull_def by auto
then show "A \<subseteq> carrier_vec n \<Longrightarrow> convex_hull A = {} \<longleftrightarrow> A = {}"
using set_in_convex_hull[of A] by auto
qed
lemma convex_hull_bound: assumes XBnd: "X \<subseteq> Bounded_vec Bnd"
and X: "X \<subseteq> carrier_vec n"
shows "convex_hull X \<subseteq> Bounded_vec Bnd"
proof
fix x
assume "x \<in> convex_hull X"
from this[unfolded convex_hull_def]
obtain Y c where fin: "finite Y" and YX: "Y \<subseteq> X" and cx: "convex_lincomb c Y x" by auto
from cx[unfolded convex_lincomb_def nonneg_lincomb_def]
have x: "x = lincomb c Y" and sum: "sum c Y = 1" and c0: "\<And> y. y \<in> Y \<Longrightarrow> c y \<ge> 0" by auto
from YX X XBnd have Y: "Y \<subseteq> carrier_vec n" and YBnd: "Y \<subseteq> Bounded_vec Bnd" by auto
from x Y have dim: "dim_vec x = n" by auto
show "x \<in> Bounded_vec Bnd" unfolding Bounded_vec_def mem_Collect_eq dim
proof (intro allI impI)
fix i
assume i: "i < n"
have "abs (x $ i) = abs (\<Sum>x\<in>Y. c x * x $ i)" unfolding x
by (subst lincomb_index[OF i Y], auto)
also have "\<dots> \<le> (\<Sum>x\<in>Y. abs (c x * x $ i))" by auto
also have "\<dots> = (\<Sum>x\<in>Y. abs (c x) * abs (x $ i))" by (simp add: abs_mult)
also have "\<dots> \<le> (\<Sum>x\<in>Y. abs (c x) * Bnd)"
by (intro sum_mono mult_left_mono, insert YBnd[unfolded Bounded_vec_def] i Y, force+)
also have "\<dots> = (\<Sum>x\<in>Y. abs (c x)) * Bnd"
by (simp add: sum_distrib_right)
also have "(\<Sum>x\<in>Y. abs (c x)) = (\<Sum>x\<in>Y. c x)"
by (rule sum.cong, insert c0, auto)
also have "\<dots> = 1" by fact
finally show "\<bar>x $ i\<bar> \<le> Bnd" by auto
qed
qed
definition "convex_hull_list Vs = {x. \<exists> c. convex_lincomb_list c Vs x}"
lemma lincomb_list_elem:
"set Vs \<subseteq> carrier_vec n \<Longrightarrow>
lincomb_list (\<lambda> j. if i=j then 1 else 0) Vs = (if i < length Vs then Vs ! i else 0\<^sub>v n)"
proof (induction Vs rule: rev_induct)
case (snoc x Vs)
have x: "x \<in> carrier_vec n" and Vs: "set Vs \<subseteq> carrier_vec n" using snoc.prems by auto
let ?f = "\<lambda> j. if i = j then 1 else 0"
have "lincomb_list ?f (Vs @ [x]) = lincomb_list ?f Vs + ?f (length Vs) \<cdot>\<^sub>v x"
using x Vs by simp
also have "\<dots> = (if i < length (Vs @ [x]) then (Vs @ [x]) ! i else 0\<^sub>v n)" (is ?goal)
using less_linear[of i "length Vs"]
proof (elim disjE)
assume i: "i < length Vs"
have "lincomb_list (\<lambda>j. if i = j then 1 else 0) Vs = Vs ! i"
using snoc.IH[OF Vs] i by auto
moreover have "(if i = length Vs then 1 else 0) \<cdot>\<^sub>v x = 0\<^sub>v n" using i x by auto
moreover have "(if i < length (Vs @ [x]) then (Vs @ [x]) ! i else 0\<^sub>v n) = Vs ! i"
using i append_Cons_nth_left by fastforce
ultimately show ?goal using Vs i lincomb_list_carrier M.r_zero by metis
next
assume i: "i = length Vs"
have "lincomb_list (\<lambda>j. if i = j then 1 else 0) Vs = 0\<^sub>v n"
using snoc.IH[OF Vs] i by auto
moreover have "(if i = length Vs then 1 else 0) \<cdot>\<^sub>v x = x" using i x by auto
moreover have "(if i < length (Vs @ [x]) then (Vs @ [x]) ! i else 0\<^sub>v n) = x"
using i append_Cons_nth_left by simp
ultimately show ?goal using x by simp
next
assume i: "i > length Vs"
have "lincomb_list (\<lambda>j. if i = j then 1 else 0) Vs = 0\<^sub>v n"
using snoc.IH[OF Vs] i by auto
moreover have "(if i = length Vs then 1 else 0) \<cdot>\<^sub>v x = 0\<^sub>v n" using i x by auto
moreover have "(if i < length (Vs @ [x]) then (Vs @ [x]) ! i else 0\<^sub>v n) = 0\<^sub>v n"
using i by simp
ultimately show ?goal by simp
qed
finally show ?case by auto
qed simp
lemma set_in_convex_hull_list: fixes Vs :: "'a vec list"
assumes "set Vs \<subseteq> carrier_vec n"
shows "set Vs \<subseteq> convex_hull_list Vs"
proof
fix x assume "x \<in> set Vs"
then obtain i where i: "i < length Vs"
and x: "x = Vs ! i" using set_conv_nth[of Vs] by auto
let ?f = "\<lambda> j. if i = j then 1 else 0 :: 'a"
have "lincomb_list ?f Vs = x" using i x lincomb_list_elem[OF assms] by auto
moreover have "\<forall> j < length Vs. ?f j \<ge> 0" by auto
moreover have "sum ?f {0..<length Vs} = 1" using i by simp
ultimately show "x \<in> convex_hull_list Vs"
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
qed
lemma convex_hull_list_combination:
assumes Vs: "set Vs \<subseteq> carrier_vec n"
and x: "x \<in> convex_hull_list Vs"
and y: "y \<in> convex_hull_list Vs"
and l0: "0 \<le> l" and l1: "l \<le> 1"
shows "l \<cdot>\<^sub>v x + (1 - l) \<cdot>\<^sub>v y \<in> convex_hull_list Vs"
proof -
from x obtain cx where x: "lincomb_list cx Vs = x" and cx0: "\<forall> i < length Vs. cx i \<ge> 0"
and cx1: "sum cx {0..<length Vs} = 1"
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
from y obtain cy where y: "lincomb_list cy Vs = y" and cy0: "\<forall> i < length Vs. cy i \<ge> 0"
and cy1: "sum cy {0..<length Vs} = 1"
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
let ?c = "\<lambda> i. l * cx i + (1 - l) * cy i"
have "set Vs \<subseteq> carrier_vec n \<Longrightarrow>
lincomb_list ?c Vs = l \<cdot>\<^sub>v lincomb_list cx Vs + (1 - l) \<cdot>\<^sub>v lincomb_list cy Vs"
proof (induction Vs rule: rev_induct)
case (snoc v Vs)
have v: "v \<in> carrier_vec n" and Vs: "set Vs \<subseteq> carrier_vec n"
using snoc.prems by auto
have "lincomb_list ?c (Vs @ [v]) = lincomb_list ?c Vs + ?c (length Vs) \<cdot>\<^sub>v v"
using snoc.prems by auto
also have "lincomb_list ?c Vs =
l \<cdot>\<^sub>v lincomb_list cx Vs + (1 - l) \<cdot>\<^sub>v lincomb_list cy Vs"
by (rule snoc.IH[OF Vs])
also have "?c (length Vs) \<cdot>\<^sub>v v =
l \<cdot>\<^sub>v (cx (length Vs) \<cdot>\<^sub>v v) + (1 - l) \<cdot>\<^sub>v (cy (length Vs) \<cdot>\<^sub>v v)"
using add_smult_distrib_vec smult_smult_assoc by metis
also have "l \<cdot>\<^sub>v lincomb_list cx Vs + (1 - l) \<cdot>\<^sub>v lincomb_list cy Vs + \<dots> =
l \<cdot>\<^sub>v (lincomb_list cx Vs + cx (length Vs) \<cdot>\<^sub>v v) +
(1 - l) \<cdot>\<^sub>v (lincomb_list cy Vs + cy (length Vs) \<cdot>\<^sub>v v)"
using lincomb_list_carrier[OF Vs] v
by (simp add: M.add.m_assoc M.add.m_lcomm smult_r_distr)
finally show ?case using Vs v by simp
qed simp
hence "lincomb_list ?c Vs = l \<cdot>\<^sub>v x + (1 - l) \<cdot>\<^sub>v y" using Vs x y by simp
moreover have "\<forall> i < length Vs. ?c i \<ge> 0" using cx0 cy0 l0 l1 by simp
moreover have "sum ?c {0..<length Vs} = 1"
proof(simp add: sum.distrib)
have "(\<Sum>i = 0..<length Vs. (1 - l) * cy i) = (1 - l) * sum cy {0..<length Vs}"
using sum_distrib_left by metis
moreover have "(\<Sum>i = 0..<length Vs. l * cx i) = l * sum cx {0..<length Vs}"
using sum_distrib_left by metis
ultimately show "(\<Sum>i = 0..<length Vs. l * cx i) + (\<Sum>i = 0..<length Vs. (1 - l) * cy i) = 1"
using cx1 cy1 by simp
qed
ultimately show ?thesis
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
qed
lemma convex_hull_list_mono:
assumes "set Ws \<subseteq> carrier_vec n"
shows "set Vs \<subseteq> set Ws \<Longrightarrow> convex_hull_list Vs \<subseteq> convex_hull_list Ws"
proof (standard, induction Vs)
case Nil
from Nil(2) show ?case unfolding convex_hull_list_def convex_lincomb_list_def by auto
next
case (Cons v Vs)
have v: "v \<in> set Ws" and Vs: "set Vs \<subseteq> set Ws" using Cons.prems(1) by auto
hence v1: "v \<in> convex_hull_list Ws" using set_in_convex_hull_list[OF assms] by auto
from Cons.prems(2) obtain c
where x: "lincomb_list c (v # Vs) = x" and c0: "\<forall> i < length Vs + 1. c i \<ge> 0"
and c1: "sum c {0..<length Vs + 1} = 1"
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
have x: "x = c 0 \<cdot>\<^sub>v v + lincomb_list (c \<circ> Suc) Vs" using Vs v assms x by auto
show ?case proof (cases)
assume P: "c 0 = 1"
hence "sum (c \<circ> Suc) {0..<length Vs} = 0"
using sum.atLeast0_lessThan_Suc_shift c1
by (metis One_nat_def R.show_r_zero add.right_neutral add_Suc_right)
moreover have "\<And> i. i \<in> {0..<length Vs} \<Longrightarrow> (c \<circ> Suc) i \<ge> 0"
using c0 by simp
ultimately have "\<forall> i \<in> {0..<length Vs}. (c \<circ> Suc) i = 0"
using sum_nonneg_eq_0_iff by blast
hence "\<And> i. i < length Vs \<Longrightarrow> (c \<circ> Suc) i \<cdot>\<^sub>v Vs ! i = 0\<^sub>v n"
using Vs assms by (simp add: subset_code(1))
hence "lincomb_list (c \<circ> Suc) Vs = 0\<^sub>v n"
using lincomb_list_eq_0 by simp
hence "x = v" using P x v assms by auto
thus ?case using v1 by auto
next
assume P: "c 0 \<noteq> 1"
have c1: "c 0 + sum (c \<circ> Suc) {0..<length Vs} = 1"
using sum.atLeast0_lessThan_Suc_shift[of c] c1 by simp
have "sum (c \<circ> Suc) {0..<length Vs} \<ge> 0" by (rule sum_nonneg, insert c0, simp)
hence "c 0 < 1" using P c1 by auto
let ?c' = "\<lambda> i. 1 / (1 - c 0) * (c \<circ> Suc) i"
have "sum ?c' {0..<length Vs} = 1 / (1 - c 0) * sum (c \<circ> Suc) {0..<length Vs}"
using c1 P sum_distrib_left by metis
hence "sum ?c' {0..<length Vs} = 1" using P c1 by simp
- moreover have "\<forall> i < length Vs. ?c' i \<ge> 0" using c0 `c 0 < 1` by simp
+ moreover have "\<forall> i < length Vs. ?c' i \<ge> 0" using c0 \<open>c 0 < 1\<close> by simp
ultimately have c': "lincomb_list ?c' Vs \<in> convex_hull_list Ws"
using Cons.IH[OF Vs]
convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by blast
have "lincomb_list ?c' Vs = 1 / (1 - c 0) \<cdot>\<^sub>v lincomb_list (c \<circ> Suc) Vs"
by(rule lincomb_list_smult, insert Vs assms, auto)
hence "(1 - c 0) \<cdot>\<^sub>v lincomb_list ?c' Vs = lincomb_list (c \<circ> Suc) Vs"
using P by auto
hence "x = c 0 \<cdot>\<^sub>v v + (1 - c 0) \<cdot>\<^sub>v lincomb_list ?c' Vs" using x by auto
thus "x \<in> convex_hull_list Ws"
- using convex_hull_list_combination[OF assms v1 c'] c0 `c 0 < 1`
+ using convex_hull_list_combination[OF assms v1 c'] c0 \<open>c 0 < 1\<close>
by simp
qed
qed
lemma convex_hull_list_eq_set:
"set Vs \<subseteq> carrier_vec n \<Longrightarrow> set Vs = set Ws \<Longrightarrow> convex_hull_list Vs = convex_hull_list Ws"
using convex_hull_list_mono by blast
lemma find_indices_empty: "(find_indices x Vs = []) = (x \<notin> set Vs)"
proof (induction Vs rule: rev_induct)
case (snoc v Vs)
show ?case
proof
assume "find_indices x (Vs @ [v]) = []"
hence "x \<noteq> v \<and> find_indices x Vs = []" by auto
thus "x \<notin> set (Vs @ [v])" using snoc by simp
next
assume "x \<notin> set (Vs @ [v])"
hence "x \<noteq> v \<and> find_indices x Vs = []" using snoc by auto
thus "find_indices x (Vs @ [v]) = []" by simp
qed
qed simp
lemma distinct_list_find_indices:
shows "\<lbrakk> i < length Vs; Vs ! i = x; distinct Vs \<rbrakk> \<Longrightarrow> find_indices x Vs = [i]"
proof (induction Vs rule: rev_induct)
case (snoc v Vs)
have dist: "distinct Vs" and xVs: "v \<notin> set Vs" using snoc.prems(3) by(simp_all)
show ?case
proof (cases)
assume i: "i = length Vs"
hence "x = v" using snoc.prems(2) by auto
thus ?case using xVs find_indices_empty i
by fastforce
next
assume "i \<noteq> length Vs"
hence i: "i < length Vs" using snoc.prems(1) by simp
hence Vsi: "Vs ! i = x" using snoc.prems(2) append_Cons_nth_left by fastforce
hence "x \<noteq> v" using snoc.prems(3) i by auto
thus ?case using snoc.IH[OF i Vsi dist] by simp
qed
qed auto
lemma finite_convex_hull_iff_convex_hull_list: assumes Vs: "Vs \<subseteq> carrier_vec n"
and id': "Vs = set Vsl'"
shows "convex_hull Vs = convex_hull_list Vsl'"
proof -
have fin: "finite Vs" unfolding id' by auto
from finite_distinct_list fin obtain Vsl
where id: "Vs = set Vsl" and dist: "distinct Vsl" by auto
from Vs id have Vsl: "set Vsl \<subseteq> carrier_vec n" by auto
{
fix c :: "nat \<Rightarrow> 'a"
have "distinct Vsl \<Longrightarrow>(\<Sum>x\<in>set Vsl. sum_list (map c (find_indices x Vsl))) =
sum c {0..<length Vsl}"
proof (induction Vsl rule: rev_induct)
case (snoc v Vsl)
let ?coef = "\<lambda> x. sum_list (map c (find_indices x (Vsl @ [v])))"
let ?coef' = "\<lambda> x. sum_list (map c (find_indices x Vsl))"
have dist: "distinct Vsl" using snoc.prems by simp
have "sum ?coef (set (Vsl @ [v])) = sum_list (map ?coef (Vsl @ [v]))"
by (rule sum.distinct_set_conv_list[OF snoc.prems, of ?coef])
also have "\<dots> = sum_list (map ?coef Vsl) + ?coef v" by simp
also have "sum_list (map ?coef Vsl) = sum ?coef (set Vsl)"
using sum.distinct_set_conv_list[OF dist, of ?coef] by auto
also have "\<dots> = sum ?coef' (set Vsl)"
proof (intro R.finsum_restrict[of ?coef] restrict_ext, standard)
fix x
assume "x \<in> set Vsl"
then obtain i where i: "i < length Vsl" and x: "x = Vsl ! i"
using in_set_conv_nth[of x Vsl] by blast
hence "(Vsl @ [v]) ! i = x" by (simp add: append_Cons_nth_left)
hence "?coef x = c i"
using distinct_list_find_indices[OF _ _ snoc.prems] i by fastforce
also have "c i = ?coef' x"
using distinct_list_find_indices[OF i _ dist] x by simp
finally show "?coef x = ?coef' x" by auto
qed
also have "\<dots> = sum c {0..<length Vsl}" by (rule snoc.IH[OF dist])
also have "?coef v = c (length Vsl)"
using distinct_list_find_indices[OF _ _ snoc.prems, of "length Vsl" v]
nth_append_length by simp
finally show ?case using sum.atLeast0_lessThan_Suc by simp
qed simp
} note sum_sumlist = this
{
fix b
assume "b \<in> convex_hull_list Vsl"
then obtain c where b: "lincomb_list c Vsl = b" and c: "(\<forall> i < length Vsl. c i \<ge> 0)"
and c1: "sum c {0..<length Vsl} = 1"
unfolding convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def
by auto
have "convex_lincomb (mk_coeff Vsl c) Vs b"
unfolding b[symmetric] convex_lincomb_def nonneg_lincomb_def
apply (subst lincomb_list_as_lincomb[OF Vsl])
by (insert c c1, auto simp: id mk_coeff_def dist sum_sumlist intro!: sum_list_nonneg)
hence "b \<in> convex_hull Vs"
unfolding convex_hull_def convex_lincomb_def using fin by blast
}
moreover
{
fix b
assume "b \<in> convex_hull Vs"
then obtain c Ws where Ws: "Ws \<subseteq> Vs" and b: "lincomb c Ws = b"
and c: "c ` Ws \<subseteq> {x. x \<ge> 0}" and c1: "sum c Ws = 1"
unfolding convex_hull_def convex_lincomb_def nonneg_lincomb_def by auto
let ?d = "\<lambda> x. if x \<in> Ws then c x else 0"
have "lincomb ?d Vs = lincomb c Ws + lincomb (\<lambda> x. 0) (Vs - Ws)"
using lincomb_union2[OF _ _ Diff_disjoint[of Ws Vs], of c "\<lambda> x. 0"]
fin Vs Diff_partition[OF Ws] by metis
also have "lincomb (\<lambda> x. 0) (Vs - Ws) = 0\<^sub>v n"
using lincomb_zero[of "Vs - Ws" "\<lambda> x. 0"] Vs by auto
finally have "lincomb ?d Vs = b" using b lincomb_closed Vs Ws by auto
moreover have "?d ` Vs \<subseteq> {t. t \<ge> 0}" using c by auto
moreover have "sum ?d Vs = 1" using c1 R.extend_sum[OF fin Ws] by auto
ultimately have "\<exists> c. convex_lincomb c Vs b"
unfolding convex_lincomb_def nonneg_lincomb_def by blast
}
moreover
{
fix b
assume "\<exists> c. convex_lincomb c Vs b"
then obtain c where b: "lincomb c Vs = b" and c: "c ` Vs \<subseteq> {x. x \<ge> 0}"
and c1: "sum c Vs = 1"
unfolding convex_lincomb_def nonneg_lincomb_def by auto
from lincomb_as_lincomb_list_distinct[OF Vsl dist, of c]
have b: "lincomb_list (\<lambda>i. c (Vsl ! i)) Vsl = b"
unfolding b[symmetric] id by simp
have "1 = sum c (set Vsl)" using c1 id by auto
also have "\<dots> = sum_list (map c Vsl)" by(rule sum.distinct_set_conv_list[OF dist])
also have "\<dots> = sum ((!) (map c Vsl)) {0..<length Vsl}"
using sum_list_sum_nth length_map by metis
also have "\<dots> = sum (\<lambda> i. c (Vsl ! i)) {0..<length Vsl}" by simp
finally have sum_1: "(\<Sum>i = 0..<length Vsl. c (Vsl ! i)) = 1" by simp
have "\<exists> c. convex_lincomb_list c Vsl b"
unfolding convex_lincomb_list_def nonneg_lincomb_list_def
by (intro exI[of _ "\<lambda>i. c (Vsl ! i)"] conjI b sum_1)
(insert c, force simp: set_conv_nth id)
hence "b \<in> convex_hull_list Vsl" unfolding convex_hull_list_def by auto
}
ultimately have "convex_hull Vs = convex_hull_list Vsl" by auto
also have "convex_hull_list Vsl = convex_hull_list Vsl'"
using convex_hull_list_eq_set[OF Vsl, of Vsl'] id id' by simp
finally show ?thesis by simp
qed
definition "convex S = (convex_hull S = S)"
lemma convex_convex_hull: "convex S \<Longrightarrow> convex_hull S = S"
unfolding convex_def by auto
lemma convex_hull_convex_hull_listD: assumes A: "A \<subseteq> carrier_vec n"
and x: "x \<in> convex_hull A"
shows "\<exists> as. set as \<subseteq> A \<and> x \<in> convex_hull_list as"
proof -
from x[unfolded convex_hull_def]
obtain X c where finX: "finite X" and XA: "X \<subseteq> A" and "convex_lincomb c X x" by auto
hence x: "x \<in> convex_hull X" unfolding convex_hull_def by auto
from finite_list[OF finX] obtain xs where X: "X = set xs" by auto
from finite_convex_hull_iff_convex_hull_list[OF _ this] x XA A have x: "x \<in> convex_hull_list xs" by auto
thus ?thesis using XA unfolding X by auto
qed
lemma convex_hull_convex_sum: assumes A: "A \<subseteq> carrier_vec n"
and x: "x \<in> convex_hull A"
and y: "y \<in> convex_hull A"
and a: "0 \<le> a" "a \<le> 1"
shows "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> convex_hull A"
proof -
from convex_hull_convex_hull_listD[OF A x] obtain xs where xs: "set xs \<subseteq> A"
and x: "x \<in> convex_hull_list xs" by auto
from convex_hull_convex_hull_listD[OF A y] obtain ys where ys: "set ys \<subseteq> A"
and y: "y \<in> convex_hull_list ys" by auto
have fin: "finite (set (xs @ ys))" by auto
have sub: "set (xs @ ys) \<subseteq> A" using xs ys by auto
from convex_hull_list_mono[of "xs @ ys" xs] x sub A have x: "x \<in> convex_hull_list (xs @ ys)" by auto
from convex_hull_list_mono[of "xs @ ys" ys] y sub A have y: "y \<in> convex_hull_list (xs @ ys)" by auto
from convex_hull_list_combination[OF _ x y a]
have "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> convex_hull_list (xs @ ys)" using sub A by auto
from finite_convex_hull_iff_convex_hull_list[of _ "xs @ ys"] this sub A
have "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> convex_hull (set (xs @ ys))" by auto
with convex_hull_mono[OF sub]
show "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> convex_hull A" by auto
qed
lemma convexI: assumes S: "S \<subseteq> carrier_vec n"
and step: "\<And> a x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> 0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> S"
shows "convex S"
unfolding convex_def
proof (standard, standard)
fix z
assume "z \<in> convex_hull S"
from this[unfolded convex_hull_def] obtain W c where "finite W" and WS: "W \<subseteq> S"
and "convex_lincomb c W z" by auto
then show "z \<in> S"
proof (induct W arbitrary: c z)
case empty
thus ?case unfolding convex_lincomb_def by auto
next
case (insert w W c z)
have "convex_lincomb c (insert w W) z" by fact
hence zl: "z = lincomb c (insert w W)" and nonneg: "\<And> w. w \<in> W \<Longrightarrow> 0 \<le> c w"
and cw: "c w \<ge> 0"
and sum: "sum c (insert w W) = 1"
unfolding convex_lincomb_def nonneg_lincomb_def by auto
have zl: "z = c w \<cdot>\<^sub>v w + lincomb c W" unfolding zl
by (rule lincomb_insert2, insert insert S, auto)
have sum: "c w + sum c W = 1" unfolding sum[symmetric]
by (subst sum.insert, insert insert, auto)
have W: "W \<subseteq> carrier_vec n" and w: "w \<in> carrier_vec n" using S insert by auto
show ?case
proof (cases "sum c W = 0")
case True
with nonneg have c0: "\<And> w. w \<in> W \<Longrightarrow> c w = 0"
using insert(1) sum_nonneg_eq_0_iff by auto
with sum have cw: "c w = 1" by auto
have lin0: "lincomb c W = 0\<^sub>v n"
by (intro lincomb_zero W, insert c0, auto)
have "z = w" unfolding zl cw lin0 using w by simp
with insert(4) show ?thesis by simp
next
case False
have "sum c W \<ge> 0" using nonneg by (metis sum_nonneg)
with False have pos: "sum c W > 0" by auto
define b where "b = (\<lambda> w. inverse (sum c W) * c w)"
have "convex_lincomb b W (lincomb b W)"
unfolding convex_lincomb_def nonneg_lincomb_def b_def
proof (intro conjI refl)
show "(\<lambda>w. inverse (sum c W) * c w) ` W \<subseteq> Collect ((\<le>) 0)" using nonneg pos by auto
show "(\<Sum>w\<in>W. inverse (sum c W) * c w) = 1" unfolding sum_distrib_left[symmetric] using False by auto
qed
from insert(3)[OF _ this] insert
have IH: "lincomb b W \<in> S" by auto
have lin: "lincomb c W = sum c W \<cdot>\<^sub>v lincomb b W"
unfolding b_def
by (subst lincomb_smult[symmetric, OF W], rule lincomb_cong[OF _ W], insert False, auto)
from sum cw pos have sum: "sum c W = 1 - c w" and cw1: "c w \<le> 1" by auto
show ?thesis unfolding zl lin unfolding sum
by (rule step[OF _ IH cw cw1], insert insert, auto)
qed
qed
next
show "S \<subseteq> convex_hull S" using S by (rule set_in_convex_hull)
qed
lemma convex_hulls_are_convex: assumes "A \<subseteq> carrier_vec n"
shows "convex (convex_hull A)"
by (intro convexI convex_hull_convex_sum convex_hull_carrier assms)
lemma convex_hull_sum: assumes A: "A \<subseteq> carrier_vec n" and B: "B \<subseteq> carrier_vec n"
shows "convex_hull (A + B) = convex_hull A + convex_hull B"
proof
note cA = convex_hull_carrier[OF A]
note cB = convex_hull_carrier[OF B]
have "convex (convex_hull A + convex_hull B)"
proof (intro convexI sum_carrier_vec convex_hull_carrier A B)
fix a :: 'a and x1 x2
assume "x1 \<in> convex_hull A + convex_hull B" "x2 \<in> convex_hull A + convex_hull B"
then obtain y1 y2 z1 z2 where
x12: "x1 = y1 + z1" "x2 = y2 + z2" and
y12: "y1 \<in> convex_hull A" "y2 \<in> convex_hull A" and
z12: "z1 \<in> convex_hull B" "z2 \<in> convex_hull B"
unfolding set_plus_def by auto
from y12 z12 cA cB have carr:
"y1 \<in> carrier_vec n" "y2 \<in> carrier_vec n"
"z1 \<in> carrier_vec n" "z2 \<in> carrier_vec n"
by auto
assume a: "0 \<le> a" "a \<le> 1"
have A: "a \<cdot>\<^sub>v y1 + (1 - a) \<cdot>\<^sub>v y2 \<in> convex_hull A" using y12 a A by (metis convex_hull_convex_sum)
have B: "a \<cdot>\<^sub>v z1 + (1 - a) \<cdot>\<^sub>v z2 \<in> convex_hull B" using z12 a B by (metis convex_hull_convex_sum)
have "a \<cdot>\<^sub>v x1 + (1 - a) \<cdot>\<^sub>v x2 = (a \<cdot>\<^sub>v y1 + a \<cdot>\<^sub>v z1) + ((1 - a) \<cdot>\<^sub>v y2 + (1 - a) \<cdot>\<^sub>v z2)" unfolding x12
using carr by (auto simp: smult_add_distrib_vec)
also have "\<dots> = (a \<cdot>\<^sub>v y1 + (1 - a) \<cdot>\<^sub>v y2) + (a \<cdot>\<^sub>v z1 + (1 - a) \<cdot>\<^sub>v z2)" using carr
by (intro eq_vecI, auto)
finally show "a \<cdot>\<^sub>v x1 + (1 - a) \<cdot>\<^sub>v x2 \<in> convex_hull A + convex_hull B"
using A B by auto
qed
from convex_convex_hull[OF this]
have id: "convex_hull (convex_hull A + convex_hull B) = convex_hull A + convex_hull B" .
show "convex_hull (A + B) \<subseteq> convex_hull A + convex_hull B"
by (subst id[symmetric], rule convex_hull_mono[OF set_plus_mono2]; intro set_in_convex_hull A B)
show "convex_hull A + convex_hull B \<subseteq> convex_hull (A + B)"
proof
fix x
assume "x \<in> convex_hull A + convex_hull B"
then obtain y z where x: "x = y + z" and y: "y \<in> convex_hull A" and z: "z \<in> convex_hull B"
by (auto simp: set_plus_def)
from convex_hull_convex_hull_listD[OF A y] obtain ys where ysA: "set ys \<subseteq> A" and
y: "y \<in> convex_hull_list ys" by auto
from convex_hull_convex_hull_listD[OF B z] obtain zs where zsB: "set zs \<subseteq> B" and
z: "z \<in> convex_hull_list zs" by auto
from y[unfolded convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def]
obtain c where yid: "y = lincomb_list c ys"
and conv_c: "(\<forall>i<length ys. 0 \<le> c i) \<and> sum c {0..<length ys} = 1"
by auto
from z[unfolded convex_hull_list_def convex_lincomb_list_def nonneg_lincomb_list_def]
obtain d where zid: "z = lincomb_list d zs"
and conv_d: "(\<forall>i<length zs. 0 \<le> d i) \<and> sum d {0..<length zs} = 1"
by auto
from ysA A have ys: "set ys \<subseteq> carrier_vec n" by auto
from zsB B have zs: "set zs \<subseteq> carrier_vec n" by auto
have [intro, simp]: "lincomb_list x ys \<in> carrier_vec n" for x using lincomb_list_carrier[OF ys] .
have [intro, simp]: "lincomb_list x zs \<in> carrier_vec n" for x using lincomb_list_carrier[OF zs] .
have dim[simp]: "dim_vec (lincomb_list d zs) = n" by auto
from yid have y: "y \<in> carrier_vec n" by auto
from zid have z: "z \<in> carrier_vec n" by auto
{
fix x
assume "x \<in> set (map ((+) y) zs)"
then obtain z where "x = y + z" and "z \<in> set zs" by auto
then obtain j where j: "j < length zs" and x: "x = y + zs ! j" unfolding set_conv_nth by auto
hence mem: "zs ! j \<in> set zs" by auto
hence zsj: "zs ! j \<in> carrier_vec n" using zs by auto
let ?list = "(map (\<lambda> y. y + zs ! j) ys)"
let ?set = "set ?list"
have set: "?set \<subseteq> carrier_vec n" using ys A zsj by auto
have lin_map: "lincomb_list c ?list \<in> carrier_vec n"
by (intro lincomb_list_carrier[OF set])
have "y + (zs ! j) = lincomb_list c ?list"
unfolding yid using zsj lin_map lincomb_list_index[OF _ set] lincomb_list_index[OF _ ys]
by (intro eq_vecI, auto simp: field_simps sum_distrib_right[symmetric] conv_c)
hence "convex_lincomb_list c ?list (y + (zs ! j))"
unfolding convex_lincomb_list_def nonneg_lincomb_list_def using conv_c by auto
hence "y + (zs ! j) \<in> convex_hull_list ?list" unfolding convex_hull_list_def by auto
with finite_convex_hull_iff_convex_hull_list[OF set refl]
have "(y + zs ! j) \<in> convex_hull ?set" by auto
also have "\<dots> \<subseteq> convex_hull (A + B)"
by (rule convex_hull_mono, insert mem ys ysA zsB, force simp: set_plus_def)
finally have "x \<in> convex_hull (A + B)" unfolding x .
} note step1 = this
{
let ?list = "map ((+) y) zs"
let ?set = "set ?list"
have set: "?set \<subseteq> carrier_vec n" using zs B y by auto
have lin_map: "lincomb_list d ?list \<in> carrier_vec n"
by (intro lincomb_list_carrier[OF set])
have [simp]: "i < n \<Longrightarrow> (\<Sum>j = 0..<length zs. d j * (y + zs ! j) $ i) =
(\<Sum>j = 0..<length zs. d j * (y $ i + zs ! j $ i))" for i
by (rule sum.cong, insert zs[unfolded set_conv_nth] y, auto)
have "y + z = lincomb_list d ?list"
unfolding zid using y zs lin_map lincomb_list_index[OF _ set] lincomb_list_index[OF _ zs]
set lincomb_list_carrier[OF zs, of d] zs[unfolded set_conv_nth]
by (intro eq_vecI, auto simp: field_simps sum_distrib_right[symmetric] conv_d)
hence "convex_lincomb_list d ?list x" unfolding x
unfolding convex_lincomb_list_def nonneg_lincomb_list_def using conv_d by auto
hence "x \<in> convex_hull_list ?list" unfolding convex_hull_list_def by auto
with finite_convex_hull_iff_convex_hull_list[OF set refl]
have "x \<in> convex_hull ?set" by auto
also have "\<dots> \<subseteq> convex_hull (convex_hull (A + B))"
by (rule convex_hull_mono, insert step1, auto)
also have "\<dots> = convex_hull (A + B)"
by (rule convex_convex_hull[OF convex_hulls_are_convex], intro sum_carrier_vec A B)
finally show "x \<in> convex_hull (A + B)" .
}
qed
qed
lemma convex_hull_in_cone:
"convex_hull C \<subseteq> cone C"
unfolding convex_hull_def cone_def convex_lincomb_def finite_cone_def by auto
lemma convex_cone:
assumes C: "C \<subseteq> carrier_vec n"
shows "convex (cone C)"
unfolding convex_def
using convex_hull_in_cone set_in_convex_hull[OF cone_carrier[OF C]] cone_cone[OF C]
by blast
end
end
\ No newline at end of file
diff --git a/thys/Linear_Inequalities/Decomposition_Theorem.thy b/thys/Linear_Inequalities/Decomposition_Theorem.thy
--- a/thys/Linear_Inequalities/Decomposition_Theorem.thy
+++ b/thys/Linear_Inequalities/Decomposition_Theorem.thy
@@ -1,805 +1,805 @@
section \<open>The Decomposition Theorem\<close>
text \<open>This theory contains a proof of the fact, that every polyhedron can be decomposed
into a convex hull of a finite set of points + a finitely generated cone, including bounds
on the numbers that are required in the decomposition.
We further prove the inverse direction of this theorem (without bounds) and
as a corollary, we derive that a polyhedron is bounded iff it is the convex hull
of finitely many points, i.e., a polytope.\<close>
theory Decomposition_Theorem
imports
Farkas_Minkowsky_Weyl
Convex_Hull
begin
context gram_schmidt
begin
definition "polytope P = (\<exists> V. V \<subseteq> carrier_vec n \<and> finite V \<and> P = convex_hull V)"
definition "polyhedron A b = {x \<in> carrier_vec n. A *\<^sub>v x \<le> b}"
lemma polyhedra_are_convex:
assumes A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr"
and P: "P = polyhedron A b"
shows "convex P"
proof (intro convexI)
show Pcarr: "P \<subseteq> carrier_vec n" using assms unfolding polyhedron_def by auto
fix a :: 'a and x y
assume xy: "x \<in> P" "y \<in> P" and a: "0 \<le> a" "a \<le> 1"
from xy[unfolded P polyhedron_def]
have x: "x \<in> carrier_vec n" and y: "y \<in> carrier_vec n" and le: "A *\<^sub>v x \<le> b" "A *\<^sub>v y \<le> b" by auto
show "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> P" unfolding P polyhedron_def
proof (intro CollectI conjI)
from x have ax: "a \<cdot>\<^sub>v x \<in> carrier_vec n" by auto
from y have ay: "(1 - a) \<cdot>\<^sub>v y \<in> carrier_vec n" by auto
show "a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y \<in> carrier_vec n" using ax ay by auto
show "A *\<^sub>v (a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y) \<le> b"
proof (intro lesseq_vecI[OF _ b])
show "A *\<^sub>v (a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y) \<in> carrier_vec nr" using A x y by auto
fix i
assume i: "i < nr"
from lesseq_vecD[OF b le(1) i] lesseq_vecD[OF b le(2) i]
have le: "(A *\<^sub>v x) $ i \<le> b $ i" "(A *\<^sub>v y) $ i \<le> b $ i" by auto
have "(A *\<^sub>v (a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y)) $ i = a * (A *\<^sub>v x) $ i + (1 - a) * (A *\<^sub>v y) $ i"
using A x y i by (auto simp: scalar_prod_add_distrib[of _ n])
also have "\<dots> \<le> a * b $ i + (1 - a) * b $ i"
by (rule add_mono; rule mult_left_mono, insert le a, auto)
also have "\<dots> = b $ i" by (auto simp: field_simps)
finally show "(A *\<^sub>v (a \<cdot>\<^sub>v x + (1 - a) \<cdot>\<^sub>v y)) $ i \<le> b $ i" .
qed
qed
qed
end
locale gram_schmidt_m = n: gram_schmidt n f_ty + m: gram_schmidt m f_ty
for n m :: nat and f_ty
begin
lemma vec_first_lincomb_list:
assumes Xs: "set Xs \<subseteq> carrier_vec n"
and nm: "m \<le> n"
shows "vec_first (n.lincomb_list c Xs) m =
m.lincomb_list c (map (\<lambda> v. vec_first v m) Xs)"
using Xs
proof (induction Xs arbitrary: c)
case Nil
show ?case by (simp add: nm)
next
case (Cons x Xs)
from Cons.prems have x: "x \<in> carrier_vec n" and Xs: "set Xs \<subseteq> carrier_vec n" by auto
have "vec_first (n.lincomb_list c (x # Xs)) m =
vec_first (c 0 \<cdot>\<^sub>v x + n.lincomb_list (c \<circ> Suc) Xs) m" by auto
also have "\<dots> = vec_first (c 0 \<cdot>\<^sub>v x) m + vec_first (n.lincomb_list (c \<circ> Suc) Xs) m"
using vec_first_add[of m "c 0 \<cdot>\<^sub>v x"] x n.lincomb_list_carrier[OF Xs, of "c \<circ> Suc"] nm
by simp
also have "vec_first (c 0 \<cdot>\<^sub>v x) m = c 0 \<cdot>\<^sub>v vec_first x m"
using vec_first_smult[OF nm, of x "c 0"] Cons.prems by auto
also have "vec_first (n.lincomb_list (c \<circ> Suc) Xs) m =
m.lincomb_list (c \<circ> Suc) (map (\<lambda> v. vec_first v m) Xs)"
using Cons by simp
also have "c 0 \<cdot>\<^sub>v vec_first x m + \<dots> =
m.lincomb_list c (map (\<lambda> v. vec_first v m) (x # Xs))"
by simp
finally show ?case by auto
qed
lemma convex_hull_next_dim:
assumes "n = m + 1"
and X: "X \<subseteq> carrier_vec n"
and "finite X"
and Xm1: "\<forall> y \<in> X. y $ m = 1"
and y_dim: "y \<in> carrier_vec n"
and y: "y $ m = 1"
shows "(vec_first y m \<in> m.convex_hull {vec_first y m | y. y \<in> X}) = (y \<in> n.cone X)"
proof -
- from `finite X` obtain Xs where Xs: "X = set Xs" using finite_list by auto
+ from \<open>finite X\<close> obtain Xs where Xs: "X = set Xs" using finite_list by auto
let ?Y = "{vec_first y m | y. y \<in> X}"
let ?Ys = "map (\<lambda> y. vec_first y m) Xs"
have Ys: "?Y = set ?Ys" using Xs by auto
define x where "x = vec_first y m"
{
have "y = vec_first y m @\<^sub>v vec_last y 1"
- using `n = m + 1` vec_first_last_append y_dim by auto
+ using \<open>n = m + 1\<close> vec_first_last_append y_dim by auto
also have "vec_last y 1 = vec_of_scal (vec_last y 1 $ 0)"
using vec_of_scal_dim_1[of "vec_last y 1"] by simp
also have "vec_last y 1 $ 0 = y $ m"
- using y_dim `n = m + 1` vec_last_index[of y m 1 0] by auto
+ using y_dim \<open>n = m + 1\<close> vec_last_index[of y m 1 0] by auto
finally have "y = x @\<^sub>v vec_of_scal 1" unfolding x_def using y by simp
} note xy = this
{
assume "y \<in> n.cone X"
then obtain c where x: "n.nonneg_lincomb c X y"
- using n.cone_iff_finite_cone[OF X] `finite X`
+ using n.cone_iff_finite_cone[OF X] \<open>finite X\<close>
unfolding n.finite_cone_def by auto
have "1 = y $ m" by (simp add: y)
also have "y = n.lincomb c X"
using x unfolding n.nonneg_lincomb_def by simp
also have "\<dots> $ m = (\<Sum>x\<in>X. c x * x $ m)"
- using n.lincomb_index[OF _ X] `n = m + 1` by simp
+ using n.lincomb_index[OF _ X] \<open>n = m + 1\<close> by simp
also have "\<dots> = sum c X"
by (rule n.R.finsum_restrict, auto, rule restrict_ext, simp add: Xm1)
finally have "y \<in> n.convex_hull X"
unfolding n.convex_hull_def n.convex_lincomb_def
- using `finite X` x by auto
+ using \<open>finite X\<close> x by auto
}
moreover have "n.convex_hull X \<subseteq> n.cone X"
unfolding n.convex_hull_def n.convex_lincomb_def n.finite_cone_def n.cone_def
- using `finite X` by auto
+ using \<open>finite X\<close> by auto
moreover have "n.convex_hull X = n.convex_hull_list Xs"
by (rule n.finite_convex_hull_iff_convex_hull_list[OF X Xs])
moreover {
assume "y \<in> n.convex_hull_list Xs"
then obtain c where c: "n.lincomb_list c Xs = y"
and c0: "\<forall> i < length Xs. c i \<ge> 0" and c1: "sum c {0..<length Xs} = 1"
unfolding n.convex_hull_list_def n.convex_lincomb_list_def
n.nonneg_lincomb_list_def by fast
have "m.lincomb_list c ?Ys = vec_first y m"
- using c vec_first_lincomb_list[of Xs c] X Xs `n = m + 1` by simp
+ using c vec_first_lincomb_list[of Xs c] X Xs \<open>n = m + 1\<close> by simp
hence "x \<in> m.convex_hull_list ?Ys"
unfolding m.convex_hull_list_def m.convex_lincomb_list_def
m.nonneg_lincomb_list_def
using x_def c0 c1 x_def by auto
} moreover {
assume "x \<in> m.convex_hull_list ?Ys"
then obtain c where x: "m.lincomb_list c ?Ys = x"
and c0: "\<forall> i < length Xs. c i \<ge> 0"
and c1: "sum c {0..<length Xs} = 1"
unfolding m.convex_hull_list_def m.convex_lincomb_list_def
m.nonneg_lincomb_list_def by auto
have "n.lincomb_list c Xs $ m = (\<Sum>j = 0..<length Xs. c j * Xs ! j $ m)"
- using n.lincomb_list_index[of m Xs c] `n = m + 1` Xs X by fastforce
+ using n.lincomb_list_index[of m Xs c] \<open>n = m + 1\<close> Xs X by fastforce
also have "\<dots> = sum c {0..<length Xs}"
apply(rule n.R.finsum_restrict, auto, rule restrict_ext)
by (simp add: Xm1 Xs)
also have "\<dots> = 1" by (rule c1)
finally have "vec_last (n.lincomb_list c Xs) 1 $ 0 = 1"
using vec_of_scal_dim_1 vec_last_index[of "n.lincomb_list c Xs" m 1 0]
- n.lincomb_list_carrier Xs X `n = m + 1` by simp
+ n.lincomb_list_carrier Xs X \<open>n = m + 1\<close> by simp
hence "vec_last (n.lincomb_list c Xs) 1 = vec_of_scal 1"
using vec_of_scal_dim_1 by auto
moreover have "vec_first (n.lincomb_list c Xs) m = x"
- using vec_first_lincomb_list `n = m + 1` Xs X x by auto
+ using vec_first_lincomb_list \<open>n = m + 1\<close> Xs X x by auto
moreover have "n.lincomb_list c Xs =
vec_first (n.lincomb_list c Xs) m @\<^sub>v vec_last (n.lincomb_list c Xs) 1"
- using vec_first_last_append Xs X n.lincomb_list_carrier `n = m + 1` by auto
+ using vec_first_last_append Xs X n.lincomb_list_carrier \<open>n = m + 1\<close> by auto
ultimately have "n.lincomb_list c Xs = y" using xy by simp
hence "y \<in> n.convex_hull_list Xs"
unfolding n.convex_hull_list_def n.convex_lincomb_list_def
n.nonneg_lincomb_list_def using c0 c1 by blast
}
moreover have "m.convex_hull ?Y = m.convex_hull_list ?Ys"
using m.finite_convex_hull_iff_convex_hull_list[OF _ Ys] by fastforce
ultimately show ?thesis unfolding x_def by blast
qed
lemma cone_next_dim:
assumes "n = m + 1"
and X: "X \<subseteq> carrier_vec n"
and "finite X"
and Xm0: "\<forall> y \<in> X. y $ m = 0"
and y_dim: "y \<in> carrier_vec n"
and y: "y $ m = 0"
shows "(vec_first y m \<in> m.cone {vec_first y m | y. y \<in> X}) = (y \<in> n.cone X)"
proof -
- from `finite X` obtain Xs where Xs: "X = set Xs" using finite_list by auto
+ from \<open>finite X\<close> obtain Xs where Xs: "X = set Xs" using finite_list by auto
let ?Y = "{vec_first y m | y. y \<in> X}"
let ?Ys = "map (\<lambda> y. vec_first y m) Xs"
have Ys: "?Y = set ?Ys" using Xs by auto
define x where "x = vec_first y m"
{
have "y = vec_first y m @\<^sub>v vec_last y 1"
- using `n = m + 1` vec_first_last_append y_dim by auto
+ using \<open>n = m + 1\<close> vec_first_last_append y_dim by auto
also have "vec_last y 1 = vec_of_scal (vec_last y 1 $ 0)"
using vec_of_scal_dim_1[of "vec_last y 1"] by simp
also have "vec_last y 1 $ 0 = y $ m"
- using y_dim `n = m + 1` vec_last_index[of y m 1 0] by auto
+ using y_dim \<open>n = m + 1\<close> vec_last_index[of y m 1 0] by auto
finally have "y = x @\<^sub>v vec_of_scal 0" unfolding x_def using y by simp
} note xy = this
have "n.cone X = n.cone_list Xs"
- using n.cone_iff_finite_cone[OF X `finite X`] n.finite_cone_iff_cone_list[OF X Xs]
+ using n.cone_iff_finite_cone[OF X \<open>finite X\<close>] n.finite_cone_iff_cone_list[OF X Xs]
by simp
moreover {
assume "y \<in> n.cone_list Xs"
then obtain c where y: "n.lincomb_list c Xs = y" and c: "\<forall> i < length Xs. c i \<ge> 0"
unfolding n.cone_list_def n.nonneg_lincomb_list_def by blast
from y have "m.lincomb_list c ?Ys = x"
unfolding x_def
- using vec_first_lincomb_list Xs X `n = m + 1` by auto
+ using vec_first_lincomb_list Xs X \<open>n = m + 1\<close> by auto
hence "x \<in> m.cone_list ?Ys" using c
unfolding m.cone_list_def m.nonneg_lincomb_list_def by auto
} moreover {
assume "x \<in> m.cone_list ?Ys"
then obtain c where x: "m.lincomb_list c ?Ys = x" and c: "\<forall> i < length Xs. c i \<ge> 0"
unfolding m.cone_list_def m.nonneg_lincomb_list_def by auto
have "vec_last (n.lincomb_list c Xs) 1 $ 0 = n.lincomb_list c Xs $ m"
- using `n = m + 1` n.lincomb_list_carrier X Xs vec_last_index[of _ m 1 0]
+ using \<open>n = m + 1\<close> n.lincomb_list_carrier X Xs vec_last_index[of _ m 1 0]
by auto
also have "\<dots> = 0"
- using n.lincomb_list_index[of m Xs c] Xs X `n = m + 1` Xm0 by simp
+ using n.lincomb_list_index[of m Xs c] Xs X \<open>n = m + 1\<close> Xm0 by simp
also have "\<dots> = vec_last y 1 $ 0"
- using y y_dim `n = m + 1` vec_last_index[of y m 1 0] by auto
+ using y y_dim \<open>n = m + 1\<close> vec_last_index[of y m 1 0] by auto
finally have "vec_last (n.lincomb_list c Xs) 1 = vec_last y 1" by fastforce
moreover have "vec_first (n.lincomb_list c Xs) m = x"
- using vec_first_lincomb_list[of Xs c] x X Xs `n = m + 1`
+ using vec_first_lincomb_list[of Xs c] x X Xs \<open>n = m + 1\<close>
unfolding x_def by simp
ultimately have "n.lincomb_list c Xs = y" unfolding x_def
- using vec_first_last_append[of _ m 1] `n = m + 1` y_dim
+ using vec_first_last_append[of _ m 1] \<open>n = m + 1\<close> y_dim
n.lincomb_list_carrier[of Xs c] Xs X
by metis
hence "y \<in> n.cone_list Xs"
unfolding n.cone_list_def n.nonneg_lincomb_list_def using c by blast
}
moreover have "m.cone_list ?Ys = m.cone ?Y"
using m.finite_cone_iff_cone_list[OF _ Ys] m.cone_iff_finite_cone[of ?Y]
- `finite X` by force
+ \<open>finite X\<close> by force
ultimately show ?thesis unfolding x_def by blast
qed
end
context gram_schmidt
begin
lemma decomposition_theorem_polyhedra_1:
assumes A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr"
and P: "P = polyhedron A b"
shows "\<exists> Q X. X \<subseteq> carrier_vec n \<and> finite X \<and>
Q \<subseteq> carrier_vec n \<and> finite Q \<and>
P = convex_hull Q + cone X \<and>
(A \<in> \<int>\<^sub>m \<inter> Bounded_mat Bnd \<longrightarrow> b \<in> \<int>\<^sub>v \<inter> Bounded_vec Bnd \<longrightarrow>
X \<subseteq> \<int>\<^sub>v \<inter> Bounded_vec (det_bound n (max 1 Bnd))
\<and> Q \<subseteq> Bounded_vec (det_bound n (max 1 Bnd)))"
proof -
interpret next_dim: gram_schmidt "n + 1" "TYPE ('a)".
interpret gram_schmidt_m "n + 1" n "TYPE('a)".
from P[unfolded polyhedron_def] have "P \<subseteq> carrier_vec n" by auto
have mcb: "mat_of_col (-b) \<in> carrier_mat nr 1" using b by auto
define M where "M = (A @\<^sub>c mat_of_col (-b)) @\<^sub>r (0\<^sub>m 1 n @\<^sub>c -1\<^sub>m 1)"
have M_top: "A @\<^sub>c mat_of_col (- b) \<in> carrier_mat nr (n + 1)"
by (rule carrier_append_cols[OF A mcb])
have M_bottom: "(0\<^sub>m 1 n @\<^sub>c -1\<^sub>m 1) \<in> carrier_mat 1 (n + 1)"
by (rule carrier_append_cols, auto)
have M_dim: "M \<in> carrier_mat (nr + 1) (n + 1)"
unfolding M_def
by (rule carrier_append_rows[OF M_top M_bottom])
{
fix x :: "'a vec" fix t assume x: "x \<in> carrier_vec n"
have "x @\<^sub>v vec_of_scal t \<in> next_dim.polyhedral_cone M =
(A *\<^sub>v x - t \<cdot>\<^sub>v b \<le> 0\<^sub>v nr \<and> t \<ge> 0)"
proof -
let ?y = "x @\<^sub>v vec_of_scal t"
have y: "?y \<in> carrier_vec (n + 1)" using x by(simp del: One_nat_def)
have "?y \<in> next_dim.polyhedral_cone M =
(M *\<^sub>v ?y \<le> 0\<^sub>v (nr + 1))"
unfolding next_dim.polyhedral_cone_def using y M_dim by auto
also have "0\<^sub>v (nr + 1) = 0\<^sub>v nr @\<^sub>v 0\<^sub>v 1" by auto
also have "M *\<^sub>v ?y \<le> 0\<^sub>v nr @\<^sub>v 0\<^sub>v 1 =
((A @\<^sub>c mat_of_col (-b)) *\<^sub>v ?y \<le> 0\<^sub>v nr \<and>
(0\<^sub>m 1 n @\<^sub>c -1\<^sub>m 1) *\<^sub>v ?y \<le> 0\<^sub>v 1)"
unfolding M_def
by (intro append_rows_le[OF M_top M_bottom _ y], auto)
also have "(A @\<^sub>c mat_of_col(-b)) *\<^sub>v ?y =
A *\<^sub>v x + mat_of_col(-b) *\<^sub>v vec_of_scal t"
by (rule mat_mult_append_cols[OF A _ x],
auto simp add: b simp del: One_nat_def)
also have "mat_of_col(-b) *\<^sub>v vec_of_scal t = t \<cdot>\<^sub>v (-b)"
by(rule mult_mat_of_row_vec_of_scal)
also have "A *\<^sub>v x + t \<cdot>\<^sub>v (-b) = A *\<^sub>v x - t \<cdot>\<^sub>v b" by auto
also have "(0\<^sub>m 1 n @\<^sub>c - 1\<^sub>m 1) *\<^sub>v (x @\<^sub>v vec_of_scal t) =
0\<^sub>m 1 n *\<^sub>v x + - 1\<^sub>m 1 *\<^sub>v vec_of_scal t"
by(rule mat_mult_append_cols, auto simp add: x simp del: One_nat_def)
also have "\<dots> = - vec_of_scal t" using x by (auto simp del: One_nat_def)
also have "(\<dots> \<le> 0\<^sub>v 1) = (t \<ge> 0)" unfolding less_eq_vec_def by auto
finally show "(?y \<in> next_dim.polyhedral_cone M) =
(A *\<^sub>v x - t \<cdot>\<^sub>v b \<le> 0\<^sub>v nr \<and> t \<ge> 0)" by auto
qed
} note M_cone_car = this
from next_dim.farkas_minkowsky_weyl_theorem_2[OF M_dim, of "max 1 Bnd"]
obtain X where X: "next_dim.polyhedral_cone M = next_dim.cone X" and
fin_X: "finite X" and X_carrier: "X \<subseteq> carrier_vec (n+1)"
and Bnd: "M \<in> \<int>\<^sub>m \<inter> Bounded_mat (max 1 Bnd) \<Longrightarrow>
X \<subseteq> \<int>\<^sub>v \<inter> Bounded_vec (det_bound n (max 1 Bnd))"
by auto
let ?f = "\<lambda> x. if x $ n = 0 then 1 else 1 / (x $ n)"
define Y where "Y = {?f x \<cdot>\<^sub>v x | x. x \<in> X}"
have "finite Y" unfolding Y_def using fin_X by auto
have Y_carrier: "Y \<subseteq> carrier_vec (n+1)" unfolding Y_def using X_carrier by auto
have "?f ` X \<subseteq> {y. y > 0}"
proof
fix y
assume "y \<in> ?f ` X"
then obtain x where x: "x \<in> X" and y: "y = ?f x" by auto
show "y \<in> {y. y > 0}"
proof cases
assume "x $ n = 0"
thus "y \<in> {y. y > 0}" using y by auto
next
assume P: "x $ n \<noteq> 0"
have "x = vec_first x n @\<^sub>v vec_last x 1"
using x X_carrier vec_first_last_append by auto
also have "vec_last x 1 = vec_of_scal (vec_last x 1 $ 0)" by auto
also have "vec_last x 1 $ 0 = x $ n"
using x X_carrier unfolding vec_last_def by auto
finally have "x = vec_first x n @\<^sub>v vec_of_scal (x $ n)" by auto
moreover have "x \<in> next_dim.polyhedral_cone M"
using x X X_carrier next_dim.set_in_cone by auto
ultimately have "x $ n \<ge> 0" using M_cone_car vec_first_carrier by metis
hence "x $ n > 0" using P by auto
thus "y \<in> {y. y > 0}" using y by auto
qed
qed
hence Y: "next_dim.cone Y = next_dim.polyhedral_cone M" unfolding Y_def
using next_dim.cone_smult_basis[OF X_carrier] X by auto
define Y0 where "Y0 = {v \<in> Y. v $ n = 0}"
define Y1 where "Y1 = Y - Y0"
have Y0_carrier: "Y0 \<subseteq> carrier_vec (n + 1)" and Y1_carrier: "Y1 \<subseteq> carrier_vec (n + 1)"
unfolding Y0_def Y1_def using Y_carrier by auto
have "finite Y0" and "finite Y1"
- unfolding Y0_def Y1_def using `finite Y` by auto
+ unfolding Y0_def Y1_def using \<open>finite Y\<close> by auto
have Y1: "\<And> y. y \<in> Y1 \<Longrightarrow> y $ n = 1"
proof -
fix y assume y: "y \<in> Y1"
hence "y \<in> Y" unfolding Y1_def by auto
then obtain x where "x \<in> X" and x: "y = ?f x \<cdot>\<^sub>v x" unfolding Y_def by auto
then have "x $ n \<noteq> 0" using x y Y1_def Y0_def by auto
then have "y = 1 / (x $ n) \<cdot>\<^sub>v x" using x by auto
- then have "y $ n = 1 / (x $ n) * x $ n" using X_carrier `x \<in> X` by auto
- thus "y $ n = 1" using `x $ n \<noteq> 0` by auto
+ then have "y $ n = 1 / (x $ n) * x $ n" using X_carrier \<open>x \<in> X\<close> by auto
+ thus "y $ n = 1" using \<open>x $ n \<noteq> 0\<close> by auto
qed
let ?Z0 = "{vec_first y n | y. y \<in> Y0}"
let ?Z1 = "{vec_first y n | y. y \<in> Y1}"
show ?thesis
proof (intro exI conjI impI)
show "?Z0 \<subseteq> carrier_vec n" by auto
show "?Z1 \<subseteq> carrier_vec n" by auto
- show "finite ?Z0" using `finite Y0` by auto
- show "finite ?Z1" using `finite Y1` by auto
+ show "finite ?Z0" using \<open>finite Y0\<close> by auto
+ show "finite ?Z1" using \<open>finite Y1\<close> by auto
show "P = convex_hull ?Z1 + cone ?Z0"
proof -
{
fix x
assume "x \<in> P"
hence xn: "x \<in> carrier_vec n" and "A *\<^sub>v x \<le> b"
using P unfolding polyhedron_def by auto
hence "A *\<^sub>v x - 1 \<cdot>\<^sub>v b \<le> 0\<^sub>v nr"
using vec_le_iff_diff_le_0 A b carrier_vecD mult_mat_vec_carrier one_smult_vec
by metis
hence "x @\<^sub>v vec_of_scal 1 \<in> next_dim.polyhedral_cone M"
using M_cone_car[OF xn] by auto
hence "x @\<^sub>v vec_of_scal 1 \<in> next_dim.cone Y" using Y by auto
hence "x @\<^sub>v vec_of_scal 1 \<in> next_dim.finite_cone Y"
- using next_dim.cone_iff_finite_cone[OF Y_carrier `finite Y`] by auto
+ using next_dim.cone_iff_finite_cone[OF Y_carrier \<open>finite Y\<close>] by auto
then obtain c where c: "next_dim.nonneg_lincomb c Y (x @\<^sub>v vec_of_scal 1)"
- unfolding next_dim.finite_cone_def using `finite Y` by auto
+ unfolding next_dim.finite_cone_def using \<open>finite Y\<close> by auto
let ?y = "next_dim.lincomb c Y1"
let ?z = "next_dim.lincomb c Y0"
have y_dim: "?y \<in> carrier_vec (n + 1)" and z_dim: "?z \<in> carrier_vec (n + 1)"
unfolding next_dim.nonneg_lincomb_def
using Y0_carrier Y1_carrier next_dim.lincomb_closed by simp_all
hence yz_dim: "?y + ?z \<in> carrier_vec (n + 1)" by auto
have "x @\<^sub>v vec_of_scal 1 = next_dim.lincomb c Y"
using c unfolding next_dim.nonneg_lincomb_def by auto
also have "Y = Y1 \<union> Y0" unfolding Y1_def using Y0_def by blast
also have "next_dim.lincomb c (Y1 \<union> Y0) = ?y + ?z"
using next_dim.lincomb_union2[of Y1 Y0]
- `finite Y0` `finite Y` Y0_carrier Y_carrier
+ \<open>finite Y0\<close> \<open>finite Y\<close> Y0_carrier Y_carrier
unfolding Y1_def by fastforce
also have "?y + ?z = vec_first (?y + ?z) n @\<^sub>v vec_last (?y + ?z) 1"
using vec_first_last_append[of "?y + ?z" n 1] add_carrier_vec yz_dim
by simp
also have "vec_last (?y + ?z) 1 = vec_of_scal ((?y + ?z) $ n)"
using vec_of_scal_dim_1 vec_last_index[OF yz_dim, of 0] by auto
finally have "x @\<^sub>v vec_of_scal 1 =
vec_first (?y + ?z) n @\<^sub>v vec_of_scal ((?y + ?z) $ n)" by auto
hence "x = vec_first (?y + ?z) n" and
yz_last: "vec_of_scal 1 = vec_of_scal ((?y + ?z) $ n)"
using append_vec_eq yz_dim xn by auto
hence xyz: "x = vec_first ?y n + vec_first ?z n"
using vec_first_add[of n ?y ?z] y_dim z_dim by simp
have "1 = ((?y + ?z) $ n)" using yz_last index_vec_of_scal
by (metis (no_types, lifting))
hence "1 = ?y $ n + ?z $ n" using y_dim z_dim by auto
moreover have zn0: "?z $ n = 0"
using next_dim.lincomb_index[OF _ Y0_carrier] Y0_def by auto
ultimately have yn1: "1 = ?y $ n" by auto
have "next_dim.nonneg_lincomb c Y1 ?y"
using c Y1_def
unfolding next_dim.nonneg_lincomb_def by auto
hence "?y \<in> next_dim.cone Y1"
- using next_dim.cone_iff_finite_cone[OF Y1_carrier] `finite Y1`
+ using next_dim.cone_iff_finite_cone[OF Y1_carrier] \<open>finite Y1\<close>
unfolding next_dim.finite_cone_def by auto
hence y: "vec_first ?y n \<in> convex_hull ?Z1"
- using convex_hull_next_dim[OF _ Y1_carrier `finite Y1` _ y_dim] Y1 yn1
+ using convex_hull_next_dim[OF _ Y1_carrier \<open>finite Y1\<close> _ y_dim] Y1 yn1
by simp
have "next_dim.nonneg_lincomb c Y0 ?z" using c Y0_def
unfolding next_dim.nonneg_lincomb_def by blast
hence "?z \<in> next_dim.cone Y0"
- using `finite Y0` next_dim.cone_iff_finite_cone[OF Y0_carrier `finite Y0`]
+ using \<open>finite Y0\<close> next_dim.cone_iff_finite_cone[OF Y0_carrier \<open>finite Y0\<close>]
unfolding next_dim.finite_cone_def
by fastforce
hence z: "vec_first ?z n \<in> cone ?Z0"
- using cone_next_dim[OF _ Y0_carrier `finite Y0` _ _ zn0] Y0_def
+ using cone_next_dim[OF _ Y0_carrier \<open>finite Y0\<close> _ _ zn0] Y0_def
next_dim.lincomb_closed[OF Y0_carrier] by blast
from xyz y z have "x \<in> convex_hull ?Z1 + cone ?Z0" by blast
} moreover {
fix x
assume "x \<in> convex_hull ?Z1 + cone ?Z0"
then obtain y z where "x = y + z" and y: "y \<in> convex_hull ?Z1"
and z: "z \<in> cone ?Z0" by (auto elim: set_plus_elim)
have yn: "y \<in> carrier_vec n"
- using y convex_hull_carrier[OF `?Z1 \<subseteq> carrier_vec n`] by blast
+ using y convex_hull_carrier[OF \<open>?Z1 \<subseteq> carrier_vec n\<close>] by blast
hence "y @\<^sub>v vec_of_scal 1 \<in> carrier_vec (n + 1)"
using vec_of_scal_dim(2) by fast
moreover have "vec_first (y @\<^sub>v vec_of_scal 1) n \<in> convex_hull ?Z1"
using vec_first_append[OF yn] y by auto
moreover have "(y @\<^sub>v vec_of_scal 1) $ n = 1" using yn by simp
ultimately have "y @\<^sub>v vec_of_scal 1 \<in> next_dim.cone Y1"
- using convex_hull_next_dim[OF _ Y1_carrier `finite Y1`] Y1 by blast
+ using convex_hull_next_dim[OF _ Y1_carrier \<open>finite Y1\<close>] Y1 by blast
hence y_cone: "y @\<^sub>v vec_of_scal 1 \<in> next_dim.cone Y"
using next_dim.cone_mono[of Y1 Y] Y1_def by blast
have zn: "z \<in> carrier_vec n" using z cone_carrier[of ?Z0] by fastforce
hence "z @\<^sub>v vec_of_scal 0 \<in> carrier_vec (n + 1)"
using vec_of_scal_dim(2) by fast
moreover have "vec_first (z @\<^sub>v vec_of_scal 0) n \<in> cone ?Z0"
using vec_first_append[OF zn] z by auto
moreover have "(z @\<^sub>v vec_of_scal 0) $ n = 0" using zn by simp
ultimately have "z @\<^sub>v vec_of_scal 0 \<in> next_dim.cone Y0"
- using cone_next_dim[OF _ Y0_carrier `finite Y0`] Y0_def by blast
+ using cone_next_dim[OF _ Y0_carrier \<open>finite Y0\<close>] Y0_def by blast
hence z_cone: "z @\<^sub>v vec_of_scal 0 \<in> next_dim.cone Y"
using Y0_def next_dim.cone_mono[of Y0 Y] by blast
- have xn: "x \<in> carrier_vec n" using `x = y + z` yn zn by blast
+ have xn: "x \<in> carrier_vec n" using \<open>x = y + z\<close> yn zn by blast
have "x @\<^sub>v vec_of_scal 1 = (y @\<^sub>v vec_of_scal 1) + (z @\<^sub>v vec_of_scal 0)"
- using `x = y + z` append_vec_add[OF yn zn]
+ using \<open>x = y + z\<close> append_vec_add[OF yn zn]
unfolding vec_of_scal_def by auto
hence "x @\<^sub>v vec_of_scal 1 \<in> next_dim.cone Y"
using next_dim.cone_elem_sum[OF Y_carrier y_cone z_cone] by simp
hence "A *\<^sub>v x - b \<le> 0\<^sub>v nr" using M_cone_car[OF xn] Y by simp
hence "A *\<^sub>v x \<le> b" using vec_le_iff_diff_le_0[of "A *\<^sub>v x" b]
dim_mult_mat_vec[of A x] A by simp
hence "x \<in> P" using P xn unfolding polyhedron_def by blast
}
ultimately show "P = convex_hull ?Z1 + cone ?Z0" by blast
qed
let ?Bnd = "det_bound n (max 1 Bnd)"
assume "A \<in> \<int>\<^sub>m \<inter> Bounded_mat Bnd"
"b \<in> \<int>\<^sub>v \<inter> Bounded_vec Bnd"
hence *: "A \<in> \<int>\<^sub>m" "A \<in> Bounded_mat Bnd" "b \<in> \<int>\<^sub>v" "b \<in> Bounded_vec Bnd" by auto
have "elements_mat M \<subseteq> elements_mat A \<union> vec_set (-b) \<union> {0,-1}"
unfolding M_def
unfolding elements_mat_append_rows[OF M_top M_bottom]
unfolding elements_mat_append_cols[OF A mcb]
by (subst elements_mat_append_cols, auto)
also have "\<dots> \<subseteq> \<int> \<inter> ({x. abs x \<le> Bnd} \<union> {0,-1})"
using *[unfolded Bounded_mat_elements_mat Ints_mat_elements_mat
Bounded_vec_vec_set Ints_vec_vec_set] by auto
also have "\<dots> \<subseteq> \<int> \<inter> ({x. abs x \<le> max 1 Bnd})" by auto
finally have "M \<in> \<int>\<^sub>m" "M \<in> Bounded_mat (max 1 Bnd)"
unfolding Bounded_mat_elements_mat Ints_mat_elements_mat by auto
hence "M \<in> \<int>\<^sub>m \<inter> Bounded_mat (max 1 Bnd)" by blast
from Bnd[OF this]
have XBnd: "X \<subseteq> \<int>\<^sub>v \<inter> Bounded_vec ?Bnd" .
{
fix y
assume y: "y \<in> Y"
then obtain x where y: "y = ?f x \<cdot>\<^sub>v x" and xX: "x \<in> X" unfolding Y_def by auto
with \<open>X \<subseteq> carrier_vec (n+1)\<close> have x: "x \<in> carrier_vec (n+1)" by auto
from XBnd xX have xI: "x \<in> \<int>\<^sub>v" and xB: "x \<in> Bounded_vec ?Bnd" by auto
{
assume "y $ n = 0"
hence "y = x" unfolding y using x by auto
hence "y \<in> \<int>\<^sub>v \<inter> Bounded_vec ?Bnd" using xI xB by auto
} note y0 = this
{
assume "y $ n \<noteq> 0"
hence x0: "x $ n \<noteq> 0" using x unfolding y by auto
from x xI have "x $ n \<in> \<int>" unfolding Ints_vec_def by auto
with x0 have "abs (x $ n) \<ge> 1" by (meson Ints_nonzero_abs_ge1)
hence abs: "abs (1 / (x $ n)) \<le> 1" by simp
{
fix a
have "abs ((1 / (x $ n)) * a) = abs (1 / (x $ n)) * abs a"
by simp
also have "\<dots> \<le> 1 * abs a"
by (rule mult_right_mono[OF abs], auto)
finally have "abs ((1 / (x $ n)) * a) \<le> abs a" by auto
} note abs = this
from x0 have y: "y = (1 / (x $ n)) \<cdot>\<^sub>v x" unfolding y by auto
have vy: "vec_set y = (\<lambda> a. (1 / (x $ n)) * a) ` vec_set x"
unfolding y by (auto simp: vec_set_def)
have "y \<in> Bounded_vec ?Bnd" using xB abs
unfolding Bounded_vec_vec_set vy
by (smt imageE max.absorb2 max.bounded_iff)
} note yn0 = this
note y0 yn0
} note BndY = this
from \<open>Y \<subseteq> carrier_vec (n+1)\<close>
have setvY: "y \<in> Y \<Longrightarrow> set\<^sub>v (vec_first y n) \<subseteq> set\<^sub>v y" for y
unfolding vec_first_def vec_set_def by auto
from BndY(1) setvY
show "?Z0 \<subseteq> \<int>\<^sub>v \<inter> Bounded_vec (det_bound n (max 1 Bnd))"
by (force simp: Bounded_vec_vec_set Ints_vec_vec_set Y0_def)
from BndY(2) setvY
show "?Z1 \<subseteq> Bounded_vec (det_bound n (max 1 Bnd))"
by (force simp: Bounded_vec_vec_set Ints_vec_vec_set Y0_def Y1_def)
qed
qed
lemma decomposition_theorem_polyhedra_2:
assumes Q: "Q \<subseteq> carrier_vec n" and fin_Q: "finite Q"
and X: "X \<subseteq> carrier_vec n" and fin_X: "finite X"
and P: "P = convex_hull Q + cone X"
shows "\<exists>A b nr. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b"
proof -
interpret next_dim: gram_schmidt "n + 1" "TYPE ('a)".
interpret gram_schmidt_m "n + 1" n "TYPE('a)".
from fin_Q obtain Qs where Qs: "Q = set Qs" using finite_list by auto
from fin_X obtain Xs where Xs: "X = set Xs" using finite_list by auto
define Y where "Y = {x @\<^sub>v vec_of_scal 1 | x. x \<in> Q}"
define Z where "Z = {x @\<^sub>v vec_of_scal 0 | x. x \<in> X}"
have fin_Y: "finite Y" unfolding Y_def using fin_Q by simp
have fin_Z: "finite Z" unfolding Z_def using fin_X by simp
have Y_dim: "Y \<subseteq> carrier_vec (n + 1)"
unfolding Y_def using Q append_carrier_vec[OF _ vec_of_scal_dim(2)[of 1]]
by blast
have Z_dim: "Z \<subseteq> carrier_vec (n + 1)"
unfolding Z_def using X append_carrier_vec[OF _ vec_of_scal_dim(2)[of 0]]
by blast
have Y_car: "Q = {vec_first x n | x. x \<in> Y}"
proof (intro equalityI subsetI)
fix x assume x: "x \<in> Q"
hence "x @\<^sub>v vec_of_scal 1 \<in> Y" unfolding Y_def by blast
thus "x \<in> {vec_first x n | x. x \<in> Y}"
using Q vec_first_append[of x n "vec_of_scal 1"] x by force
next
fix x assume "x \<in> {vec_first x n | x. x \<in> Y}"
then obtain y where "y \<in> Q" and "x = vec_first (y @\<^sub>v vec_of_scal 1) n"
unfolding Y_def by blast
thus "x \<in> Q" using Q vec_first_append[of y] by auto
qed
have Z_car: "X = {vec_first x n | x. x \<in> Z}"
proof (intro equalityI subsetI)
fix x assume x: "x \<in> X"
hence "x @\<^sub>v vec_of_scal 0 \<in> Z" unfolding Z_def by blast
thus "x \<in> {vec_first x n | x. x \<in> Z}"
using X vec_first_append[of x n "vec_of_scal 0"] x by force
next
fix x assume "x \<in> {vec_first x n | x. x \<in> Z}"
then obtain y where "y \<in> X" and "x = vec_first (y @\<^sub>v vec_of_scal 0) n"
unfolding Z_def by blast
thus "x \<in> X" using X vec_first_append[of y] by auto
qed
have Y_last: "\<forall> x \<in> Y. x $ n = 1" unfolding Y_def using Q by auto
have Z_last: "\<forall> x \<in> Z. x $ n = 0" unfolding Z_def using X by auto
have "finite (Y \<union> Z)" using fin_Y fin_Z by blast
moreover have "Y \<union> Z \<subseteq> carrier_vec (n + 1)" using Y_dim Z_dim by blast
ultimately obtain B nr
where B: "next_dim.cone (Y \<union> Z) = next_dim.polyhedral_cone B"
and B_carrier: "B \<in> carrier_mat nr (n + 1)"
using next_dim.farkas_minkowsky_weyl_theorem[of "next_dim.cone (Y \<union> Z)"]
by blast
define A where "A = mat_col_first B n"
define b where "b = col B n"
have B_blocks: "B = A @\<^sub>c mat_of_col b"
unfolding A_def b_def
using mat_col_first_last_append[of B n 1] B_carrier
mat_of_col_dim_col_1[of "mat_col_last B 1"] by auto
have A_carrier: "A \<in> carrier_mat nr n" unfolding A_def using B_carrier by force
have b_carrier: "b \<in> carrier_vec nr" unfolding b_def using B_carrier by force
{
fix x assume "x \<in> P"
then obtain y z where x: "x = y + z" and y: "y \<in> convex_hull Q" and z: "z \<in> cone X"
using P by (auto elim: set_plus_elim)
have yn: "y \<in> carrier_vec n" using y convex_hull_carrier[OF Q] by blast
moreover have zn: "z \<in> carrier_vec n" using z cone_carrier[OF X] by blast
ultimately have xn: "x \<in> carrier_vec n" using x by blast
have yn1: "y @\<^sub>v vec_of_scal 1 \<in> carrier_vec (n + 1)"
using append_carrier_vec[OF yn] vec_of_scal_dim by fast
have y_last: "(y @\<^sub>v vec_of_scal 1) $ n = 1" using yn by force
have "vec_first (y @\<^sub>v vec_of_scal 1) n = y"
using vec_first_append[OF yn] by simp
hence "y @\<^sub>v vec_of_scal 1 \<in> next_dim.cone Y"
using convex_hull_next_dim[OF _ Y_dim fin_Y Y_last yn1 y_last] Y_car y by argo
hence y_cone: "y @\<^sub>v vec_of_scal 1 \<in> next_dim.cone (Y \<union> Z)"
using next_dim.cone_mono[of Y "Y \<union> Z"] by blast
have zn1: "z @\<^sub>v vec_of_scal 0 \<in> carrier_vec (n + 1)"
using append_carrier_vec[OF zn] vec_of_scal_dim by fast
have z_last: "(z @\<^sub>v vec_of_scal 0) $ n = 0" using zn by force
have "vec_first (z @\<^sub>v vec_of_scal 0) n = z"
using vec_first_append[OF zn] by simp
hence "z @\<^sub>v vec_of_scal 0 \<in> next_dim.cone Z"
using cone_next_dim[OF _ Z_dim fin_Z Z_last zn1 z_last] Z_car z by argo
hence z_cone: "z @\<^sub>v vec_of_scal 0 \<in> next_dim.cone (Y \<union> Z)"
using next_dim.cone_mono[of Z "Y \<union> Z"] by blast
- from `x = y + z`
+ from \<open>x = y + z\<close>
have "x @\<^sub>v vec_of_scal 1 = (y @\<^sub>v vec_of_scal 1) + (z @\<^sub>v vec_of_scal 0)"
using append_vec_add[OF yn zn] vec_of_scal_dim_1
unfolding vec_of_scal_def by auto
hence "x @\<^sub>v vec_of_scal 1 \<in> next_dim.cone (Y \<union> Z) \<and> x \<in> carrier_vec n"
using next_dim.cone_elem_sum[OF _ y_cone z_cone] Y_dim Z_dim xn by auto
} moreover {
fix x assume "x @\<^sub>v vec_of_scal 1 \<in> next_dim.cone (Y \<union> Z)"
then obtain c where x: "next_dim.lincomb c (Y \<union> Z) = x @\<^sub>v vec_of_scal 1"
and c: "c ` (Y \<union> Z) \<subseteq> {t. t \<ge> 0}"
using next_dim.cone_iff_finite_cone Y_dim Z_dim fin_Y fin_Z
unfolding next_dim.finite_cone_def next_dim.nonneg_lincomb_def by auto
let ?y = "next_dim.lincomb c Y"
let ?z = "next_dim.lincomb c Z"
have xyz: "x @\<^sub>v vec_of_scal 1 = ?y + ?z"
using x next_dim.lincomb_union[OF Y_dim Z_dim _ fin_Y fin_Z] Y_last Z_last
by fastforce
have y_dim: "?y \<in> carrier_vec (n + 1)" using next_dim.lincomb_closed[OF Y_dim]
by blast
have z_dim: "?z \<in> carrier_vec (n + 1)" using next_dim.lincomb_closed[OF Z_dim]
by blast
have "x @\<^sub>v vec_of_scal 1 \<in> carrier_vec (n + 1)"
using xyz add_carrier_vec[OF y_dim z_dim] by argo
hence x_dim: "x \<in> carrier_vec n"
using carrier_dim_vec[of x n] carrier_dim_vec[of _ "n + 1"]
by force
have z_last: "?z $ n = 0" using Z_last next_dim.lincomb_index[OF _ Z_dim, of n]
by force
have "?y $ n + ?z $ n = (x @\<^sub>v vec_of_scal 1) $ n"
using xyz index_add_vec(1) z_dim by simp
also have "\<dots> = 1" using x_dim by auto
finally have y_last: "?y $ n = 1" using z_last by algebra
have "?y \<in> next_dim.cone Y"
using next_dim.cone_iff_finite_cone[OF Y_dim] fin_Y c
unfolding next_dim.finite_cone_def next_dim.nonneg_lincomb_def by auto
hence y_cone: "vec_first ?y n \<in> convex_hull Q"
using convex_hull_next_dim[OF _ Y_dim fin_Y Y_last y_dim y_last] Y_car
by blast
have "?z \<in> next_dim.cone Z"
using next_dim.cone_iff_finite_cone[OF Z_dim] fin_Z c
unfolding next_dim.finite_cone_def next_dim.nonneg_lincomb_def by auto
hence z_cone: "vec_first ?z n \<in> cone X"
using cone_next_dim[OF _ Z_dim fin_Z Z_last z_dim z_last] Z_car
by blast
have "x = vec_first (x @\<^sub>v vec_of_scal 1) n" using vec_first_append[OF x_dim] by simp
also have "\<dots> = vec_first ?y n + vec_first ?z n"
using xyz vec_first_add[of n ?y ?z] y_dim z_dim carrier_dim_vec by auto
finally have "x \<in> P"
using y_cone z_cone P by blast
} moreover {
fix x :: "'a vec"
assume xn: "x \<in> carrier_vec n"
hence "(x @\<^sub>v vec_of_scal 1 \<in> next_dim.polyhedral_cone B) =
(B *\<^sub>v (x @\<^sub>v vec_of_scal 1) \<le> 0\<^sub>v nr)"
unfolding next_dim.polyhedral_cone_def using B_carrier
using append_carrier_vec[OF _ vec_of_scal_dim(2)[of 1]] by auto
also have "\<dots> = ((A @\<^sub>c mat_of_col b) *\<^sub>v (x @\<^sub>v vec_of_scal 1) \<le> 0\<^sub>v nr)"
using B_blocks by blast
also have "(A @\<^sub>c mat_of_col b) *\<^sub>v (x @\<^sub>v vec_of_scal 1) =
A *\<^sub>v x + mat_of_col b *\<^sub>v vec_of_scal 1"
by (rule mat_mult_append_cols, insert A_carrier b_carrier xn, auto simp del: One_nat_def)
also have "mat_of_col b *\<^sub>v vec_of_scal 1 = b"
using mult_mat_of_row_vec_of_scal[of b 1] by simp
also have "A *\<^sub>v x + b = A *\<^sub>v x - -b" by auto
finally have "(x @\<^sub>v vec_of_scal 1 \<in> next_dim.polyhedral_cone B) = (A *\<^sub>v x \<le> -b)"
using vec_le_iff_diff_le_0[of "A *\<^sub>v x" "-b"] A_carrier by simp
}
ultimately have "P = polyhedron A (-b)"
unfolding polyhedron_def using B by blast
moreover have "-b \<in> carrier_vec nr" using b_carrier by simp
ultimately show ?thesis using A_carrier by blast
qed
lemma decomposition_theorem_polyhedra:
"(\<exists> A b nr. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b) \<longleftrightarrow>
(\<exists> Q X. Q \<union> X \<subseteq> carrier_vec n \<and> finite (Q \<union> X) \<and> P = convex_hull Q + cone X)" (is "?l = ?r")
proof
assume ?l
then obtain A b nr where A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr" and P: "P = polyhedron A b" by auto
from decomposition_theorem_polyhedra_1[OF this] obtain Q X
where *: "X \<subseteq> carrier_vec n" "finite X" "Q \<subseteq> carrier_vec n" "finite Q" "P = convex_hull Q + cone X"
by meson
show ?r
by (rule exI[of _ Q], rule exI[of _ X], insert *, auto simp: polytope_def)
next
assume ?r
then obtain Q X where QX_carrier: "Q \<union> X \<subseteq> carrier_vec n"
and QX_fin: "finite (Q \<union> X)"
and P: "P = convex_hull Q + cone X" by blast
from QX_carrier have Q: "Q \<subseteq> carrier_vec n" and X: "X \<subseteq> carrier_vec n" by simp_all
from QX_fin have fin_Q: "finite Q" and fin_X: "finite X" by simp_all
show ?l using decomposition_theorem_polyhedra_2[OF Q fin_Q X fin_X P] by blast
qed
lemma polytope_equiv_bounded_polyhedron:
"polytope P \<longleftrightarrow>
(\<exists>A b nr bnd. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b \<and> P \<subseteq> Bounded_vec bnd)"
proof
assume polyP: "polytope P"
from this obtain Q where Qcarr: "Q \<subseteq> carrier_vec n" and finQ: "finite Q"
and PconvhQ: "P = convex_hull Q" unfolding polytope_def by auto
let ?X = "{}"
have "convex_hull Q + {0\<^sub>v n} = convex_hull Q" using Qcarr add_0_right_vecset[of "convex_hull Q"]
by (simp add: convex_hull_carrier)
hence "P = convex_hull Q + cone ?X" using PconvhQ by simp
hence "Q \<union> ?X \<subseteq> carrier_vec n \<and> finite (Q \<union> ?X) \<and> P = convex_hull Q + cone ?X"
using Qcarr finQ PconvhQ by simp
hence "\<exists> A b nr. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b"
using decomposition_theorem_polyhedra by blast
hence Ppolyh: "\<exists>A b nr. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b" by blast
from finite_Bounded_vec_Max[OF Qcarr finQ] obtain bnd where "Q \<subseteq> Bounded_vec bnd" by auto
hence Pbnd: "P \<subseteq> Bounded_vec bnd" using convex_hull_bound PconvhQ Qcarr by auto
from Ppolyh Pbnd show "\<exists>A b nr bnd. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr
\<and> P = polyhedron A b \<and> P \<subseteq> Bounded_vec bnd" by auto
next
assume "\<exists>A b nr bnd. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b
\<and> P \<subseteq> Bounded_vec bnd"
from this obtain A b nr bnd where Adim: "A \<in> carrier_mat nr n" and bdim: "b \<in> carrier_vec nr"
and Ppolyh: "P = polyhedron A b" and Pbnd: "P \<subseteq> Bounded_vec bnd" by auto
have "\<exists> A b nr. A \<in> carrier_mat nr n \<and> b \<in> carrier_vec nr \<and> P = polyhedron A b"
using Adim bdim Ppolyh by blast
hence "\<exists> Q X. Q \<union> X \<subseteq> carrier_vec n \<and> finite (Q \<union> X) \<and> P = convex_hull Q + cone X"
using decomposition_theorem_polyhedra by simp
from this obtain Q X where QXcarr: "Q \<union> X \<subseteq> carrier_vec n"
and finQX: "finite (Q \<union> X)" and Psum: "P = convex_hull Q + cone X" by auto
from QXcarr have Qcarr: "convex_hull Q \<subseteq> carrier_vec n" by (simp add: convex_hull_carrier)
from QXcarr have Xcarr: "cone X \<subseteq> carrier_vec n" by (simp add: gram_schmidt.cone_carrier)
from Pbnd have Pcarr: "P \<subseteq> carrier_vec n" using Ppolyh unfolding polyhedron_def by simp
have "P = convex_hull Q"
proof(cases "Q = {}")
case True
then show "P = convex_hull Q" unfolding Psum by (auto simp: set_plus_def)
next
case False
hence convnotempty: "convex_hull Q \<noteq> {}" using QXcarr by simp
have Pbndex: "\<exists>bnd. P \<subseteq> Bounded_vec bnd" using Pbnd
using QXcarr by auto
from False have "(\<exists> bndc. cone X \<subseteq> Bounded_vec bndc)"
using bounded_vecset_sum[OF Qcarr Xcarr Psum Pbndex] False convnotempty by blast
hence "cone X = {0\<^sub>v n}" using bounded_cone_is_zero QXcarr by auto
thus ?thesis unfolding Psum using Qcarr by (auto simp: add_0_right_vecset)
qed
thus "polytope P" using finQX QXcarr unfolding polytope_def by auto
qed
end
end
\ No newline at end of file
diff --git a/thys/Linear_Inequalities/Integer_Hull.thy b/thys/Linear_Inequalities/Integer_Hull.thy
--- a/thys/Linear_Inequalities/Integer_Hull.thy
+++ b/thys/Linear_Inequalities/Integer_Hull.thy
@@ -1,364 +1,364 @@
section \<open>Integer Hull\<close>
text \<open>We define the integer hull of a polyhedron, i.e., the convex hull of all integer solutions.
Moreover, we prove the result of Meyer that the integer hull of a polyhedron defined by an
integer matrix is again a polyhedron, and give bounds for a corresponding decomposition theorem.\<close>
theory Integer_Hull
imports
Decomposition_Theorem
Mixed_Integer_Solutions (* for gram-schmidt-floor *)
begin
context gram_schmidt
begin
definition "integer_hull P = convex_hull (P \<inter> \<int>\<^sub>v)"
lemma integer_hull_mono: "P \<subseteq> Q \<Longrightarrow> integer_hull P \<subseteq> integer_hull Q"
unfolding integer_hull_def
by (intro convex_hull_mono, auto)
end
lemma abs_neg_floor: "\<bar>of_int b\<bar> \<le> Bnd \<Longrightarrow> - (floor Bnd) \<le> b"
using abs_le_D2 floor_mono by fastforce
lemma abs_pos_floor: "\<bar>of_int b\<bar> \<le> Bnd \<Longrightarrow> b \<le> floor Bnd"
using abs_le_D1 le_floor_iff by auto
context gram_schmidt_floor
begin
lemma integer_hull_integer_cone: assumes C: "C \<subseteq> carrier_vec n"
and CI: "C \<subseteq> \<int>\<^sub>v"
shows "integer_hull (cone C) = cone C"
proof
have "cone C \<inter> \<int>\<^sub>v \<subseteq> cone C" by blast
thus "integer_hull (cone C) \<subseteq> cone C"
using cone_cone[OF C] convex_cone[OF C] convex_hull_mono
unfolding integer_hull_def convex_def by metis
{
fix x
assume "x \<in> cone C"
then obtain D where finD: "finite D" and DC: "D \<subseteq> C" and x: "x \<in> finite_cone D"
unfolding cone_def by auto
from DC C CI have D: "D \<subseteq> carrier_vec n" and DI: "D \<subseteq> \<int>\<^sub>v" by auto
from D x finD have "x \<in> finite_cone (D \<union> {0\<^sub>v n})" using finite_cone_mono[of "D \<union> {0\<^sub>v n}" D] by auto
then obtain l where x: "lincomb l (D \<union> {0\<^sub>v n}) = x"
and l: "l ` (D \<union> {0\<^sub>v n}) \<subseteq> {t. t \<ge> 0}"
using finD unfolding finite_cone_def nonneg_lincomb_def by auto
define L where "L = sum l (D \<union> {0\<^sub>v n})"
define L_sup :: 'a where "L_sup = of_int (floor L + 1)"
have "L_sup \<ge> L" using floor_correct[of L] unfolding L_sup_def by linarith
have "L \<ge> 0" unfolding L_def using sum_nonneg[of _ l] l by blast
hence "L_sup \<ge> 1" unfolding L_sup_def by simp
hence "L_sup > 0" by fastforce
let ?f = "\<lambda> y. if y = 0\<^sub>v n then L_sup - L else 0"
have "lincomb ?f {0\<^sub>v n} = 0\<^sub>v n"
using already_in_span[of "{}" "0\<^sub>v n"] lincomb_in_span local.span_empty
by auto
moreover have "lincomb ?f (D - {0\<^sub>v n}) = 0\<^sub>v n"
by(rule lincomb_zero, insert D, auto)
ultimately have "lincomb ?f (D \<union> {0\<^sub>v n}) = 0\<^sub>v n"
using lincomb_vec_diff_add[of "D \<union> {0\<^sub>v n}" "{0\<^sub>v n}"] D finD by simp
hence lcomb_f: "lincomb (\<lambda> y. l y + ?f y) (D \<union> {0\<^sub>v n}) = x"
using lincomb_sum[of "D \<union> {0\<^sub>v n}" l ?f] finD D x by simp
have "sum ?f (D \<union> {0\<^sub>v n}) = L_sup - L"
by (simp add: sum.subset_diff[of "{0\<^sub>v n}" "D \<union> {0\<^sub>v n}" ?f] finD)
hence "sum (\<lambda> y. l y + ?f y) (D \<union> {0\<^sub>v n}) = L_sup"
using l L_def by auto
moreover have "(\<lambda> y. l y + ?f y) ` (D \<union> {0\<^sub>v n}) \<subseteq> {t. t \<ge> 0}"
- using `L \<le> L_sup` l by force
+ using \<open>L \<le> L_sup\<close> l by force
ultimately obtain l' where x: "lincomb l' (D \<union> {0\<^sub>v n}) = x"
and l': "l' ` (D \<union> {0\<^sub>v n}) \<subseteq> {t. t \<ge> 0}"
and sum_l': "sum l' (D \<union> {0\<^sub>v n}) = L_sup"
using lcomb_f by blast
let ?D' = "{L_sup \<cdot>\<^sub>v v | v. v \<in> D \<union> {0\<^sub>v n}}"
have Did: "?D' = (\<lambda> v. L_sup \<cdot>\<^sub>v v) ` (D \<union> {0\<^sub>v n})" by force
define l'' where "l'' = (\<lambda> y. l' ((1 / L_sup) \<cdot>\<^sub>v y) / L_sup)"
obtain lD where dist: "distinct lD" and lD: "D \<union> {0\<^sub>v n} = set lD"
using finite_distinct_list[of "D \<union> {0\<^sub>v n}"] finD by auto
let ?lD' = "map ((\<cdot>\<^sub>v) L_sup) lD"
have dist': "distinct ?lD'"
- using distinct_smult_nonneg[OF _ dist] `L_sup > 0` by fastforce
+ using distinct_smult_nonneg[OF _ dist] \<open>L_sup > 0\<close> by fastforce
have x': "lincomb l'' ?D' = x" unfolding x[symmetric] l''_def
unfolding lincomb_def Did
proof (subst finsum_reindex)
- from `L_sup > 0` smult_vec_nonneg_eq[of L_sup] show "inj_on ((\<cdot>\<^sub>v) L_sup) (D \<union> {0\<^sub>v n})"
+ from \<open>L_sup > 0\<close> smult_vec_nonneg_eq[of L_sup] show "inj_on ((\<cdot>\<^sub>v) L_sup) (D \<union> {0\<^sub>v n})"
by (auto simp: inj_on_def)
show "(\<lambda>v. l' (1 / L_sup \<cdot>\<^sub>v v) / L_sup \<cdot>\<^sub>v v) \<in> (\<cdot>\<^sub>v) L_sup ` (D \<union> {0\<^sub>v n}) \<rightarrow> carrier_vec n"
using D by auto
- from `L_sup > 0` have "L_sup \<noteq> 0" by auto
+ from \<open>L_sup > 0\<close> have "L_sup \<noteq> 0" by auto
then show "(\<Oplus>\<^bsub>V\<^esub>x\<in>D \<union> {0\<^sub>v n}. l' (1 / L_sup \<cdot>\<^sub>v (L_sup \<cdot>\<^sub>v x)) / L_sup \<cdot>\<^sub>v (L_sup \<cdot>\<^sub>v x)) =
(\<Oplus>\<^bsub>V\<^esub>v\<in>D \<union> {0\<^sub>v n}. l' v \<cdot>\<^sub>v v)"
by (intro finsum_cong, insert D, auto simp: smult_smult_assoc)
qed
have "D \<union> {0\<^sub>v n} \<subseteq> cone C" using set_in_cone[OF C] DC zero_in_cone by blast
- hence D': "?D' \<subseteq> cone C" using cone_smult[of L_sup, OF _ C] `L_sup > 0` by auto
+ hence D': "?D' \<subseteq> cone C" using cone_smult[of L_sup, OF _ C] \<open>L_sup > 0\<close> by auto
have "D \<union> {0\<^sub>v n} \<subseteq> \<int>\<^sub>v" unfolding zero_vec_def using DI Ints_vec_def by auto
moreover have "L_sup \<in> \<int>" unfolding L_sup_def by auto
ultimately have D'I: "?D' \<subseteq> \<int>\<^sub>v" unfolding Ints_vec_def by force
- have "1 = sum l' (D \<union> {0\<^sub>v n}) * (1 / L_sup)" using sum_l' `L_sup > 0` by auto
+ have "1 = sum l' (D \<union> {0\<^sub>v n}) * (1 / L_sup)" using sum_l' \<open>L_sup > 0\<close> by auto
also have "sum l' (D \<union> {0\<^sub>v n}) = sum_list (map l' lD)"
using sum.distinct_set_conv_list[OF dist] lD by auto
also have "map l' lD = map (l' \<circ> ((\<cdot>\<^sub>v) (1 / L_sup))) ?lD'"
- using smult_smult_assoc[of "1 / L_sup" L_sup] `L_sup > 0`
+ using smult_smult_assoc[of "1 / L_sup" L_sup] \<open>L_sup > 0\<close>
by (simp add: comp_assoc)
also have "l' \<circ> ((\<cdot>\<^sub>v) (1 / L_sup)) = (\<lambda> x. l' ((1 / L_sup) \<cdot>\<^sub>v x))" by (rule comp_def)
also have "sum_list (map \<dots> ?lD') * (1 / L_sup) =
sum_list (map (\<lambda>y. l' (1 / L_sup \<cdot>\<^sub>v y) * (1 / L_sup)) ?lD')"
using sum_list_mult_const[of _ "1 / L_sup" ?lD'] by presburger
also have "\<dots> = sum_list (map l'' ?lD')"
- unfolding l''_def using `L_sup > 0` by simp
+ unfolding l''_def using \<open>L_sup > 0\<close> by simp
also have "\<dots> = sum l'' (set ?lD')" using sum.distinct_set_conv_list[OF dist'] by metis
also have "set ?lD' = ?D'" using lD by auto
finally have sum_l': "sum l'' ?D' = 1" by auto
moreover have "l'' ` ?D' \<subseteq> {t. t \<ge> 0}"
proof
fix y
assume "y \<in> l'' ` ?D'"
then obtain x where y: "y = l'' x" and "x \<in> ?D'" by blast
then obtain v where "v \<in> D \<union> {0\<^sub>v n}" and x: "x = L_sup \<cdot>\<^sub>v v" by blast
- hence "0 \<le> l' v / L_sup" using l' `L_sup > 0` by fastforce
+ hence "0 \<le> l' v / L_sup" using l' \<open>L_sup > 0\<close> by fastforce
also have "\<dots> = l'' x" unfolding x l''_def
- using smult_smult_assoc[of "1 / L_sup" "L_sup" v] `L_sup > 0` by simp
+ using smult_smult_assoc[of "1 / L_sup" "L_sup" v] \<open>L_sup > 0\<close> by simp
finally show "y \<in> {t. t \<ge> 0}" using y by blast
qed
moreover have "finite ?D'" using finD by simp
ultimately have "x \<in> integer_hull (cone C)"
unfolding integer_hull_def convex_hull_def
using x' D' D'I convex_lincomb_def[of l'' ?D' x]
nonneg_lincomb_def[of l'' ?D' x] by fast
}
thus "cone C \<subseteq> integer_hull (cone C)" by blast
qed
theorem decomposition_theorem_integer_hull_of_polyhedron: assumes A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr"
and AI: "A \<in> \<int>\<^sub>m"
and bI: "b \<in> \<int>\<^sub>v"
and P: "P = polyhedron A b"
and Bnd: "Bnd = Max (abs ` (elements_mat A \<union> vec_set b))"
shows "\<exists> H C. H \<union> C \<subseteq> carrier_vec n \<inter> \<int>\<^sub>v
\<and> H \<subseteq> Bounded_vec (fact (n+1) * (max 1 Bnd)^n)
\<and> C \<subseteq> Bounded_vec (fact n * (max 1 Bnd)^n)
\<and> finite (H \<union> C)
\<and> integer_hull P = convex_hull H + cone C"
proof -
define DBnd where "DBnd = det_bound n (max 1 Bnd)"
define nBnd where "nBnd = of_nat (n+1) * DBnd"
have DBnd: "DBnd = fact n * (max 1 Bnd)^n" unfolding DBnd_def det_bound_def by auto
have nBnd: "nBnd = fact (n+1) * (max 1 Bnd)^n" unfolding nBnd_def DBnd
by auto
have DBnd0: "DBnd \<ge> 0" unfolding DBnd_def det_bound_def by auto
have Pn: "P \<subseteq> carrier_vec n" unfolding P polyhedron_def by auto
have "A \<in> Bounded_mat Bnd \<and> b \<in> Bounded_vec Bnd"
unfolding Bnd Bounded_mat_elements_mat Bounded_vec_vec_set
by (intro ballI conjI Max_ge finite_imageI imageI finite_UnI, auto)
from decomposition_theorem_polyhedra_1[OF A b P, of Bnd] AI bI this
obtain QQ Q C where C: "C \<subseteq> carrier_vec n" and finC: "finite C"
and QQ: "QQ \<subseteq> carrier_vec n" and finQ: "finite QQ" and BndQQ: "QQ \<subseteq> Bounded_vec DBnd"
and P: "P = Q + cone C"
and Q_def: "Q = convex_hull QQ"
and CI: "C \<subseteq> \<int>\<^sub>v" and BndC: "C \<subseteq> Bounded_vec DBnd"
by (auto simp: DBnd_def)
define Bnd' where "Bnd' = of_nat n * DBnd"
note coneC = cone_iff_finite_cone[OF C finC]
have Q: "Q \<subseteq> carrier_vec n" unfolding Q_def using convex_hull_carrier[OF QQ] .
define B where "B = {x. \<exists> a D. nonneg_lincomb a D x \<and> D \<subseteq> C \<and> lin_indpt D \<and> (\<forall> d \<in> D. a d \<le> 1)}"
{
fix b
assume "b \<in> B"
then obtain a D where b: "b = lincomb a D" and DC: "D \<subseteq> C"
and linD: "lin_indpt D" and bnd_a: "\<forall> d \<in> D. 0 \<le> a d \<and> a d \<le> 1"
by (force simp: B_def nonneg_lincomb_def)
from DC C have D: "D \<subseteq> carrier_vec n" by auto
from DC finC have finD: "finite D" by (metis finite_subset)
from D linD finD have cardD: "card D \<le> n" using dim_is_n li_le_dim(2) by auto
from BndC DC have BndD: "D \<subseteq> Bounded_vec DBnd" by auto
from lincomb_card_bound[OF this D DBnd0 _ cardD, of a, folded b] bnd_a
have "b \<in> Bounded_vec Bnd'" unfolding Bnd'_def by force
}
hence BndB: "B \<subseteq> Bounded_vec Bnd'" ..
from BndQQ have BndQ: "Q \<subseteq> Bounded_vec DBnd" unfolding Q_def using QQ by (metis convex_hull_bound)
have B: "B \<subseteq> carrier_vec n"
unfolding B_def nonneg_lincomb_def using C by auto
from Q B have QB: "Q + B \<subseteq> carrier_vec n" by (auto elim!: set_plus_elim)
from sum_in_Bounded_vecI[of _ DBnd _ Bnd' n] BndQ BndB B Q
have "Q + B \<subseteq> Bounded_vec (DBnd + Bnd')" by (auto elim!: set_plus_elim)
also have "DBnd + Bnd' = nBnd" unfolding nBnd_def Bnd'_def by (simp add: algebra_simps)
finally have QB_Bnd: "Q + B \<subseteq> Bounded_vec nBnd" by blast
have finQBZ: "finite ((Q + B) \<inter> \<int>\<^sub>v)"
proof (rule finite_subset[OF subsetI])
define ZBnd where "ZBnd = floor nBnd"
let ?vecs = "set (map vec_of_list (concat_lists (map (\<lambda> i. map (of_int :: _ \<Rightarrow> 'a) [-ZBnd..ZBnd]) [0..<n])))"
have id: "?vecs = vec_of_list `
{as. length as = n \<and> (\<forall>i<n. \<exists> b. as ! i = of_int b \<and> b \<in> {- ZBnd..ZBnd})}"
unfolding set_map by (rule image_cong, auto)
show "finite ?vecs" by (rule finite_set)
fix x
assume "x \<in> (Q + B) \<inter> \<int>\<^sub>v"
hence xQB: "x \<in> Q + B" and xI: "x \<in> \<int>\<^sub>v" by auto
from xQB QB_Bnd QB have xBnd: "x \<in> Bounded_vec nBnd" and x: "x \<in> carrier_vec n" by auto
have xid: "x = vec_of_list (list_of_vec x)" by auto
show "x \<in> ?vecs" unfolding id
proof (subst xid, intro imageI CollectI conjI allI impI)
show "length (list_of_vec x) = n" using x by auto
fix i
assume i: "i < n"
have id: "list_of_vec x ! i = x $ i" using i x by auto
from xBnd[unfolded Bounded_vec_def] i x have xiBnd: "abs (x $ i) \<le> nBnd" by auto
from xI[unfolded Ints_vec_def] i x have "x $ i \<in> \<int>" by auto
then obtain b where b: "x $ i = of_int b" unfolding Ints_def by blast
show "\<exists>b. list_of_vec x ! i = of_int b \<and> b \<in> {- ZBnd..ZBnd}" unfolding id ZBnd_def
using xiBnd unfolding b by (intro exI[of _ b], auto intro!: abs_neg_floor abs_pos_floor)
qed
qed
have QBZ: "(Q + B) \<inter> \<int>\<^sub>v \<subseteq> carrier_vec n" using QB by auto
from decomposition_theorem_polyhedra_2[OF QBZ finQBZ, folded integer_hull_def, OF C finC refl]
obtain A' b' nr' where A': "A' \<in> carrier_mat nr' n" and b': "b' \<in> carrier_vec nr'"
and IH: "integer_hull (Q + B) + cone C = polyhedron A' b'"
by auto
{
fix p
assume "p \<in> P \<inter> \<int>\<^sub>v"
hence pI: "p \<in> \<int>\<^sub>v" and p: "p \<in> Q + cone C" unfolding P by auto
from set_plus_elim[OF p] obtain q c where
pqc: "p = q + c" and qQ: "q \<in> Q" and cC: "c \<in> cone C" by auto
from qQ Q have q: "q \<in> carrier_vec n" by auto
from Caratheodory_theorem[OF C] cC
obtain D where cD: "c \<in> finite_cone D" and DC: "D \<subseteq> C" and linD: "lin_indpt D" by auto
from DC C have D: "D \<subseteq> carrier_vec n" by auto
from DC finC have finD: "finite D" by (metis finite_subset)
from cD finD
obtain a where "nonneg_lincomb a D c" unfolding finite_cone_def by auto
hence caD: "c = lincomb a D" and a0: "\<And> d. d \<in> D \<Longrightarrow> a d \<ge> 0"
unfolding nonneg_lincomb_def by auto
define a1 where "a1 = (\<lambda> c. a c - of_int (floor (a c)))"
define a2 where "a2 = (\<lambda> c. of_int (floor (a c)) :: 'a)"
define d where "d = lincomb a2 D"
define b where "b = lincomb a1 D"
have b: "b \<in> carrier_vec n" and d: "d \<in> carrier_vec n" unfolding d_def b_def using D by auto
have bB: "b \<in> B" unfolding B_def b_def nonneg_lincomb_def
proof (intro CollectI exI[of _ a1] exI[of _ D] conjI ballI refl subsetI linD)
show "x \<in> a1 ` D \<Longrightarrow> 0 \<le> x" for x using a0 unfolding a1_def by auto
show "a1 c \<le> 1" for c unfolding a1_def by linarith
qed (insert DC, auto)
have cbd: "c = b + d" unfolding b_def d_def caD lincomb_sum[OF finD D, symmetric]
by (rule lincomb_cong[OF refl D], auto simp: a1_def a2_def)
have "nonneg_lincomb a2 D d" unfolding d_def nonneg_lincomb_def
by (intro allI conjI refl subsetI, insert a0, auto simp: a2_def)
hence dC: "d \<in> cone C" unfolding cone_def finite_cone_def using finC finD DC by auto
have p: "p = (q + b) + d" unfolding pqc cbd using q b d by auto
have dI: "d \<in> \<int>\<^sub>v" using CI DC C unfolding d_def indexed_Ints_vec_UNIV
by (intro lincomb_indexed_Ints_vec, auto simp: a2_def)
from diff_indexed_Ints_vec[of _ _ _ UNIV, folded indexed_Ints_vec_UNIV, OF _ d pI dI, unfolded p]
have "q + b + d - d \<in> \<int>\<^sub>v" using q b d by auto
also have "q + b + d - d = q + b" using q b d by auto
finally have qbI: "q + b \<in> \<int>\<^sub>v" by auto
have "p \<in> integer_hull (Q + B) + cone C" unfolding p integer_hull_def
by (intro set_plus_intro dC set_mp[OF set_in_convex_hull] IntI qQ bB qbI, insert Q B,
auto elim!: set_plus_elim)
}
hence "P \<inter> \<int>\<^sub>v \<subseteq> integer_hull (Q + B) + cone C" by auto
hence one_dir: "integer_hull P \<subseteq> integer_hull (Q + B) + cone C" unfolding IH
unfolding integer_hull_def using convex_convex_hull[OF polyhedra_are_convex[OF A' b' refl]]
convex_hull_mono by blast
have "integer_hull (Q + B) + cone C \<subseteq> integer_hull P + cone C" unfolding P
proof (intro set_plus_mono2 subset_refl integer_hull_mono)
show "B \<subseteq> cone C" unfolding B_def cone_def finite_cone_def using finite_subset[OF _ finC] by auto
qed
also have "\<dots> = integer_hull P + integer_hull (cone C)"
using integer_hull_integer_cone[OF C CI] by simp
also have "\<dots> = convex_hull (P \<inter> \<int>\<^sub>v) + convex_hull (cone C \<inter> \<int>\<^sub>v)"
unfolding integer_hull_def by simp
also have "\<dots> = convex_hull ((P \<inter> \<int>\<^sub>v) + (cone C \<inter> \<int>\<^sub>v))"
by (rule convex_hull_sum[symmetric], insert Pn cone_carrier[OF C], auto)
also have "\<dots> \<subseteq> convex_hull ((P + cone C) \<inter> \<int>\<^sub>v)"
proof (rule convex_hull_mono)
show "P \<inter> \<int>\<^sub>v + cone C \<inter> \<int>\<^sub>v \<subseteq> (P + cone C) \<inter> \<int>\<^sub>v"
using add_indexed_Ints_vec[of _ n _ UNIV, folded indexed_Ints_vec_UNIV] cone_carrier[OF C] Pn
by (auto elim!: set_plus_elim)
qed
also have "\<dots> = integer_hull (P + cone C)" unfolding integer_hull_def ..
also have "P + cone C = P"
proof -
have CC: "cone C \<subseteq> carrier_vec n" using C by (rule cone_carrier)
have "P + cone C = Q + (cone C + cone C)" unfolding P
by (rule assoc_add_vecset[symmetric, OF Q CC CC])
also have "cone C + cone C = cone C" by (rule cone_add_cone[OF C])
finally show ?thesis unfolding P .
qed
finally have "integer_hull (Q + B) + cone C \<subseteq> integer_hull P" .
with one_dir have id: "integer_hull P = integer_hull (Q + B) + cone C" by auto
show ?thesis unfolding id unfolding integer_hull_def nBnd[symmetric] DBnd[symmetric]
proof (rule exI[of _ "(Q + B) \<inter> \<int>\<^sub>v"], intro exI[of _ C] conjI refl BndC)
from QB_Bnd show "(Q + B) \<inter> \<int>\<^sub>v \<subseteq> Bounded_vec nBnd" by auto
show "(Q + B) \<inter> \<int>\<^sub>v \<union> C \<subseteq> carrier_vec n \<inter> \<int>\<^sub>v"
using QB C CI by auto
show "finite ((Q + B) \<inter> \<int>\<^sub>v \<union> C)" using finQBZ finC by auto
qed
qed
corollary integer_hull_of_polyhedron: assumes A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr"
and AI: "A \<in> \<int>\<^sub>m"
and bI: "b \<in> \<int>\<^sub>v"
and P: "P = polyhedron A b"
shows "\<exists> A' b' nr'. A' \<in> carrier_mat nr' n \<and> b' \<in> carrier_vec nr' \<and> integer_hull P = polyhedron A' b'"
proof -
from decomposition_theorem_integer_hull_of_polyhedron[OF A b AI bI P refl]
obtain H C
where HC: "H \<union> C \<subseteq> carrier_vec n \<inter> \<int>\<^sub>v" "finite (H \<union> C)"
and decomp: "integer_hull P = convex_hull H + cone C" by auto
show ?thesis
by (rule decomposition_theorem_polyhedra_2[OF _ _ _ _ decomp], insert HC, auto)
qed
corollary small_integer_solution_nonstrict_via_decomp: fixes A :: "'a mat"
assumes A: "A \<in> carrier_mat nr n"
and b: "b \<in> carrier_vec nr"
and AI: "A \<in> \<int>\<^sub>m"
and bI: "b \<in> \<int>\<^sub>v"
and Bnd: "Bnd = Max (abs ` (elements_mat A \<union> vec_set b))"
and x: "x \<in> carrier_vec n"
and xI: "x \<in> \<int>\<^sub>v"
and sol: "A *\<^sub>v x \<le> b"
shows "\<exists> y.
y \<in> carrier_vec n \<and>
y \<in> \<int>\<^sub>v \<and>
A *\<^sub>v y \<le> b \<and>
y \<in> Bounded_vec (fact (n+1) * (max 1 Bnd)^n)"
proof -
from x sol have "x \<in> polyhedron A b" unfolding polyhedron_def by auto
with xI x have xsol: "x \<in> integer_hull (polyhedron A b)" unfolding integer_hull_def
by (meson IntI convex_hull_mono in_mono inf_sup_ord(1) inf_sup_ord(2) set_in_convex_hull)
from decomposition_theorem_integer_hull_of_polyhedron[OF A b AI bI refl Bnd]
obtain H C where HC: "H \<union> C \<subseteq> carrier_vec n \<inter> \<int>\<^sub>v"
"H \<subseteq> Bounded_vec (fact (n + 1) * max 1 Bnd ^ n)"
"finite (H \<union> C)" and
id: "integer_hull (polyhedron A b) = convex_hull H + cone C"
by auto
from xsol[unfolded id] have "H \<noteq> {}" unfolding set_plus_def by auto
then obtain h where hH: "h \<in> H" by auto
with set_in_convex_hull have "h \<in> convex_hull H" using HC by auto
moreover have "0\<^sub>v n \<in> cone C" by (intro zero_in_cone)
ultimately have "h + 0\<^sub>v n \<in> integer_hull (polyhedron A b)" unfolding id by auto
also have "h + 0\<^sub>v n = h" using hH HC by auto
also have "integer_hull (polyhedron A b) \<subseteq> convex_hull (polyhedron A b)"
unfolding integer_hull_def by (rule convex_hull_mono, auto)
also have "convex_hull (polyhedron A b) = polyhedron A b" using A b
using convex_convex_hull polyhedra_are_convex by blast
finally have h: "h \<in> carrier_vec n" "A *\<^sub>v h \<le> b" unfolding polyhedron_def by auto
show ?thesis
by (intro exI[of _ h] conjI h, insert HC hH, auto)
qed
end
end
\ No newline at end of file
diff --git a/thys/Optics/Two.thy b/thys/Optics/Two.thy
--- a/thys/Optics/Two.thy
+++ b/thys/Optics/Two.thy
@@ -1,50 +1,50 @@
-section {* Types of Cardinality 2 or Greater *}
+section \<open>Types of Cardinality 2 or Greater\<close>
theory Two
imports HOL.Real
begin
-text {* The two class states that a type's carrier is either infinite, or else it has a finite
+text \<open>The two class states that a type's carrier is either infinite, or else it has a finite
cardinality of at least 2. It is needed when we depend on having at least two distinguishable
- elements. *}
+ elements.\<close>
class two =
assumes card_two: "infinite (UNIV :: 'a set) \<or> card (UNIV :: 'a set) \<ge> 2"
begin
lemma two_diff: "\<exists> x y :: 'a. x \<noteq> y"
proof -
obtain A where "finite A" "card A = 2" "A \<subseteq> (UNIV :: 'a set)"
proof (cases "infinite (UNIV :: 'a set)")
case True
with infinite_arbitrarily_large[of "UNIV :: 'a set" 2] that
show ?thesis by auto
next
case False
with card_two that
show ?thesis
by (metis UNIV_bool card_UNIV_bool card_image card_le_inj finite.intros(1) finite_insert finite_subset)
qed
thus ?thesis
by (metis (full_types) One_nat_def Suc_1 UNIV_eq_I card.empty card.insert finite.intros(1) insertCI nat.inject nat.simps(3))
qed
end
instance bool :: two
by (intro_classes, auto)
instance nat :: two
by (intro_classes, auto)
instance int :: two
by (intro_classes, auto simp add: infinite_UNIV_int)
instance rat :: two
by (intro_classes, auto simp add: infinite_UNIV_char_0)
instance real :: two
by (intro_classes, auto simp add: infinite_UNIV_char_0)
instance list :: (type) two
by (intro_classes, auto simp add: infinite_UNIV_listI)
end
diff --git a/thys/Poincare_Disc/Poincare_Between.thy b/thys/Poincare_Disc/Poincare_Between.thy
--- a/thys/Poincare_Disc/Poincare_Between.thy
+++ b/thys/Poincare_Disc/Poincare_Between.thy
@@ -1,1262 +1,1262 @@
theory Poincare_Between
imports Poincare_Distance
begin
(* ------------------------------------------------------------------ *)
section\<open>H-betweenness in the Poincar\'e model\<close>
(* ------------------------------------------------------------------ *)
subsection \<open>H-betwenness expressed by a cross-ratio\<close>
text\<open>The point $v$ is h-between $u$ and $w$ if the cross-ratio between the pairs $u$ and $w$ and $v$
and inverse of $v$ is real and negative.\<close>
definition poincare_between :: "complex_homo \<Rightarrow> complex_homo \<Rightarrow> complex_homo \<Rightarrow> bool" where
"poincare_between u v w \<longleftrightarrow>
u = v \<or> v = w \<or>
(let cr = cross_ratio u v w (inversion v)
in is_real (to_complex cr) \<and> Re (to_complex cr) < 0)"
subsubsection \<open>H-betwenness is preserved by h-isometries\<close>
text \<open>Since they preserve cross-ratio and inversion, h-isometries (unit disc preserving Möbius
transformations and conjugation) preserve h-betweeness.\<close>
lemma unit_disc_fix_moebius_preserve_poincare_between [simp]:
assumes "unit_disc_fix M" and "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M w) \<longleftrightarrow>
poincare_between u v w"
proof (cases "u = v \<or> v = w")
case True
thus ?thesis
using assms
unfolding poincare_between_def
by auto
next
case False
moreover
hence "moebius_pt M u \<noteq> moebius_pt M v \<and> moebius_pt M v \<noteq> moebius_pt M w"
by auto
moreover
have "v \<noteq> inversion v" "w \<noteq> inversion v"
using inversion_noteq_unit_disc[of v w]
using inversion_noteq_unit_disc[of v v]
using \<open>v \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>
by auto
ultimately
show ?thesis
using assms
using unit_circle_fix_moebius_pt_inversion[of M v, symmetric]
unfolding poincare_between_def
by (simp del: unit_circle_fix_moebius_pt_inversion)
qed
lemma conjugate_preserve_poincare_between [simp]:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_between (conjugate u) (conjugate v) (conjugate w) \<longleftrightarrow>
poincare_between u v w"
proof (cases "u = v \<or> v = w")
case True
thus ?thesis
using assms
unfolding poincare_between_def
by auto
next
case False
moreover
hence "conjugate u \<noteq> conjugate v \<and> conjugate v \<noteq> conjugate w"
using conjugate_inj by blast
moreover
have "v \<noteq> inversion v" "w \<noteq> inversion v"
using inversion_noteq_unit_disc[of v w]
using inversion_noteq_unit_disc[of v v]
using \<open>v \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>
by auto
ultimately
show ?thesis
using assms
using conjugate_cross_ratio[of v w "inversion v" u]
unfolding poincare_between_def
by (metis conjugate_id_iff conjugate_involution inversion_def inversion_sym o_apply)
qed
subsubsection \<open>Some elementary properties of h-betwenness\<close>
lemma poincare_between_nonstrict [simp]:
shows "poincare_between u u v" and "poincare_between u v v"
by (simp_all add: poincare_between_def)
lemma poincare_between_sandwich:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
assumes "poincare_between u v u"
shows "u = v"
proof (rule ccontr)
assume "\<not> ?thesis"
thus False
using assms
using inversion_noteq_unit_disc[of v u]
using cross_ratio_1[of v u "inversion v"]
unfolding poincare_between_def Let_def
by auto
qed
lemma poincare_between_rev:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_between u v w \<longleftrightarrow> poincare_between w v u"
using assms
using inversion_noteq_unit_disc[of v w]
using inversion_noteq_unit_disc[of v u]
using cross_ratio_commute_13[of u v w "inversion v"]
using cross_ratio_not_inf[of w "inversion v" v u]
using cross_ratio_not_zero[of w v u "inversion v"]
using inf_or_of_complex[of "cross_ratio w v u (inversion v)"]
unfolding poincare_between_def
by (auto simp add: Let_def Im_complex_div_eq_0 Re_divide divide_less_0_iff)
subsubsection \<open>H-betwenness and h-collinearity\<close>
text\<open>Three points can be in an h-between relation only when they are h-collinear.\<close>
lemma poincare_between_poincare_collinear [simp]:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc" "w \<in> unit_disc"
assumes betw: "poincare_between u v w"
shows "poincare_collinear {u, v, w}"
proof (cases "u = v \<or> v = w")
case True
thus ?thesis
using assms
by auto
next
case False
hence distinct: "distinct [u, v, w, inversion v]"
using in_disc inversion_noteq_unit_disc[of v v] inversion_noteq_unit_disc[of v u] inversion_noteq_unit_disc[of v w]
using betw poincare_between_sandwich[of w v]
by (auto simp add: poincare_between_def Let_def)
then obtain H where *: "{u, v, w, inversion v} \<subseteq> circline_set H"
using assms
unfolding poincare_between_def
using four_points_on_circline_iff_cross_ratio_real[of u v w "inversion v"]
by auto
hence "H = poincare_line u v"
using assms distinct
using unique_circline_set[of u v "inversion v"]
using poincare_line[of u v] poincare_line_inversion[of u v]
unfolding circline_set_def
by auto
thus ?thesis
using * assms False
unfolding poincare_collinear_def
by (rule_tac x="poincare_line u v" in exI) simp
qed
lemma poincare_between_poincare_line_uvz:
assumes "u \<noteq> v" and "u \<in> unit_disc" and "v \<in> unit_disc" and
"z \<in> unit_disc" and "poincare_between u v z"
shows "z \<in> circline_set (poincare_line u v)"
using assms
using poincare_between_poincare_collinear[of u v z]
using unique_poincare_line[OF assms(1-3)]
unfolding poincare_collinear_def
by auto
lemma poincare_between_poincare_line_uzv:
assumes "u \<noteq> v" and "u \<in> unit_disc" and "v \<in> unit_disc" and
"z \<in> unit_disc" "poincare_between u z v"
shows "z \<in> circline_set (poincare_line u v)"
using assms
using poincare_between_poincare_collinear[of u z v]
using unique_poincare_line[OF assms(1-3)]
unfolding poincare_collinear_def
by auto
subsubsection \<open>H-betweeness on Euclidean segments\<close>
text\<open>If the three points lie on an h-line that is a Euclidean line (e.g., if it contains zero),
h-betweenness can be characterized much simpler than in the definition.\<close>
lemma poincare_between_x_axis_u0v:
assumes "is_real u'" and "u' \<noteq> 0" and "v' \<noteq> 0"
shows "poincare_between (of_complex u') 0\<^sub>h (of_complex v') \<longleftrightarrow> is_real v' \<and> Re u' * Re v' < 0"
proof-
have "Re u' \<noteq> 0"
using \<open>is_real u'\<close> \<open>u' \<noteq> 0\<close>
using complex_eq_if_Re_eq
by auto
have nz: "of_complex u' \<noteq> 0\<^sub>h" "of_complex v' \<noteq> 0\<^sub>h"
by (simp_all add: \<open>u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close>)
hence "0\<^sub>h \<noteq> of_complex v'"
by metis
let ?cr = "cross_ratio (of_complex u') 0\<^sub>h (of_complex v') \<infinity>\<^sub>h"
have "is_real (to_complex ?cr) \<and> Re (to_complex ?cr) < 0 \<longleftrightarrow> is_real v' \<and> Re u' * Re v' < 0"
using cross_ratio_0inf[of v' u'] \<open>v' \<noteq> 0\<close> \<open>u' \<noteq> 0\<close> \<open>is_real u'\<close>
by (metis Re_complex_div_lt_0 Re_mult_real complex_cnj_divide divide_cancel_left eq_cnj_iff_real to_complex_of_complex)
thus ?thesis
unfolding poincare_between_def inversion_zero
using \<open>of_complex u' \<noteq> 0\<^sub>h\<close> \<open>0\<^sub>h \<noteq> of_complex v'\<close>
by simp
qed
lemma poincare_between_u0v:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> 0\<^sub>h" and "v \<noteq> 0\<^sub>h"
shows "poincare_between u 0\<^sub>h v \<longleftrightarrow> (\<exists> k < 0. to_complex u = cor k * to_complex v)" (is "?P u v")
proof (cases "u = v")
case True
thus ?thesis
using assms
using inf_or_of_complex[of v]
using poincare_between_sandwich[of u "0\<^sub>h"]
by auto
next
case False
have "\<forall> u. u \<in> unit_disc \<and> u \<noteq> 0\<^sub>h \<longrightarrow> ?P u v" (is "?P' v")
proof (rule wlog_rotation_to_positive_x_axis)
fix \<phi> v
let ?M = "moebius_pt (moebius_rotation \<phi>)"
assume 1: "v \<in> unit_disc" "v \<noteq> 0\<^sub>h"
assume 2: "?P' (?M v)"
show "?P' v"
proof (rule allI, rule impI, (erule conjE)+)
fix u
assume 3: "u \<in> unit_disc" "u \<noteq> 0\<^sub>h"
have "poincare_between (?M u) 0\<^sub>h (?M v) \<longleftrightarrow> poincare_between u 0\<^sub>h v"
using \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close>
using unit_disc_fix_moebius_preserve_poincare_between unit_disc_fix_rotation zero_in_unit_disc
by fastforce
thus "?P u v"
using 1 2[rule_format, of "?M u"] 3
using inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
qed
next
fix x
assume 1: "is_real x" "0 < Re x" "Re x < 1"
hence "x \<noteq> 0"
by auto
show "?P' (of_complex x)"
proof (rule allI, rule impI, (erule conjE)+)
fix u
assume 2: "u \<in> unit_disc" "u \<noteq> 0\<^sub>h"
then obtain u' where "u = of_complex u'"
using inf_or_of_complex[of u]
by auto
show "?P u (of_complex x)"
using 1 2 \<open>x \<noteq> 0\<close> \<open>u = of_complex u'\<close>
using poincare_between_rev[of u "0\<^sub>h" "of_complex x"]
using poincare_between_x_axis_u0v[of x u'] \<open>is_real x\<close>
apply (auto simp add: cmod_eq_Re)
apply (rule_tac x="Re u' / Re x" in exI, simp add: divide_neg_pos algebra_split_simps)
using mult_neg_pos mult_pos_neg
by blast
qed
qed fact+
thus ?thesis
using assms
by auto
qed
lemma poincare_between_u0v_polar_form:
assumes "x \<in> unit_disc" and "y \<in> unit_disc" and "x \<noteq> 0\<^sub>h" and "y \<noteq> 0\<^sub>h" and
"to_complex x = cor rx * cis \<phi>" "to_complex y = cor ry * cis \<phi>"
shows "poincare_between x 0\<^sub>h y \<longleftrightarrow> rx * ry < 0" (is "?P x y rx ry")
proof-
from assms have "rx \<noteq> 0" "ry \<noteq> 0"
using inf_or_of_complex[of x] inf_or_of_complex[of y]
by auto
have "(\<exists>k<0. cor rx = cor k * cor ry ) = (rx * ry < 0)"
proof
assume "\<exists>k<0. cor rx = cor k * cor ry"
then obtain k where "k < 0" "cor rx = cor k * cor ry"
by auto
hence "rx = k * ry"
using of_real_eq_iff
by fastforce
thus "rx * ry < 0"
using \<open>k < 0\<close> \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (smt divisors_zero mult_nonneg_nonpos mult_nonpos_nonpos zero_less_mult_pos2)
next
assume "rx * ry < 0"
hence "rx = (rx/ry)*ry" "rx / ry < 0"
using \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (auto simp add: divide_less_0_iff algebra_split_simps)
thus "\<exists>k<0. cor rx = cor k * cor ry"
using \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (rule_tac x="rx / ry" in exI, simp)
qed
thus ?thesis
using assms
using poincare_between_u0v[OF assms(1-4)]
by auto
qed
lemma poincare_between_x_axis_0uv:
fixes x y :: real
assumes "-1 < x" and "x < 1" and "x \<noteq> 0"
assumes "-1 < y" and "y < 1" and "y \<noteq> 0"
shows "poincare_between 0\<^sub>h (of_complex x) (of_complex y) \<longleftrightarrow>
(x < 0 \<and> y < 0 \<and> y \<le> x) \<or> (x > 0 \<and> y > 0 \<and> x \<le> y)" (is "?lhs \<longleftrightarrow> ?rhs")
proof (cases "x = y")
case True
thus ?thesis
using assms
unfolding poincare_between_def
by auto
next
case False
let ?x = "of_complex x" and ?y = "of_complex y"
have "?x \<in> unit_disc" "?y \<in> unit_disc"
using assms
by auto
have distinct: "distinct [0\<^sub>h, ?x, ?y, inversion ?x]"
using \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>x \<noteq> y\<close> \<open>?x \<in> unit_disc\<close> \<open>?y \<in> unit_disc\<close>
using inversion_noteq_unit_disc[of ?x ?y]
using inversion_noteq_unit_disc[of ?x ?x]
using inversion_noteq_unit_disc[of ?x "0\<^sub>h"]
using of_complex_inj[of x y]
by (metis distinct_length_2_or_more distinct_singleton of_complex_zero_iff of_real_eq_0_iff of_real_eq_iff zero_in_unit_disc)
let ?cr = "cross_ratio 0\<^sub>h ?x ?y (inversion ?x)"
have "Re (to_complex ?cr) = x\<^sup>2 * (x*y - 1) / (x * (y - x))"
using \<open>x \<noteq> 0\<close> \<open>x \<noteq> y\<close>
unfolding inversion_def
by simp (transfer, transfer, auto simp add: vec_cnj_def power2_eq_square field_simps split: if_split_asm)
moreover
{
fix a b :: real
assume "b \<noteq> 0"
hence "a < 0 \<longleftrightarrow> b\<^sup>2 * a < (0::real)"
by (metis mult.commute mult_eq_0_iff mult_neg_pos mult_pos_pos not_less_iff_gr_or_eq not_real_square_gt_zero power2_eq_square)
}
hence "x\<^sup>2 * (x*y - 1) < 0"
using assms
by (smt minus_mult_minus mult_le_cancel_left1)
moreover
have "x * (y - x) > 0 \<longleftrightarrow> ?rhs"
using \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>x \<noteq> y\<close>
by (smt mult_le_0_iff)
ultimately
have *: "Re (to_complex ?cr) < 0 \<longleftrightarrow> ?rhs"
by (simp add: divide_less_0_iff)
show ?thesis
proof
assume ?lhs
have "is_real (to_complex ?cr)" "Re (to_complex ?cr) < 0"
using \<open>?lhs\<close> distinct
unfolding poincare_between_def Let_def
by auto
thus ?rhs
using *
by simp
next
assume ?rhs
hence "Re (to_complex ?cr) < 0"
using *
by simp
moreover
have "{0\<^sub>h, of_complex (cor x), of_complex (cor y), inversion (of_complex (cor x))} \<subseteq> circline_set x_axis"
using \<open>x \<noteq> 0\<close> is_real_inversion[of "cor x"]
using inf_or_of_complex[of "inversion ?x"]
by (auto simp del: inversion_of_complex)
hence "is_real (to_complex ?cr)"
using four_points_on_circline_iff_cross_ratio_real[OF distinct]
by auto
ultimately
show ?lhs
using distinct
unfolding poincare_between_def Let_def
by auto
qed
qed
lemma poincare_between_0uv:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> 0\<^sub>h" and "v \<noteq> 0\<^sub>h"
shows "poincare_between 0\<^sub>h u v \<longleftrightarrow>
(let u' = to_complex u; v' = to_complex v in arg u' = arg v' \<and> cmod u' \<le> cmod v')" (is "?P u v")
proof (cases "u = v")
case True
thus ?thesis
by simp
next
case False
have "\<forall> v. v \<in> unit_disc \<and> v \<noteq> 0\<^sub>h \<and> v \<noteq> u \<longrightarrow> (poincare_between 0\<^sub>h u v \<longleftrightarrow> (let u' = to_complex u; v' = to_complex v in arg u' = arg v' \<and> cmod u' \<le> cmod v'))" (is "?P' u")
proof (rule wlog_rotation_to_positive_x_axis)
show "u \<in> unit_disc" "u \<noteq> 0\<^sub>h"
by fact+
next
fix x
assume *: "is_real x" "0 < Re x" "Re x < 1"
hence "of_complex x \<in> unit_disc" "of_complex x \<noteq> 0\<^sub>h" "of_complex x \<in> circline_set x_axis"
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re)
show "?P' (of_complex x)"
proof safe
fix v
assume "v \<in> unit_disc" "v \<noteq> 0\<^sub>h" "v \<noteq> of_complex x" "poincare_between 0\<^sub>h (of_complex x) v"
hence "v \<in> circline_set x_axis"
using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex x" v]
using poincare_line_0_real_is_x_axis[of "of_complex x"]
using \<open>of_complex x \<noteq> 0\<^sub>h\<close> \<open>v \<noteq> 0\<^sub>h\<close> \<open>v \<noteq> of_complex x\<close> \<open>of_complex x \<in> unit_disc\<close> \<open>of_complex x \<in> circline_set x_axis\<close>
by auto
obtain v' where "v = of_complex v'"
using \<open>v \<in> unit_disc\<close>
using inf_or_of_complex[of v]
by auto
hence **: "v = of_complex v'" "-1 < Re v'" "Re v' < 1" "Re v' \<noteq> 0" "is_real v'"
using \<open>v \<in> unit_disc\<close> \<open>v \<noteq> 0\<^sub>h\<close> \<open>v \<in> circline_set x_axis\<close> of_complex_inj[of v']
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re real_imag_0)
show "let u' = to_complex (of_complex x); v' = to_complex v in arg u' = arg v' \<and> cmod u' \<le> cmod v'"
using poincare_between_x_axis_0uv[of "Re x" "Re v'"] * **
using \<open>poincare_between 0\<^sub>h (of_complex x) v\<close>
using arg_complex_of_real_positive[of "Re x"] arg_complex_of_real_negative[of "Re x"]
using arg_complex_of_real_positive[of "Re v'"] arg_complex_of_real_negative[of "Re v'"]
by (auto simp add: cmod_eq_Re)
next
fix v
assume "v \<in> unit_disc" "v \<noteq> 0\<^sub>h" "v \<noteq> of_complex x"
then obtain v' where **: "v = of_complex v'" "v' \<noteq> 0" "v' \<noteq> x"
using inf_or_of_complex[of v]
by auto blast
assume "let u' = to_complex (of_complex x); v' = to_complex v in arg u' = arg v' \<and> cmod u' \<le> cmod v'"
hence ***: "Re x < 0 \<and> Re v' < 0 \<and> Re v' \<le> Re x \<or> 0 < Re x \<and> 0 < Re v' \<and> Re x \<le> Re v'" "is_real v'"
using arg_pi_iff[of x] arg_pi_iff[of v']
using arg_0_iff[of x] arg_0_iff[of v']
using * **
by (smt cmod_Re_le_iff to_complex_of_complex)+
have "-1 < Re v'" "Re v' < 1" "Re v' \<noteq> 0" "is_real v'"
using \<open>v \<in> unit_disc\<close> ** \<open>is_real v'\<close>
by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq)
thus "poincare_between 0\<^sub>h (of_complex x) v"
using poincare_between_x_axis_0uv[of "Re x" "Re v'"] * ** ***
by simp
qed
next
fix \<phi> u
assume "u \<in> unit_disc" "u \<noteq> 0\<^sub>h"
let ?M = "moebius_rotation \<phi>"
assume *: "?P' (moebius_pt ?M u)"
show "?P' u"
proof (rule allI, rule impI, (erule conjE)+)
fix v
assume "v \<in> unit_disc" "v \<noteq> 0\<^sub>h" "v \<noteq> u"
have "moebius_pt ?M v \<noteq> moebius_pt ?M u"
using \<open>v \<noteq> u\<close>
by auto
obtain u' v' where "v = of_complex v'" "u = of_complex u'" "v' \<noteq> 0" "u' \<noteq> 0"
using inf_or_of_complex[of u] inf_or_of_complex[of v]
using \<open>v \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close> \<open>v \<noteq> 0\<^sub>h\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
thus "?P u v"
using *[rule_format, of "moebius_pt ?M v"]
using \<open>moebius_pt ?M v \<noteq> moebius_pt ?M u\<close>
using unit_disc_fix_moebius_preserve_poincare_between[of ?M "0\<^sub>h" u v]
using \<open>v \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close> \<open>v \<noteq> 0\<^sub>h\<close> \<open>u \<noteq> 0\<^sub>h\<close>
using arg_mult_eq[of "cis \<phi>" u' v']
by simp (auto simp add: arg_mult)
qed
qed
thus ?thesis
using assms False
by auto
qed
lemma poincare_between_y_axis_0uv:
fixes x y :: real
assumes "-1 < x" and "x < 1" and "x \<noteq> 0"
assumes "-1 < y" and "y < 1" and "y \<noteq> 0"
shows "poincare_between 0\<^sub>h (of_complex (\<i> * x)) (of_complex (\<i> * y)) \<longleftrightarrow>
(x < 0 \<and> y < 0 \<and> y \<le> x) \<or> (x > 0 \<and> y > 0 \<and> x \<le> y)" (is "?lhs \<longleftrightarrow> ?rhs")
using assms
using poincare_between_0uv[of "of_complex (\<i> * x)" "of_complex (\<i> * y)"]
using arg_pi2_iff[of "\<i> * cor x"] arg_pi2_iff[of "\<i> * cor y"]
using arg_minus_pi2_iff[of "\<i> * cor x"] arg_minus_pi2_iff[of "\<i> * cor y"]
apply simp
apply (cases "x > 0")
apply (cases "y > 0", simp, simp)
apply (cases "y > 0")
apply simp
using pi_gt_zero apply linarith
apply simp
done
lemma poincare_between_x_axis_uvw:
fixes x y z :: real
assumes "-1 < x" and "x < 1"
assumes "-1 < y" and "y < 1" and "y \<noteq> x"
assumes "-1 < z" and "z < 1" and "z \<noteq> x"
shows "poincare_between (of_complex x) (of_complex y) (of_complex z) \<longleftrightarrow>
(y < x \<and> z < x \<and> z \<le> y) \<or> (y > x \<and> z > x \<and> y \<le> z)" (is "?lhs \<longleftrightarrow> ?rhs")
proof (cases "x = 0 \<or> y = 0 \<or> z = 0")
case True
thus ?thesis
proof (cases "x = 0")
case True
thus ?thesis
using poincare_between_x_axis_0uv assms
by simp
next
case False
show ?thesis
proof (cases "z = 0")
case True
thus ?thesis
using poincare_between_x_axis_0uv assms poincare_between_rev
by (smt norm_of_real of_complex_zero of_real_0 poincare_between_nonstrict(2) unit_disc_iff_cmod_lt_1)
next
case False
have "y = 0"
- using `x \<noteq> 0` `z \<noteq> 0` `x = 0 \<or> y = 0 \<or> z = 0`
+ using \<open>x \<noteq> 0\<close> \<open>z \<noteq> 0\<close> \<open>x = 0 \<or> y = 0 \<or> z = 0\<close>
by simp
have "poincare_between (of_complex x) 0\<^sub>h (of_complex z) = (is_real z \<and> x * z < 0)"
- using `x \<noteq> 0` `z \<noteq> 0` poincare_between_x_axis_u0v
+ using \<open>x \<noteq> 0\<close> \<open>z \<noteq> 0\<close> poincare_between_x_axis_u0v
by auto
moreover
have "x * z < 0 \<longleftrightarrow> ?rhs"
using True \<open>x \<noteq> 0\<close> \<open>z \<noteq> 0\<close>
by (smt zero_le_mult_iff)
ultimately
show ?thesis
- using `y = 0`
+ using \<open>y = 0\<close>
by auto
qed
qed
next
case False
thus ?thesis
proof (cases "z = y")
case True
thus ?thesis
using assms
unfolding poincare_between_def
by auto
next
case False
let ?x = "of_complex x" and ?y = "of_complex y" and ?z = "of_complex z"
have "?x \<in> unit_disc" "?y \<in> unit_disc" "?z \<in> unit_disc"
using assms
by auto
have distinct: "distinct [?x, ?y, ?z, inversion ?y]"
using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> False \<open>?x \<in> unit_disc\<close> \<open>?y \<in> unit_disc\<close> \<open>?z \<in> unit_disc\<close>
using inversion_noteq_unit_disc[of ?y ?y]
using inversion_noteq_unit_disc[of ?y ?x]
using inversion_noteq_unit_disc[of ?y ?z]
using of_complex_inj[of x y] of_complex_inj[of y z] of_complex_inj[of x z]
by auto
have "cor y * cor x \<noteq> 1"
using assms
by (smt minus_mult_minus mult_less_cancel_left2 mult_less_cancel_right2 of_real_1 of_real_eq_iff of_real_mult)
let ?cr = "cross_ratio ?x ?y ?z (inversion ?y)"
have "Re (to_complex ?cr) = (x - y) * (z*y - 1)/ ((x*y - 1)*(z - y))"
proof-
have " \<And>y x z. \<lbrakk>y \<noteq> x; z \<noteq> x; z \<noteq> y; cor y * cor x \<noteq> 1; x \<noteq> 0; y \<noteq> 0; z \<noteq> 0\<rbrakk> \<Longrightarrow>
(y * y + y * (y * (x * z)) - (y * x + y * (y * (y * z)))) /
(y * y + y * (y * (x * z)) - (y * z + y * (y * (y * x)))) =
(y + y * (x * z) - (x + y * (y * z))) / (y + y * (x * z) - (z + y * (y * x)))"
by (metis (no_types, hide_lams) ab_group_add_class.ab_diff_conv_add_uminus distrib_left mult_divide_mult_cancel_left_if mult_minus_right)
thus ?thesis
using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> False \<open>\<not> (x = 0 \<or> y = 0 \<or> z = 0)\<close>
using \<open>cor y * cor x \<noteq> 1\<close>
unfolding inversion_def
by (transfer, transfer, auto simp add: vec_cnj_def power2_eq_square field_simps split: if_split_asm)
qed
moreover
have "(x*y - 1) < 0"
using assms
by (smt minus_mult_minus mult_less_cancel_right2 zero_less_mult_iff)
moreover
have "(z*y - 1) < 0"
using assms
by (smt minus_mult_minus mult_less_cancel_right2 zero_less_mult_iff)
moreover
have "(x - y) / (z - y) < 0 \<longleftrightarrow> ?rhs"
using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> False \<open>\<not> (x = 0 \<or> y = 0 \<or> z = 0)\<close>
by (smt divide_less_cancel divide_nonneg_nonpos divide_nonneg_pos divide_nonpos_nonneg divide_nonpos_nonpos)
ultimately
have *: "Re (to_complex ?cr) < 0 \<longleftrightarrow> ?rhs"
by (smt algebra_split_simps(24) minus_divide_left zero_less_divide_iff zero_less_mult_iff)
show ?thesis
proof
assume ?lhs
have "is_real (to_complex ?cr)" "Re (to_complex ?cr) < 0"
using \<open>?lhs\<close> distinct
unfolding poincare_between_def Let_def
by auto
thus ?rhs
using *
by simp
next
assume ?rhs
hence "Re (to_complex ?cr) < 0"
using *
by simp
moreover
have "{of_complex (cor x), of_complex (cor y), of_complex (cor z), inversion (of_complex (cor y))} \<subseteq> circline_set x_axis"
using \<open>\<not> (x = 0 \<or> y = 0 \<or> z = 0)\<close> is_real_inversion[of "cor y"]
using inf_or_of_complex[of "inversion ?y"]
by (auto simp del: inversion_of_complex)
hence "is_real (to_complex ?cr)"
using four_points_on_circline_iff_cross_ratio_real[OF distinct]
by auto
ultimately
show ?lhs
using distinct
unfolding poincare_between_def Let_def
by auto
qed
qed
qed
subsubsection \<open>H-betweenness and h-collinearity\<close>
text\<open>For three h-collinear points at least one of the three possible h-betweeness relations must
hold.\<close>
lemma poincare_collinear3_between:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
assumes "poincare_collinear {u, v, w}"
shows "poincare_between u v w \<or> poincare_between u w v \<or> poincare_between v u w" (is "?P' u v w")
proof (cases "u=v")
case True
thus ?thesis
using assms
by auto
next
case False
have "\<forall> w. w \<in> unit_disc \<and> poincare_collinear {u, v, w} \<longrightarrow> ?P' u v w" (is "?P u v")
proof (rule wlog_positive_x_axis[where P="?P"])
fix x
assume x: "is_real x" "0 < Re x" "Re x < 1"
hence "x \<noteq> 0"
using complex.expand[of x 0]
by auto
hence *: "poincare_line 0\<^sub>h (of_complex x) = x_axis"
using x poincare_line_0_real_is_x_axis[of "of_complex x"]
unfolding circline_set_x_axis
by auto
have "of_complex x \<in> unit_disc"
using x
by (auto simp add: cmod_eq_Re)
have "of_complex x \<noteq> 0\<^sub>h"
using \<open>x \<noteq> 0\<close>
by auto
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix w
assume "w \<in> unit_disc"
assume "poincare_collinear {0\<^sub>h, of_complex x, w}"
hence "w \<in> circline_set x_axis"
using * unique_poincare_line[of "0\<^sub>h" "of_complex x"] \<open>of_complex x \<in> unit_disc\<close> \<open>x \<noteq> 0\<close> \<open>of_complex x \<noteq> 0\<^sub>h\<close>
unfolding poincare_collinear_def
by auto
then obtain w' where w': "w = of_complex w'" "is_real w'"
using \<open>w \<in> unit_disc\<close>
using inf_or_of_complex[of w]
unfolding circline_set_x_axis
by auto
hence "-1 < Re w'" "Re w' < 1"
using \<open>w \<in> unit_disc\<close>
by (auto simp add: cmod_eq_Re)
assume 1: "\<not> poincare_between (of_complex x) 0\<^sub>h w"
hence "w \<noteq> 0\<^sub>h" "w' \<noteq> 0"
using w'
unfolding poincare_between_def
by auto
hence "Re w' \<noteq> 0"
using w' complex.expand[of w' 0]
by auto
have "Re w' \<ge> 0"
using 1 poincare_between_x_axis_u0v[of x w'] \<open>Re x > 0\<close> \<open>is_real x\<close> \<open>x \<noteq> 0\<close> \<open>w' \<noteq> 0\<close> w'
using mult_pos_neg
by force
moreover
assume "\<not> poincare_between 0\<^sub>h (of_complex x) w"
hence "Re w' < Re x"
using poincare_between_x_axis_0uv[of "Re x" "Re w'"]
using w' x \<open>-1 < Re w'\<close> \<open>Re w' < 1\<close> \<open>Re w' \<noteq> 0\<close>
by auto
ultimately
show "poincare_between 0\<^sub>h w (of_complex x)"
using poincare_between_x_axis_0uv[of "Re w'" "Re x"]
using w' x \<open>-1 < Re w'\<close> \<open>Re w' < 1\<close> \<open>Re w' \<noteq> 0\<close>
by auto
qed
next
show "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
by fact+
next
fix M u v
assume 1: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v"
assume 2: "?P ?Mu ?Mv"
show "?P u v"
proof safe
fix w
assume "w \<in> unit_disc" "poincare_collinear {u, v, w}" "\<not> poincare_between u v w" "\<not> poincare_between v u w"
thus "poincare_between u w v"
using 1 2[rule_format, of "moebius_pt M w"]
by simp
qed
qed
thus ?thesis
using assms
by simp
qed
lemma poincare_collinear3_iff:
assumes "u \<in> unit_disc" "v \<in> unit_disc" "w \<in> unit_disc"
shows "poincare_collinear {u, v, w} \<longleftrightarrow> poincare_between u v w \<or> poincare_between v u w \<or> poincare_between v w u"
using assms
by (metis poincare_collinear3_between insert_commute poincare_between_poincare_collinear poincare_between_rev)
subsection \<open>Some properties of betweenness\<close>
lemma poincare_between_transitivity:
assumes "a \<in> unit_disc" and "x \<in> unit_disc" and "b \<in> unit_disc" and "y \<in> unit_disc" and
"poincare_between a x b" and "poincare_between a b y"
shows "poincare_between x b y"
proof(cases "a = b")
case True
thus ?thesis
using assms
using poincare_between_sandwich by blast
next
case False
have "\<forall> x. \<forall> y. poincare_between a x b \<and> poincare_between a b y \<and> x \<in> unit_disc
\<and> y \<in> unit_disc \<longrightarrow> poincare_between x b y" (is "?P a b")
proof (rule wlog_positive_x_axis[where P="?P"])
show "a \<in> unit_disc"
using assms by simp
next
show "b \<in> unit_disc"
using assms by simp
next
show "a \<noteq> b"
using False by simp
next
fix M u v
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
"\<forall>x y. poincare_between (moebius_pt M u) x (moebius_pt M v) \<and>
poincare_between (moebius_pt M u) (moebius_pt M v) y \<and>
x \<in> unit_disc \<and> y \<in> unit_disc \<longrightarrow>
poincare_between x (moebius_pt M v) y"
show "\<forall>x y. poincare_between u x v \<and> poincare_between u v y \<and> x \<in> unit_disc \<and> y \<in> unit_disc
\<longrightarrow> poincare_between x v y"
proof safe
fix x y
assume "poincare_between u x v" "poincare_between u v y" " x \<in> unit_disc" "y \<in> unit_disc"
have "poincare_between (moebius_pt M u) (moebius_pt M x) (moebius_pt M v)"
using \<open>poincare_between u x v\<close> \<open>unit_disc_fix M\<close> \<open>x \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close>
by simp
moreover
have "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M y)"
using \<open>poincare_between u v y\<close> \<open>unit_disc_fix M\<close> \<open>y \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close>
by simp
moreover
have "(moebius_pt M x) \<in> unit_disc"
using \<open>unit_disc_fix M\<close> \<open>x \<in> unit_disc\<close> by simp
moreover
have "(moebius_pt M y) \<in> unit_disc"
using \<open>unit_disc_fix M\<close> \<open>y \<in> unit_disc\<close> by simp
ultimately
have "poincare_between (moebius_pt M x) (moebius_pt M v) (moebius_pt M y)"
using * by blast
thus "poincare_between x v y"
using \<open>y \<in> unit_disc\<close> * \<open>x \<in> unit_disc\<close> by simp
qed
next
fix x
assume xx: "is_real x" "0 < Re x" "Re x < 1"
hence "of_complex x \<in> unit_disc"
using cmod_eq_Re by auto
hence "of_complex x \<noteq> \<infinity>\<^sub>h"
by simp
have " of_complex x \<noteq> 0\<^sub>h"
using xx by auto
have "of_complex x \<in> circline_set x_axis"
using xx by simp
show "\<forall>m n. poincare_between 0\<^sub>h m (of_complex x) \<and> poincare_between 0\<^sub>h (of_complex x) n \<and>
m \<in> unit_disc \<and> n \<in> unit_disc \<longrightarrow> poincare_between m (of_complex x) n"
proof safe
fix m n
assume **: "poincare_between 0\<^sub>h m (of_complex x)" "poincare_between 0\<^sub>h (of_complex x) n"
"m \<in> unit_disc" " n \<in> unit_disc"
show "poincare_between m (of_complex x) n"
proof(cases "m = 0\<^sub>h")
case True
thus ?thesis
using ** by auto
next
case False
hence "m \<in> circline_set x_axis"
using poincare_between_poincare_line_uzv[of "0\<^sub>h" "of_complex x" m]
using poincare_line_0_real_is_x_axis[of "of_complex x"]
using \<open>of_complex x \<in> unit_disc\<close> \<open>of_complex x \<noteq> \<infinity>\<^sub>h\<close> \<open>of_complex x \<noteq> 0\<^sub>h\<close>
using \<open>of_complex x \<in> circline_set x_axis\<close> \<open>m \<in> unit_disc\<close> **(1)
by simp
then obtain m' where "m = of_complex m'" "is_real m'"
using inf_or_of_complex[of m] \<open>m \<in> unit_disc\<close>
unfolding circline_set_x_axis
by auto
hence "Re m' \<le> Re x"
using \<open>poincare_between 0\<^sub>h m (of_complex x)\<close> xx \<open>of_complex x \<noteq> 0\<^sub>h\<close>
using False ** \<open>of_complex x \<in> unit_disc\<close>
using cmod_Re_le_iff poincare_between_0uv by auto
have "n \<noteq> 0\<^sub>h"
using **(2, 4) \<open>of_complex x \<noteq> 0\<^sub>h\<close> \<open>of_complex x \<in> unit_disc\<close>
using poincare_between_sandwich by fastforce
have "n \<in> circline_set x_axis"
using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex x" n]
using poincare_line_0_real_is_x_axis[of "of_complex x"]
using \<open>of_complex x \<in> unit_disc\<close> \<open>of_complex x \<noteq> \<infinity>\<^sub>h\<close> \<open>of_complex x \<noteq> 0\<^sub>h\<close>
using \<open>of_complex x \<in> circline_set x_axis\<close> \<open>n \<in> unit_disc\<close> **(2)
by simp
then obtain n' where "n = of_complex n'" "is_real n'"
using inf_or_of_complex[of n] \<open>n \<in> unit_disc\<close>
unfolding circline_set_x_axis
by auto
hence "Re x \<le> Re n'"
using \<open>poincare_between 0\<^sub>h (of_complex x) n\<close> xx \<open>of_complex x \<noteq> 0\<^sub>h\<close>
using False ** \<open>of_complex x \<in> unit_disc\<close> \<open>n \<noteq> 0\<^sub>h\<close>
using cmod_Re_le_iff poincare_between_0uv
by (metis Re_complex_of_real arg_0_iff rcis_cmod_arg rcis_zero_arg to_complex_of_complex)
have "poincare_between (of_complex m') (of_complex x) (of_complex n')"
using \<open>Re x \<le> Re n'\<close> \<open>Re m' \<le> Re x\<close>
using poincare_between_x_axis_uvw[of "Re m'" "Re x" "Re n'"]
using \<open>is_real n'\<close> \<open>is_real m'\<close> \<open>n \<in> unit_disc\<close> \<open>n = of_complex n'\<close>
using xx \<open>m = of_complex m'\<close> \<open>m \<in> unit_disc\<close>
by (smt complex_of_real_Re norm_of_real poincare_between_def unit_disc_iff_cmod_lt_1)
thus ?thesis
using \<open>n = of_complex n'\<close> \<open>m = of_complex m'\<close>
by auto
qed
qed
qed
thus ?thesis
using assms
by blast
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Poincare between - sum distances\<close>
(* ------------------------------------------------------------------ *)
text\<open>Another possible definition of the h-betweenness relation is given in terms of h-distances
between pairs of points. We prove it as a characterization equivalent to our cross-ratio based
definition.\<close>
lemma poincare_between_sum_distances_x_axis_u0v:
assumes "of_complex u' \<in> unit_disc" "of_complex v' \<in> unit_disc"
assumes "is_real u'" "u' \<noteq> 0" "v' \<noteq> 0"
shows "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') = poincare_distance (of_complex u') (of_complex v') \<longleftrightarrow>
is_real v' \<and> Re u' * Re v' < 0" (is "?P u' v' \<longleftrightarrow> ?Q u' v'")
proof-
have "Re u' \<noteq> 0"
using \<open>is_real u'\<close> \<open>u' \<noteq> 0\<close>
using complex_eq_if_Re_eq
by simp
let ?u = "cmod u'" and ?v = "cmod v'" and ?uv = "cmod (u' - v')"
have disc: "?u\<^sup>2 < 1" "?v\<^sup>2 < 1"
using unit_disc_cmod_square_lt_1[OF assms(1)]
using unit_disc_cmod_square_lt_1[OF assms(2)]
by auto
have "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') =
arcosh (((1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v) / ((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2)))" (is "_ = arcosh ?r1")
using poincare_distance_formula_zero_sum[OF assms(1-2)]
by (simp add: Let_def)
moreover
have "poincare_distance (of_complex u') (of_complex v') =
arcosh (((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) + 2 * ?uv\<^sup>2) / ((1 - ?u\<^sup>2) * (1 - ?v\<^sup>2)))" (is "_ = arcosh ?r2")
using disc
using poincare_distance_formula[OF assms(1-2)]
by (subst add_divide_distrib) simp
moreover
have "arcosh ?r1 = arcosh ?r2 \<longleftrightarrow> ?Q u' v'"
proof
assume "arcosh ?r1 = arcosh ?r2"
hence "?r1 = ?r2"
proof (subst (asm) arcosh_eq_iff)
show "?r1 \<ge> 1"
proof-
have "(1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) \<le> (1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v"
by (simp add: field_simps)
thus ?thesis
using disc
by simp
qed
next
show "?r2 \<ge> 1"
using disc
by simp
qed
hence "(1 + ?u\<^sup>2) * (1 + ?v\<^sup>2) + 4 * ?u * ?v = (1 - ?u\<^sup>2) * (1 - ?v\<^sup>2) + 2 * ?uv\<^sup>2"
using disc
by auto
hence "(cmod (u' - v'))\<^sup>2 = (cmod u' + cmod v')\<^sup>2"
by (simp add: field_simps power2_eq_square)
hence *: "Re u' * Re v' + \<bar>Re u'\<bar> * sqrt ((Im v')\<^sup>2 + (Re v')\<^sup>2) = 0"
using \<open>is_real u'\<close>
unfolding cmod_power2 cmod_def
by (simp add: field_simps) (simp add: power2_eq_square field_simps)
hence "sqrt ((Im v')\<^sup>2 + (Re v')\<^sup>2) = \<bar>Re v'\<bar>"
using \<open>Re u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close>
by (smt complex_neq_0 mult.commute mult_cancel_right mult_minus_left real_sqrt_gt_0_iff)
hence "Im v' = 0"
by (smt Im_eq_0 norm_complex_def)
moreover
hence "Re u' * Re v' = - \<bar>Re u'\<bar> * \<bar>Re v'\<bar>"
using *
by simp
hence "Re u' * Re v' < 0"
using \<open>Re u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close>
by (simp add: \<open>is_real v'\<close> complex_eq_if_Re_eq)
ultimately
show "?Q u' v'"
by simp
next
assume "?Q u' v'"
hence "is_real v'" "Re u' * Re v' < 0"
by auto
have "?r1 = ?r2"
proof (cases "Re u' > 0")
case True
hence "Re v' < 0"
using \<open>Re u' * Re v' < 0\<close>
by (smt zero_le_mult_iff)
show ?thesis
using disc \<open>is_real u'\<close> \<open>is_real v'\<close>
using \<open>Re u' > 0\<close> \<open>Re v' < 0\<close>
unfolding cmod_power2 cmod_def
by simp (simp add: power2_eq_square field_simps)
next
case False
hence "Re u' < 0"
using \<open>Re u' \<noteq> 0\<close>
by simp
hence "Re v' > 0"
using \<open>Re u' * Re v' < 0\<close>
by (smt zero_le_mult_iff)
show ?thesis
using disc \<open>is_real u'\<close> \<open>is_real v'\<close>
using \<open>Re u' < 0\<close> \<open>Re v' > 0\<close>
unfolding cmod_power2 cmod_def
by simp (simp add: power2_eq_square field_simps)
qed
thus "arcosh ?r1 = arcosh ?r2"
by metis
qed
ultimately
show ?thesis
by simp
qed
text\<open>
Different proof of the previous theorem relying on the cross-ratio definition, and not the distance formula.
We suppose that this could be also used to prove the triangle inequality.
\<close>
lemma poincare_between_sum_distances_x_axis_u0v_different_proof:
assumes "of_complex u' \<in> unit_disc" "of_complex v' \<in> unit_disc"
assumes "is_real u'" "u' \<noteq> 0" "v' \<noteq> 0" (* additional condition *) "is_real v'"
shows "poincare_distance (of_complex u') 0\<^sub>h + poincare_distance 0\<^sub>h (of_complex v') = poincare_distance (of_complex u') (of_complex v') \<longleftrightarrow>
Re u' * Re v' < 0" (is "?P u' v' \<longleftrightarrow> ?Q u' v'")
proof-
have "-1 < Re u'" "Re u' < 1" "Re u' \<noteq> 0"
using assms
by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq)
have "-1 < Re v'" "Re v' < 1" "Re v' \<noteq> 0"
using assms
by (auto simp add: cmod_eq_Re complex_eq_if_Re_eq)
have "\<bar>ln (Re ((1 - u') / (1 + u')))\<bar> + \<bar>ln (Re ((1 - v') / (1 + v')))\<bar> =
\<bar>ln (Re ((1 + u') * (1 - v') / ((1 - u') * (1 + v'))))\<bar> \<longleftrightarrow> Re u' * Re v' < 0" (is "\<bar>ln ?a1\<bar> + \<bar>ln ?a2\<bar> = \<bar>ln ?a3\<bar> \<longleftrightarrow> _")
proof-
have 1: "0 < ?a1" "ln ?a1 > 0 \<longleftrightarrow> Re u' < 0"
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using complex_is_Real_iff
by auto
have 2: "0 < ?a2" "ln ?a2 > 0 \<longleftrightarrow> Re v' < 0"
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
by auto
have 3: "0 < ?a3" "ln ?a3 > 0 \<longleftrightarrow> Re v' < Re u'"
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
by auto (simp add: field_simps)+
show ?thesis
proof
assume *: "Re u' * Re v' < 0"
show "\<bar>ln ?a1\<bar> + \<bar>ln ?a2\<bar> = \<bar>ln ?a3\<bar>"
proof (cases "Re u' > 0")
case True
hence "Re v' < 0"
using *
by (smt mult_nonneg_nonneg)
show ?thesis
using 1 2 3 \<open>Re u' > 0\<close> \<open>Re v' < 0\<close>
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
using ln_div ln_mult
by simp
next
case False
hence "Re v' > 0" "Re u' < 0"
using *
by (smt zero_le_mult_iff)+
show ?thesis
using 1 2 3 \<open>Re u' < 0\<close> \<open>Re v' > 0\<close>
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
using ln_div ln_mult
by simp
qed
next
assume *: "\<bar>ln ?a1\<bar> + \<bar>ln ?a2\<bar> = \<bar>ln ?a3\<bar>"
{
assume "Re u' > 0" "Re v' > 0"
hence False
using * 1 2 3
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
using ln_mult ln_div
by (cases "Re v' < Re u'") auto
}
moreover
{
assume "Re u' < 0" "Re v' < 0"
hence False
using * 1 2 3
using \<open>Re u' < 1\<close> \<open>Re u' > -1\<close> \<open>is_real u'\<close>
using \<open>Re v' < 1\<close> \<open>Re v' > -1\<close> \<open>is_real v'\<close>
using complex_is_Real_iff
using ln_mult ln_div
by (cases "Re v' < Re u'") auto
}
ultimately
show "Re u' * Re v' < 0"
using \<open>Re u' \<noteq> 0\<close> \<open>Re v' \<noteq> 0\<close>
by (smt divisors_zero mult_le_0_iff)
qed
qed
thus ?thesis
using assms
apply (subst poincare_distance_sym, simp, simp)
apply (subst poincare_distance_zero_x_axis, simp, simp add: circline_set_x_axis)
apply (subst poincare_distance_zero_x_axis, simp, simp add: circline_set_x_axis)
apply (subst poincare_distance_x_axis_x_axis, simp, simp, simp add: circline_set_x_axis, simp add: circline_set_x_axis)
apply simp
done
qed
lemma poincare_between_sum_distances:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_between u v w \<longleftrightarrow>
poincare_distance u v + poincare_distance v w = poincare_distance u w" (is "?P' u v w")
proof (cases "u = v")
case True
thus ?thesis
using assms
by simp
next
case False
have "\<forall> w. w \<in> unit_disc \<longrightarrow> (poincare_between u v w \<longleftrightarrow> poincare_distance u v + poincare_distance v w = poincare_distance u w)" (is "?P u v")
proof (rule wlog_positive_x_axis)
fix x
assume "is_real x" "0 < Re x" "Re x < 1"
have "of_complex x \<in> circline_set x_axis"
using \<open>is_real x\<close>
by (auto simp add: circline_set_x_axis)
have "of_complex x \<in> unit_disc"
using \<open>is_real x\<close> \<open>0 < Re x\<close> \<open>Re x < 1\<close>
by (simp add: cmod_eq_Re)
have "x \<noteq> 0"
using \<open>is_real x\<close> \<open>Re x > 0\<close>
by auto
show "?P (of_complex x) 0\<^sub>h"
proof (rule allI, rule impI)
fix w
assume "w \<in> unit_disc"
then obtain w' where "w = of_complex w'"
using inf_or_of_complex[of w]
by auto
show "?P' (of_complex x) 0\<^sub>h w"
proof (cases "w = 0\<^sub>h")
case True
thus ?thesis
by simp
next
case False
hence "w' \<noteq> 0"
using \<open>w = of_complex w'\<close>
by auto
show ?thesis
using \<open>is_real x\<close> \<open>x \<noteq> 0\<close> \<open>w = of_complex w'\<close> \<open>w' \<noteq> 0\<close>
using \<open>of_complex x \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>
apply simp
apply (subst poincare_between_x_axis_u0v, simp_all)
apply (subst poincare_between_sum_distances_x_axis_u0v, simp_all)
done
qed
qed
next
show "v \<in> unit_disc" "u \<in> unit_disc"
using assms
by auto
next
show "v \<noteq> u"
using \<open>u \<noteq> v\<close>
by simp
next
fix M u v
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v" and
**: "?P (moebius_pt M v) (moebius_pt M u)"
show "?P v u"
proof (rule allI, rule impI)
fix w
assume "w \<in> unit_disc"
hence "moebius_pt M w \<in> unit_disc"
using \<open>unit_disc_fix M\<close>
by auto
thus "?P' v u w"
using \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close> \<open>unit_disc_fix M\<close>
using **[rule_format, of "moebius_pt M w"]
by auto
qed
qed
thus ?thesis
using assms
by simp
qed
subsection \<open>Some more properties of h-betweenness.\<close>
text \<open>Some lemmas proved earlier are proved almost directly using the sum of distances characterization.\<close>
lemma unit_disc_fix_moebius_preserve_poincare_between':
assumes "unit_disc_fix M" and "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_between (moebius_pt M u) (moebius_pt M v) (moebius_pt M w) \<longleftrightarrow>
poincare_between u v w"
using assms
using poincare_between_sum_distances
by simp
lemma conjugate_preserve_poincare_between':
assumes "u \<in> unit_disc" "v \<in> unit_disc" "w \<in> unit_disc"
shows "poincare_between (conjugate u) (conjugate v) (conjugate w) \<longleftrightarrow> poincare_between u v w"
using assms
using poincare_between_sum_distances
by simp
text \<open>There is a unique point on a ray on the given distance from the given starting point\<close>
lemma unique_poincare_distance_on_ray:
assumes "d \<ge> 0" "u \<noteq> v" "u \<in> unit_disc" "v \<in> unit_disc"
assumes "y \<in> unit_disc" "poincare_distance u y = d" "poincare_between u v y"
assumes "z \<in> unit_disc" "poincare_distance u z = d" "poincare_between u v z"
shows "y = z"
proof-
have "\<forall> d y z. d \<ge> 0 \<and>
y \<in> unit_disc \<and> poincare_distance u y = d \<and> poincare_between u v y \<and>
z \<in> unit_disc \<and> poincare_distance u z = d \<and> poincare_between u v z \<longrightarrow> y = z" (is "?P u v")
proof (rule wlog_positive_x_axis[where P="?P"])
fix x
assume x: "is_real x" "0 < Re x" "Re x < 1"
hence "x \<noteq> 0"
using complex.expand[of x 0]
by auto
hence *: "poincare_line 0\<^sub>h (of_complex x) = x_axis"
using x poincare_line_0_real_is_x_axis[of "of_complex x"]
unfolding circline_set_x_axis
by auto
have "of_complex x \<in> unit_disc"
using x
by (auto simp add: cmod_eq_Re)
have "arg x = 0"
using x
using arg_0_iff by blast
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix y z
assume "y \<in> unit_disc" "z \<in> unit_disc"
then obtain y' z' where yz: "y = of_complex y'" "z = of_complex z'"
using inf_or_of_complex[of y] inf_or_of_complex[of z]
by auto
assume betw: "poincare_between 0\<^sub>h (of_complex x) y" "poincare_between 0\<^sub>h (of_complex x) z"
hence "y \<noteq> 0\<^sub>h" "z \<noteq> 0\<^sub>h"
using \<open>x \<noteq> 0\<close> \<open>of_complex x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
using poincare_between_sandwich[of "0\<^sub>h" "of_complex x"]
using of_complex_zero_iff[of x]
by force+
hence "arg y' = 0" "cmod y' \<ge> cmod x" "arg z' = 0" "cmod z' \<ge> cmod x"
using poincare_between_0uv[of "of_complex x" y] poincare_between_0uv[of "of_complex x" z]
using \<open>of_complex x \<in> unit_disc\<close> \<open>x \<noteq> 0\<close> \<open>arg x = 0\<close> \<open>y \<in> unit_disc\<close> \<open>z \<in> unit_disc\<close> betw yz
by (simp_all add: Let_def)
hence *: "is_real y'" "is_real z'" "Re y' > 0" "Re z' > 0"
using arg_0_iff[of y'] arg_0_iff[of z'] x \<open>y \<noteq> 0\<^sub>h\<close> \<open>z \<noteq> 0\<^sub>h\<close> yz
by auto
assume "poincare_distance 0\<^sub>h z = poincare_distance 0\<^sub>h y" "0 \<le> poincare_distance 0\<^sub>h y"
thus "y = z"
using * yz \<open>y \<in> unit_disc\<close> \<open>z \<in> unit_disc\<close>
using unique_x_axis_poincare_distance_positive[of "poincare_distance 0\<^sub>h y"]
by (auto simp add: cmod_eq_Re unit_disc_to_complex_inj)
qed
next
show "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
by fact+
next
fix M u v
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assume **: "?P (moebius_pt M u) (moebius_pt M v)"
show "?P u v"
proof safe
fix d y z
assume ***: "0 \<le> poincare_distance u y"
"y \<in> unit_disc" "poincare_between u v y"
"z \<in> unit_disc" "poincare_between u v z"
"poincare_distance u z = poincare_distance u y"
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v" and ?My = "moebius_pt M y" and ?Mz = "moebius_pt M z"
have "?Mu \<in> unit_disc" "?Mv \<in> unit_disc" "?My \<in> unit_disc" "?Mz \<in> unit_disc"
using \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>z \<in> unit_disc\<close>
using \<open>unit_disc_fix M\<close>
by auto
hence "?My = ?Mz"
using * ***
using **[rule_format, of "poincare_distance ?Mu ?My" ?My ?Mz]
by simp
thus "y = z"
using bij_moebius_pt[of M]
unfolding bij_def inj_on_def
by blast
qed
qed
thus ?thesis
using assms
by auto
qed
end
\ No newline at end of file
diff --git a/thys/Poincare_Disc/Poincare_Circles.thy b/thys/Poincare_Disc/Poincare_Circles.thy
--- a/thys/Poincare_Disc/Poincare_Circles.thy
+++ b/thys/Poincare_Disc/Poincare_Circles.thy
@@ -1,618 +1,618 @@
theory Poincare_Circles
imports Poincare_Distance
begin
(* -------------------------------------------------------------------------- *)
section\<open>H-circles in the Poincar\'e model\<close>
(* -------------------------------------------------------------------------- *)
text\<open>Circles consist of points that are at the same distance from the center.\<close>
definition poincare_circle :: "complex_homo \<Rightarrow> real \<Rightarrow> complex_homo set" where
"poincare_circle z r = {z'. z' \<in> unit_disc \<and> poincare_distance z z' = r}"
text\<open>Each h-circle in the Poincar\'e model is represented by an Euclidean circle in the model ---
the center and radius of that euclidean circle are determined by the following formulas.\<close>
definition poincare_circle_euclidean :: "complex_homo \<Rightarrow> real \<Rightarrow> euclidean_circle" where
"poincare_circle_euclidean z r =
(let R = (cosh r - 1) / 2;
z' = to_complex z;
cz = 1 - (cmod z')\<^sup>2;
k = cz * R + 1
in (z' / k, cz * sqrt(R * (R + 1)) / k))"
text\<open>That Euclidean circle has a positive radius and is always fully within the disc.\<close>
lemma poincare_circle_in_disc:
assumes "r > 0" and "z \<in> unit_disc" and "(ze, re) = poincare_circle_euclidean z r"
shows "cmod ze < 1" "re > 0" "\<forall> x \<in> circle ze re. cmod x < 1"
proof-
let ?R = "(cosh r - 1) / 2"
let ?z' = "to_complex z"
let ?cz = "1 - (cmod ?z')\<^sup>2"
let ?k = "?cz * ?R + 1"
let ?ze = "?z' / ?k"
let ?re = "?cz * sqrt(?R * (?R + 1)) / ?k"
from \<open>z \<in> unit_disc\<close>
obtain z' where z': "z = of_complex z'"
using inf_or_of_complex[of z]
by auto
hence "z' = ?z'"
by simp
obtain cz where cz: "cz = (1 - (cmod z')\<^sup>2)"
by auto
have "cz > 0" "cz \<le> 1"
using \<open>z \<in> unit_disc\<close> z' cz
using unit_disc_cmod_square_lt_1
by fastforce+
obtain R where R: "R = ?R"
by blast
have "R > 0"
using cosh_gt_1[of r] \<open>r > 0\<close>
by (subst R) simp
obtain k where k: "k = cz * R + 1"
by auto
have "k > 1"
using k \<open>R > 0\<close> \<open>cz > 0\<close>
by simp
hence "cmod k = k"
by simp
let ?RR = "cz * sqrt(R * (R + 1)) / k"
have "cmod z' + cz * sqrt(R * (R + 1)) < k"
proof-
have "((R+1)-R)\<^sup>2 > 0"
by simp
hence "(R+1)\<^sup>2 - 2*R*(R+1) + R\<^sup>2 > 0"
unfolding power2_diff
by (simp add: field_simps)
hence "(R+1)\<^sup>2 + 2*R*(R+1) + R\<^sup>2 - 4*R*(R+1) > 0"
by simp
hence "(2*R+1)\<^sup>2 / 4 > R*(R+1)"
using power2_sum[of "R+1" R]
by (simp add: field_simps)
hence "sqrt(R*(R+1)) < (2*R+1) / 2"
using \<open>R > 0\<close>
by (smt arith_geo_mean_sqrt power_divide real_sqrt_four real_sqrt_pow2 zero_le_mult_iff)
hence "sqrt(R*(R+1)) - R < 1/2"
by (simp add: field_simps)
hence "(1 + (cmod z')) * (sqrt(R*(R+1)) - R) < (1 + (cmod z')) * (1 / 2)"
by (subst mult_strict_left_mono, simp, smt norm_not_less_zero, simp)
also have "... < 1"
using \<open>z \<in> unit_disc\<close> z'
by auto
finally have "(1 - cmod z') * ((1 + cmod z') * (sqrt(R*(R+1)) - R)) < (1 - cmod z') * 1"
using \<open>z \<in> unit_disc\<close> z'
by (subst mult_strict_left_mono, simp_all)
hence "cz * (sqrt (R*(R+1)) - R) < 1 - cmod z'"
using square_diff_square_factored[of 1 "cmod z'"]
by (subst cz, subst (asm) mult.assoc[symmetric], simp add: power2_eq_square field_simps)
hence "cmod z' + cz * sqrt(R*(R+1)) < 1 + R * cz"
by (simp add: field_simps)
thus ?thesis
using k
by (simp add: field_simps)
qed
hence "cmod z' / k + cz * sqrt(R * (R + 1)) / k < 1"
using \<open>k > 1\<close>
unfolding add_divide_distrib[symmetric]
by simp
hence "cmod (z' / k) + cz * sqrt(R * (R + 1)) / k < 1"
using \<open>k > 1\<close>
by simp
hence "cmod ?ze + ?re < 1"
using k cz \<open>R = ?R\<close> z'
by simp
moreover
have "cz * sqrt(R * (R + 1)) / k > 0"
using \<open>cz > 0\<close> \<open>R > 0\<close> \<open>k > 1\<close>
by auto
hence "?re > 0"
using k cz \<open>R = ?R\<close> z'
by simp
moreover
have "cmod ?ze < 1"
using \<open>cmod ?ze + ?re < 1\<close> \<open>?re > 0\<close>
by simp
moreover
have "ze = ?ze" "re = ?re"
using \<open>(ze, re) = poincare_circle_euclidean z r\<close>
unfolding poincare_circle_euclidean_def Let_def
by simp_all
moreover
have "\<forall> x \<in> circle ze re. cmod x \<le> cmod ze + re"
using norm_triangle_ineq2[of _ ze]
unfolding circle_def
by (smt mem_Collect_eq)
ultimately
show "cmod ze < 1" "re > 0" "\<forall> x \<in> circle ze re. cmod x < 1"
by auto
qed
text\<open>The connection between the points on the h-circle and its corresponding Euclidean circle.\<close>
lemma poincare_circle_is_euclidean_circle:
assumes "z \<in> unit_disc" and "r > 0"
shows "let (Ze, Re) = poincare_circle_euclidean z r
in of_complex ` (circle Ze Re) = poincare_circle z r"
proof-
{
fix x
let ?z = "to_complex z"
from assms obtain z' where z': "z = of_complex z'" "cmod z' < 1"
using inf_or_of_complex[of z]
by auto
have *: "\<And> x. cmod x < 1 \<Longrightarrow> 1 - (cmod x)\<^sup>2 > 0"
by (metis less_iff_diff_less_0 minus_diff_eq mult.left_neutral neg_less_0_iff_less norm_mult_less norm_power power2_eq_square)
let ?R = "(cosh r - 1) / 2"
obtain R where R: "R = ?R"
by blast
let ?cx = "1 - (cmod x)\<^sup>2" and ?cz = "1 - (cmod z')\<^sup>2" and ?czx = "(cmod (z' - x))\<^sup>2"
let ?k = "1 + R * ?cz"
obtain k where k: "k = ?k"
by blast
have "R > 0"
using R cosh_gt_1[OF \<open>r > 0\<close>]
by simp
hence "k > 1"
using assms z' k *[of z']
by auto
hence **: "cor k \<noteq> 0"
by (smt of_real_eq_0_iff)
have "of_complex x \<in> poincare_circle z r \<longleftrightarrow> cmod x < 1 \<and> poincare_distance z (of_complex x) = r"
unfolding poincare_circle_def
by auto
also have "... \<longleftrightarrow> cmod x < 1 \<and> poincare_distance_formula' ?z x = cosh r"
using poincare_distance_formula[of z "of_complex x"] cosh_dist[of z "of_complex x"]
unfolding poincare_distance_formula_def
using assms
using arcosh_cosh_real
by auto
also have "... \<longleftrightarrow> cmod x < 1 \<and> ?czx / (?cz * ?cx) = ?R"
using z'
by (simp add: field_simps)
also have "... \<longleftrightarrow> cmod x < 1 \<and> ?czx = ?R * ?cx * ?cz"
using assms z' *[of z'] *[of x]
using nonzero_divide_eq_eq[of "(1 - (cmod x)\<^sup>2) * (1 - (cmod z')\<^sup>2)" "(cmod (z' - x))\<^sup>2" ?R]
by (auto, simp add: field_simps)
also have "... \<longleftrightarrow> cmod x < 1 \<and> (z' - x) * (cnj z' - cnj x) = R * ?cz * (1 - x * cnj x)" (is "_ \<longleftrightarrow> _ \<and> ?l = ?r")
proof-
let ?l = "(z' - x) * (cnj z' - cnj x)" and ?r = "R * (1 - Re (z' * cnj z')) * (1 - x * cnj x)"
have "is_real ?l"
using eq_cnj_iff_real[of "?l"]
by simp
moreover
have "is_real ?r"
using eq_cnj_iff_real[of "1 - x * cnj x"]
using Im_complex_of_real[of "R * (1 - Re (z' * cnj z'))"]
by simp
ultimately
show ?thesis
apply (subst R[symmetric])
apply (subst cmod_square)+
apply (subst complex_eq_if_Re_eq, simp_all add: field_simps)
done
qed
also have "... \<longleftrightarrow> cmod x < 1 \<and> z' * cnj z' - x * cnj z' - cnj x * z' + x * cnj x = R * ?cz - R * ?cz * x * cnj x"
unfolding right_diff_distrib left_diff_distrib
by (simp add: field_simps)
also have "... \<longleftrightarrow> cmod x < 1 \<and> k * (x * cnj x) - x * cnj z' - cnj x * z' + z' * cnj z' = R * ?cz" (is "_ \<longleftrightarrow> _ \<and> ?lhs = ?rhs")
by (subst k) (auto simp add: field_simps)
also have "... \<longleftrightarrow> cmod x < 1 \<and> (k * x * cnj x - x * cnj z' - cnj x * z' + z' * cnj z') / k = (R * ?cz) / k"
using **
by (auto simp add: Groups.mult_ac(1))
also have "... \<longleftrightarrow> cmod x < 1 \<and> x * cnj x - x * cnj z' / k - cnj x * z' / k + z' * cnj z' / k = (R * ?cz) / k"
using **
unfolding add_divide_distrib diff_divide_distrib
by auto
also have "... \<longleftrightarrow> cmod x < 1 \<and> (x - z'/k) * cnj(x - z'/k) = (R * ?cz) / k + (z' / k) * cnj(z' / k) - z' * cnj z' / k"
by (auto simp add: field_simps diff_divide_distrib)
also have "... \<longleftrightarrow> cmod x < 1 \<and> (cmod (x - z'/k))\<^sup>2 = (R * ?cz) / k + (cmod z')\<^sup>2 / k\<^sup>2 - (cmod z')\<^sup>2 / k"
apply (subst complex_mult_cnj_cmod)+
apply (subst complex_eq_if_Re_eq)
apply (simp_all add: power_divide)
done
also have "... \<longleftrightarrow> cmod x < 1 \<and> (cmod (x - z'/k))\<^sup>2 = (R * ?cz * k + (cmod z')\<^sup>2 - (cmod z')\<^sup>2 * k) / k\<^sup>2"
using **
unfolding add_divide_distrib diff_divide_distrib
by (simp add: power2_eq_square)
also have "... \<longleftrightarrow> cmod x < 1 \<and> (cmod (x - z'/k))\<^sup>2 = ?cz\<^sup>2 * R * (R + 1) / k\<^sup>2" (is "_ \<longleftrightarrow> _ \<and> ?a\<^sup>2 = ?b")
proof-
have *: "R * (1 - (cmod z')\<^sup>2) * k + (cmod z')\<^sup>2 - (cmod z')\<^sup>2 * k = (1 - (cmod z')\<^sup>2)\<^sup>2 * R * (R + 1)"
by (subst k)+ (simp add: field_simps power2_diff)
thus ?thesis
by (subst *, simp)
qed
also have "... \<longleftrightarrow> cmod x < 1 \<and> cmod (x - z'/k) = ?cz * sqrt (R * (R + 1)) / k"
using \<open>R > 0\<close> *[of z'] ** \<open>k > 1\<close> \<open>z \<in> unit_disc\<close> z'
using real_sqrt_unique[of ?a ?b, symmetric]
by (auto simp add: real_sqrt_divide real_sqrt_mult power_divide power_mult_distrib)
finally
have "of_complex x \<in> poincare_circle z r \<longleftrightarrow> cmod x < 1 \<and> x \<in> circle (z'/k) (?cz * sqrt(R * (R+1)) / k)"
unfolding circle_def z' k R
by simp
hence "of_complex x \<in> poincare_circle z r \<longleftrightarrow> (let (Ze, Re) = poincare_circle_euclidean z r in cmod x < 1 \<and> x \<in> circle Ze Re)"
unfolding poincare_circle_euclidean_def Let_def circle_def
using z' R k
by (simp add: field_simps)
hence "of_complex x \<in> poincare_circle z r \<longleftrightarrow> (let (Ze, Re) = poincare_circle_euclidean z r in x \<in> circle Ze Re)"
using poincare_circle_in_disc[OF \<open>r > 0\<close> \<open>z \<in> unit_disc\<close>]
by auto
} note * = this
show ?thesis
unfolding Let_def
proof safe
fix Ze Re x
assume "poincare_circle_euclidean z r = (Ze, Re)" "x \<in> circle Ze Re"
thus "of_complex x \<in> poincare_circle z r"
using *[of x]
by simp
next
fix Ze Re x
assume **: "poincare_circle_euclidean z r = (Ze, Re)" "x \<in> poincare_circle z r"
then obtain x' where x': "x = of_complex x'"
unfolding poincare_circle_def
using inf_or_of_complex[of x]
by auto
hence "x' \<in> circle Ze Re"
using *[of x'] **
by simp
thus "x \<in> of_complex ` circle Ze Re"
using x'
by auto
qed
qed
subsection \<open>Intersection of circles in special positions\<close>
text \<open>Two h-circles centered at the x-axis intersect at mutually conjugate points\<close>
lemma intersect_poincare_circles_x_axis:
assumes z: "is_real z1" and "is_real z2" and "r1 > 0" and "r2 > 0" and
"-1 < Re z1" and "Re z1 < 1" and "-1 < Re z2" and "Re z2 < 1" and
"z1 \<noteq> z2"
assumes x1: "x1 \<in> poincare_circle (of_complex z1) r1 \<inter> poincare_circle (of_complex z2) r2" and
x2: "x2 \<in> poincare_circle (of_complex z1) r1 \<inter> poincare_circle (of_complex z2) r2" and
"x1 \<noteq> x2"
shows "x1 = conjugate x2"
proof-
have in_disc: "of_complex z1 \<in> unit_disc" "of_complex z2 \<in> unit_disc"
using assms
by (auto simp add: cmod_eq_Re)
obtain x1' x2' where x': "x1 = of_complex x1'" "x2 = of_complex x2'"
using x1 x2
using inf_or_of_complex[of x1] inf_or_of_complex[of x2]
unfolding poincare_circle_def
by auto
obtain Ze1 Re1 where 1: "(Ze1, Re1) = poincare_circle_euclidean (of_complex z1) r1"
by (metis poincare_circle_euclidean_def)
obtain Ze2 Re2 where 2: "(Ze2, Re2) = poincare_circle_euclidean (of_complex z2) r2"
by (metis poincare_circle_euclidean_def)
have circle: "x1' \<in> circle Ze1 Re1 \<inter> circle Ze2 Re2" "x2' \<in> circle Ze1 Re1 \<inter> circle Ze2 Re2"
using poincare_circle_is_euclidean_circle[of "of_complex z1" r1]
using poincare_circle_is_euclidean_circle[of "of_complex z2" r2]
using assms 1 2 \<open>of_complex z1 \<in> unit_disc\<close> \<open>of_complex z2 \<in> unit_disc\<close> x'
by auto (metis image_iff of_complex_inj)+
have "is_real Ze1" "is_real Ze2"
using 1 2 \<open>is_real z1\<close> \<open>is_real z2\<close>
by (simp_all add: poincare_circle_euclidean_def Let_def)
have "Re1 > 0" "Re2 > 0"
using 1 2 in_disc \<open>r1 > 0\<close> \<open>r2 > 0\<close>
using poincare_circle_in_disc(2)[of r1 "of_complex z1" Ze1 Re1]
using poincare_circle_in_disc(2)[of r2 "of_complex z2" Ze2 Re2]
by auto
have "Ze1 \<noteq> Ze2"
proof (rule ccontr)
assume "\<not> ?thesis"
hence eq: "Ze1 = Ze2" "Re1 = Re2"
using circle(1)
unfolding circle_def
by auto
let ?A = "Ze1 - Re1" and ?B = "Ze1 + Re1"
have "?A \<in> circle Ze1 Re1" "?B \<in> circle Ze1 Re1"
using \<open>Re1 > 0\<close>
unfolding circle_def
by simp_all
hence "of_complex ?A \<in> poincare_circle (of_complex z1) r1" "of_complex ?B \<in> poincare_circle (of_complex z1) r1"
"of_complex ?A \<in> poincare_circle (of_complex z2) r2" "of_complex ?B \<in> poincare_circle (of_complex z2) r2"
using eq
using poincare_circle_is_euclidean_circle[OF \<open>of_complex z1 \<in> unit_disc\<close> \<open>r1 > 0\<close>]
using poincare_circle_is_euclidean_circle[OF \<open>of_complex z2 \<in> unit_disc\<close> \<open>r2 > 0\<close>]
using 1 2
by auto blast+
hence "poincare_distance (of_complex z1) (of_complex ?A) = poincare_distance (of_complex z1) (of_complex ?B)"
"poincare_distance (of_complex z2) (of_complex ?A) = poincare_distance (of_complex z2) (of_complex ?B)"
"-1 < Re (Ze1 - Re1)" "Re (Ze1 - Re1) < 1" "-1 < Re (Ze1 + Re1)" "Re (Ze1 + Re1) < 1"
using \<open>is_real Ze1\<close> \<open>is_real Ze2\<close>
unfolding poincare_circle_def
by (auto simp add: cmod_eq_Re)
hence "z1 = z2"
using unique_midpoint_x_axis[of "Ze1 - Re1" "Ze1 + Re1"]
using \<open>is_real Ze1\<close> \<open>is_real z1\<close> \<open>is_real z2\<close> \<open>Re1 > 0\<close> \<open>-1 < Re z1\<close> \<open>Re z1 < 1\<close> \<open>-1 < Re z2\<close> \<open>Re z2 < 1\<close>
by auto
thus False
using \<open>z1 \<noteq> z2\<close>
by simp
qed
hence *: "(Re x1')\<^sup>2 + (Im x1')\<^sup>2 - 2 * Re x1' * Ze1 + Ze1 * Ze1 - cor (Re1 * Re1) = 0"
"(Re x1')\<^sup>2 + (Im x1')\<^sup>2 - 2 * Re x1' * Ze2 + Ze2 * Ze2 - cor (Re2 * Re2) = 0"
"(Re x2')\<^sup>2 + (Im x2')\<^sup>2 - 2 * Re x2' * Ze1 + Ze1 * Ze1 - cor (Re1 * Re1) = 0"
"(Re x2')\<^sup>2 + (Im x2')\<^sup>2 - 2 * Re x2' * Ze2 + Ze2 * Ze2 - cor (Re2 * Re2) = 0"
using circle_equation[of Re1 Ze1] circle_equation[of Re2 Ze2] circle
using eq_cnj_iff_real[of Ze1] \<open>is_real Ze1\<close> \<open>Re1 > 0\<close>
using eq_cnj_iff_real[of Ze2] \<open>is_real Ze2\<close> \<open>Re2 > 0\<close>
using complex_add_cnj[of x1'] complex_add_cnj[of x2']
using distrib_left[of Ze1 x1' "cnj x1'"] distrib_left[of Ze2 x1' "cnj x1'"]
using distrib_left[of Ze1 x2' "cnj x2'"] distrib_left[of Ze2 x2' "cnj x2'"]
by (auto simp add: complex_mult_cnj power2_eq_square field_simps)
hence "- 2 * Re x1' * Ze1 + Ze1 * Ze1 - cor (Re1 * Re1) = - 2 * Re x1' * Ze2 + Ze2 * Ze2 - cor (Re2 * Re2)"
"- 2 * Re x2' * Ze1 + Ze1 * Ze1 - cor (Re1 * Re1) = - 2 * Re x2' * Ze2 + Ze2 * Ze2 - cor (Re2 * Re2)"
by (smt add_diff_cancel_right' add_diff_eq eq_iff_diff_eq_0 minus_diff_eq mult_minus_left of_real_minus)+
hence "2 * Re x1' * (Ze2 - Ze1) = (Ze2 * Ze2 - cor (Re2 * Re2)) - (Ze1 * Ze1 - cor (Re1 * Re1))"
"2 * Re x2' * (Ze2 - Ze1) = (Ze2 * Ze2 - cor (Re2 * Re2)) - (Ze1 * Ze1 - cor (Re1 * Re1))"
by simp_all (simp add: field_simps)+
hence "2 * Re x1' * (Ze2 - Ze1) = 2 * Re x2' * (Ze2 - Ze1)"
by simp
hence "Re x1' = Re x2'"
using \<open>Ze1 \<noteq> Ze2\<close>
by simp
moreover
hence "(Im x1')\<^sup>2 = (Im x2')\<^sup>2"
using *(1) *(3)
by (simp add: \<open>is_real Ze1\<close> complex_eq_if_Re_eq)
hence "Im x1' = Im x2' \<or> Im x1' = -Im x2'"
using power2_eq_iff
by blast
ultimately
show ?thesis
- using x' `x1 \<noteq> x2`
+ using x' \<open>x1 \<noteq> x2\<close>
using complex.expand
by (metis cnj.code complex_surj conjugate_of_complex)
qed
text \<open>Two h-circles of the same radius centered at mutually conjugate points intersect at the x-axis\<close>
lemma intersect_poincare_circles_conjugate_centers:
assumes in_disc: "z1 \<in> unit_disc" "z2 \<in> unit_disc" and
"z1 \<noteq> z2" and "z1 = conjugate z2" and "r > 0" and
u: "u \<in> poincare_circle z1 r \<inter> poincare_circle z2 r"
shows "is_real (to_complex u)"
proof-
obtain z1e r1e z2e r2e where
euclidean: "(z1e, r1e) = poincare_circle_euclidean z1 r"
"(z2e, r2e) = poincare_circle_euclidean z2 r"
by (metis poincare_circle_euclidean_def)
obtain z1' z2' where z': "z1 = of_complex z1'" "z2 = of_complex z2'"
using inf_or_of_complex[of z1] inf_or_of_complex[of z2] in_disc
by auto
obtain u' where u': "u = of_complex u'"
using u inf_or_of_complex[of u]
by (auto simp add: poincare_circle_def)
have "z1' = cnj z2'"
using \<open>z1 = conjugate z2\<close> z'
by (auto simp add: of_complex_inj)
moreover
let ?cz = "1 - (cmod z2')\<^sup>2"
let ?den = "?cz * (cosh r - 1) / 2 + 1"
have "?cz > 0"
using in_disc z'
by (simp add: cmod_def)
hence "?den \<ge> 1"
using cosh_gt_1[OF \<open>r > 0\<close>]
by auto
hence "?den \<noteq> 0"
by simp
hence "cor ?den \<noteq> 0"
using of_real_eq_0_iff
by blast
ultimately
have "r1e = r2e" "z1e = cnj z2e" "z1e \<noteq> z2e"
using z' euclidean \<open>z1 \<noteq> z2\<close>
unfolding poincare_circle_euclidean_def Let_def
by simp_all metis
hence "u' \<in> circle (cnj z2e) r2e \<inter> circle z2e r2e" "z2e \<noteq> cnj z2e"
using euclidean u u'
using poincare_circle_is_euclidean_circle[of z1 r]
using poincare_circle_is_euclidean_circle[of z2 r]
using in_disc \<open>r > 0\<close>
by auto (metis image_iff of_complex_inj)+
hence "(cmod (u' - z2e))\<^sup>2 = (cmod(u' - cnj z2e))\<^sup>2"
by (simp add: circle_def)
hence "(u' - z2e) * (cnj u' - cnj z2e) = (u' - cnj z2e) * (cnj u' - z2e)"
by (metis complex_cnj_cnj complex_cnj_diff complex_norm_square)
hence "(z2e - cnj z2e) * (u' - cnj u') = 0"
by (simp add: field_simps)
thus ?thesis
using u' \<open>z2e \<noteq> cnj z2e\<close> eq_cnj_iff_real[of u']
by simp
qed
subsection \<open>Congruent triangles\<close>
text\<open>For every pair of triangles such that its three pairs of sides are pairwise equal there is an
h-isometry (a unit disc preserving Möbius transform, eventually composed with a conjugation) that
maps one triangle onto the other.\<close>
lemma unit_disc_fix_f_congruent_triangles:
assumes
in_disc: "u \<in> unit_disc" "v \<in> unit_disc" "w \<in> unit_disc" and
in_disc': "u' \<in> unit_disc" "v' \<in> unit_disc" "w' \<in> unit_disc" and
d: "poincare_distance u v = poincare_distance u' v'"
"poincare_distance v w = poincare_distance v' w'"
"poincare_distance u w = poincare_distance u' w'"
shows
"\<exists> M. unit_disc_fix_f M \<and> M u = u' \<and> M v = v' \<and> M w = w'"
proof (cases "u = v \<or> u = w \<or> v = w")
case True
thus ?thesis
using assms
using poincare_distance_eq_0_iff[of u' v']
using poincare_distance_eq_0_iff[of v' w']
using poincare_distance_eq_0_iff[of u' w']
using poincare_distance_eq_ex_moebius[of v w v' w']
using poincare_distance_eq_ex_moebius[of u w u' w']
using poincare_distance_eq_ex_moebius[of u v u' v']
by (metis unit_disc_fix_f_def)
next
case False
have "\<forall> w u' v' w'. w \<in> unit_disc \<and> u' \<in> unit_disc \<and> v' \<in> unit_disc \<and> w' \<in> unit_disc \<and> w \<noteq> u \<and> w \<noteq> v \<and>
poincare_distance u v = poincare_distance u' v' \<and>
poincare_distance v w = poincare_distance v' w' \<and>
poincare_distance u w = poincare_distance u' w' \<longrightarrow>
(\<exists> M. unit_disc_fix_f M \<and> M u = u' \<and> M v = v' \<and> M w = w')" (is "?P u v")
proof (rule wlog_positive_x_axis[where P="?P"])
show "v \<in> unit_disc" "u \<in> unit_disc"
by fact+
next
show "u \<noteq> v"
using False
by simp
next
fix x
assume x: "is_real x" "0 < Re x" "Re x < 1"
hence "of_complex x \<noteq> 0\<^sub>h"
using of_complex_zero_iff[of x]
by (auto simp add: complex.expand)
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix w u' v' w'
assume in_disc: "w \<in> unit_disc" "u' \<in> unit_disc" "v' \<in> unit_disc" "w' \<in> unit_disc"
assume "poincare_distance 0\<^sub>h (of_complex x) = poincare_distance u' v'"
then obtain M' where M': "unit_disc_fix M'" "moebius_pt M' u' = 0\<^sub>h" "moebius_pt M' v' = (of_complex x)"
using poincare_distance_eq_ex_moebius[of u' v' "0\<^sub>h" "of_complex x"] in_disc x
by (auto simp add: cmod_eq_Re)
let ?w = "moebius_pt M' w'"
have "?w \<in> unit_disc"
using \<open>unit_disc_fix M'\<close> \<open>w' \<in> unit_disc\<close>
by simp
assume "w \<noteq> 0\<^sub>h" "w \<noteq> of_complex x"
hence dist_gt_0: "poincare_distance 0\<^sub>h w > 0" "poincare_distance (of_complex x) w > 0"
using poincare_distance_eq_0_iff[of "0\<^sub>h" w] in_disc poincare_distance_ge0[of "0\<^sub>h" w]
using poincare_distance_eq_0_iff[of "of_complex x" w] in_disc poincare_distance_ge0[of "of_complex x" w]
using x
by (simp_all add: cmod_eq_Re)
assume "poincare_distance (of_complex x) w = poincare_distance v' w'"
"poincare_distance 0\<^sub>h w = poincare_distance u' w'"
hence "poincare_distance 0\<^sub>h ?w = poincare_distance 0\<^sub>h w"
"poincare_distance (of_complex x) ?w = poincare_distance (of_complex x) w"
using M'(1) M'(2)[symmetric] M'(3)[symmetric] in_disc
using unit_disc_fix_preserve_poincare_distance[of M' u' w']
using unit_disc_fix_preserve_poincare_distance[of M' v' w']
by simp_all
hence "?w \<in> poincare_circle 0\<^sub>h (poincare_distance 0\<^sub>h w) \<inter> poincare_circle (of_complex x) (poincare_distance (of_complex x) w)"
"w \<in> poincare_circle 0\<^sub>h (poincare_distance 0\<^sub>h w) \<inter> poincare_circle (of_complex x) (poincare_distance (of_complex x) w)"
using \<open>?w \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>
unfolding poincare_circle_def
by simp_all
hence "?w = w \<or> ?w = conjugate w"
using intersect_poincare_circles_x_axis[of 0 x "poincare_distance 0\<^sub>h w" "poincare_distance (of_complex x) w" ?w w] x
using \<open>of_complex x \<noteq> 0\<^sub>h\<close> dist_gt_0
using poincare_distance_eq_0_iff
by auto
thus "\<exists>M. unit_disc_fix_f M \<and> M 0\<^sub>h = u' \<and> M (of_complex x) = v' \<and> M w = w'"
proof
assume "moebius_pt M' w' = w"
thus ?thesis
using M'
using moebius_pt_invert[of M' u' "0\<^sub>h"]
using moebius_pt_invert[of M' v' "of_complex x"]
using moebius_pt_invert[of M' w' "w"]
apply (rule_tac x="moebius_pt (-M')" in exI)
apply (simp add: unit_disc_fix_f_def)
apply (rule_tac x="-M'" in exI, simp)
done
next
let ?M = "moebius_pt (-M') \<circ> conjugate"
assume "moebius_pt M' w' = conjugate w"
hence "?M w = w'"
using moebius_pt_invert[of M' w' "conjugate w"]
by simp
moreover
have "?M 0\<^sub>h = u'" "?M (of_complex x) = v'"
using moebius_pt_invert[of M' u' "0\<^sub>h"]
using moebius_pt_invert[of M' v' "of_complex x"]
using M' \<open>is_real x\<close> eq_cnj_iff_real[of x]
by simp_all
moreover
have "unit_disc_fix_f ?M"
using \<open>unit_disc_fix M'\<close>
unfolding unit_disc_fix_f_def
by (rule_tac x="-M'" in exI, simp)
ultimately
show ?thesis
by blast
qed
qed
next
fix M u v
assume 1: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v"
assume 2: "?P ?Mu ?Mv"
show "?P u v"
proof safe
fix w u' v' w'
let ?Mw = "moebius_pt M w" and ?Mu' = "moebius_pt M u'" and ?Mv' = "moebius_pt M v'" and ?Mw' = "moebius_pt M w'"
assume "w \<in> unit_disc" "u' \<in> unit_disc" "v' \<in> unit_disc" "w' \<in> unit_disc" "w \<noteq> u" "w \<noteq> v"
"poincare_distance u v = poincare_distance u' v'"
"poincare_distance v w = poincare_distance v' w'"
"poincare_distance u w = poincare_distance u' w'"
then obtain M' where M': "unit_disc_fix_f M'" "M' ?Mu = ?Mu'" "M' ?Mv = ?Mv'" "M' ?Mw = ?Mw'"
using 1 2[rule_format, of ?Mw ?Mu' ?Mv' ?Mw']
by auto
let ?M = "moebius_pt (-M) \<circ> M' \<circ> moebius_pt M"
show "\<exists>M. unit_disc_fix_f M \<and> M u = u' \<and> M v = v' \<and> M w = w'"
proof (rule_tac x="?M" in exI, safe)
show "unit_disc_fix_f ?M"
using M'(1) \<open>unit_disc_fix M\<close>
by (subst unit_disc_fix_f_comp, subst unit_disc_fix_f_comp, simp_all)
next
show "?M u = u'" "?M v = v'" "?M w = w'"
using M'
by auto
qed
qed
qed
thus ?thesis
using assms False
by auto
qed
end
\ No newline at end of file
diff --git a/thys/Poincare_Disc/Poincare_Distance.thy b/thys/Poincare_Disc/Poincare_Distance.thy
--- a/thys/Poincare_Disc/Poincare_Distance.thy
+++ b/thys/Poincare_Disc/Poincare_Distance.thy
@@ -1,1567 +1,1567 @@
theory Poincare_Distance
imports Poincare_Lines_Ideal_Points Hyperbolic_Functions
begin
(* ------------------------------------------------------------------ *)
section \<open>H-distance in the Poincar\'e model\<close>
(* ------------------------------------------------------------------ *)
text\<open>Informally, the \emph{h-distance} between the two h-points is defined as the absolute value of
the logarithm of the cross ratio between those two points and the two ideal points.\<close>
abbreviation Re_cross_ratio where "Re_cross_ratio z u v w \<equiv> Re (to_complex (cross_ratio z u v w))"
definition calc_poincare_distance :: "complex_homo \<Rightarrow> complex_homo \<Rightarrow> complex_homo \<Rightarrow> complex_homo \<Rightarrow> real" where
[simp]: "calc_poincare_distance u i1 v i2 = abs (ln (Re_cross_ratio u i1 v i2))"
definition poincare_distance_pred :: "complex_homo \<Rightarrow> complex_homo \<Rightarrow> real \<Rightarrow> bool" where
[simp]: "poincare_distance_pred u v d \<longleftrightarrow>
(u = v \<and> d = 0) \<or> (u \<noteq> v \<and> (\<forall> i1 i2. ideal_points (poincare_line u v) = {i1, i2} \<longrightarrow> d = calc_poincare_distance u i1 v i2))"
definition poincare_distance :: "complex_homo \<Rightarrow> complex_homo \<Rightarrow> real" where
"poincare_distance u v = (THE d. poincare_distance_pred u v d)"
text\<open>We shown that the described cross-ratio is always finite,
positive real number.\<close>
lemma distance_cross_ratio_real_positive:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> v"
shows "\<forall> i1 i2. ideal_points (poincare_line u v) = {i1, i2} \<longrightarrow>
cross_ratio u i1 v i2 \<noteq> \<infinity>\<^sub>h \<and> is_real (to_complex (cross_ratio u i1 v i2)) \<and> Re_cross_ratio u i1 v i2 > 0" (is "?P u v")
proof (rule wlog_positive_x_axis[OF assms])
fix x
assume *: "is_real x" "0 < Re x" "Re x < 1"
hence "x \<noteq> -1" "x \<noteq> 1"
by auto
hence **: "of_complex x \<noteq> \<infinity>\<^sub>h" "of_complex x \<noteq> 0\<^sub>h" "of_complex x \<noteq> of_complex (-1)" "of_complex 1 \<noteq> of_complex x"
"of_complex x \<in> circline_set x_axis"
using *
unfolding circline_set_x_axis
by (auto simp add: of_complex_inj)
have ***: "0\<^sub>h \<noteq> of_complex (-1)" "0\<^sub>h \<noteq> of_complex 1"
by (metis of_complex_zero_iff zero_neq_neg_one, simp)
have ****: "- x - 1 \<noteq> 0" "x - 1 \<noteq> 0"
using \<open>x \<noteq> -1\<close> \<open>x \<noteq> 1\<close>
by (metis add.inverse_inverse eq_iff_diff_eq_0, simp)
have "poincare_line 0\<^sub>h (of_complex x) = x_axis"
using **
by (simp add: poincare_line_0_real_is_x_axis)
thus "?P 0\<^sub>h (of_complex x)"
using * ** *** ****
using cross_ratio_not_inf[of "0\<^sub>h" "of_complex 1" "of_complex (-1)" "of_complex x"]
using cross_ratio_not_inf[of "0\<^sub>h" "of_complex (-1)" "of_complex 1" "of_complex x"]
using cross_ratio_real[of 0 "-1" x 1] cross_ratio_real[of 0 1 x "-1"]
apply (auto simp add: poincare_line_0_real_is_x_axis doubleton_eq_iff circline_set_x_axis)
apply (subst cross_ratio, simp_all, subst Re_complex_div_gt_0, simp, subst mult_neg_neg, simp_all)+
done
next
fix M u v
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v"
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
"?P ?Mu ?Mv"
show "?P u v"
proof safe
fix i1 i2
let ?cr = "cross_ratio u i1 v i2"
assume **: "ideal_points (poincare_line u v) = {i1, i2}"
have "i1 \<noteq> u" "i1 \<noteq> v" "i2 \<noteq> u" "i2 \<noteq> v" "i1 \<noteq> i2"
using ideal_points_different[OF *(2-3), of i1 i2] ** \<open>u \<noteq> v\<close>
by auto
hence "0 < Re (to_complex ?cr) \<and> is_real (to_complex ?cr) \<and> ?cr \<noteq> \<infinity>\<^sub>h"
using * **
apply (erule_tac x="moebius_pt M i1" in allE)
apply (erule_tac x="moebius_pt M i2" in allE)
apply (subst (asm) ideal_points_poincare_line_moebius[of M u v i1 i2], simp_all)
done
thus "0 < Re (to_complex ?cr)" "is_real (to_complex ?cr)" "?cr = \<infinity>\<^sub>h \<Longrightarrow> False"
by simp_all
qed
qed
text\<open>Next we can show that for every different points from the unit disc there is exactly one number
that satisfies the h-distance predicate.\<close>
lemma distance_unique:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "\<exists>! d. poincare_distance_pred u v d"
proof (cases "u = v")
case True
thus ?thesis
by auto
next
case False
obtain i1 i2 where *: "i1 \<noteq> i2" "ideal_points (poincare_line u v) = {i1, i2}"
using obtain_ideal_points[OF is_poincare_line_poincare_line] \<open>u \<noteq> v\<close>
by blast
let ?d = "calc_poincare_distance u i1 v i2"
show ?thesis
proof (rule ex1I)
show "poincare_distance_pred u v ?d"
using * \<open>u \<noteq> v\<close>
proof (simp del: calc_poincare_distance_def, safe)
fix i1' i2'
assume "{i1, i2} = {i1', i2'}"
hence **: "(i1' = i1 \<and> i2' = i2) \<or> (i1' = i2 \<and> i2' = i1)"
using doubleton_eq_iff[of i1 i2 i1' i2']
by blast
have all_different: "u \<noteq> i1" "u \<noteq> i2" "v \<noteq> i1" "v \<noteq> i2" "u \<noteq> i1'" "u \<noteq> i2'" "v \<noteq> i1'" "v \<noteq> i2'" "i1 \<noteq> i2"
using ideal_points_different[OF assms, of i1 i2] * ** \<open>u \<noteq> v\<close>
by auto
show "calc_poincare_distance u i1 v i2 = calc_poincare_distance u i1' v i2'"
proof-
let ?cr = "cross_ratio u i1 v i2"
let ?cr' = "cross_ratio u i1' v i2'"
have "Re (to_complex ?cr) > 0" "is_real (to_complex ?cr)"
"Re (to_complex ?cr') > 0" "is_real (to_complex ?cr')"
using False distance_cross_ratio_real_positive[OF assms(1-2)] * **
by auto
thus ?thesis
using **
using cross_ratio_not_zero cross_ratio_not_inf all_different
by auto (subst cross_ratio_commute_24, subst reciprocal_real, simp_all add: ln_div)
qed
qed
next
fix d
assume "poincare_distance_pred u v d"
thus "d = ?d"
using * \<open>u \<noteq> v\<close>
by auto
qed
qed
lemma poincare_distance_satisfies_pred [simp]:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance_pred u v (poincare_distance u v)"
using distance_unique[OF assms] theI'[of "poincare_distance_pred u v"]
unfolding poincare_distance_def
by blast
lemma poincare_distance_I:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> v" and "ideal_points (poincare_line u v) = {i1, i2}"
shows "poincare_distance u v = calc_poincare_distance u i1 v i2"
using assms
using poincare_distance_satisfies_pred[OF assms(1-2)]
by simp
lemma poincare_distance_refl [simp]:
assumes "u \<in> unit_disc"
shows "poincare_distance u u = 0"
using assms
using poincare_distance_satisfies_pred[OF assms assms]
by simp
text\<open>Unit disc preserving Möbius transformations preserve h-distance. \<close>
lemma unit_disc_fix_preserve_poincare_distance [simp]:
assumes "unit_disc_fix M" and "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance (moebius_pt M u) (moebius_pt M v) = poincare_distance u v"
proof (cases "u = v")
case True
have "moebius_pt M u \<in> unit_disc" "moebius_pt M v \<in> unit_disc"
using unit_disc_fix_iff[OF assms(1), symmetric] assms
by blast+
thus ?thesis
using assms \<open>u = v\<close>
by simp
next
case False
obtain i1 i2 where *: "ideal_points (poincare_line u v) = {i1, i2}"
using \<open>u \<noteq> v\<close>
by (rule obtain_ideal_points[OF is_poincare_line_poincare_line[of u v]])
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v" and ?Mi1 = "moebius_pt M i1" and ?Mi2 = "moebius_pt M i2"
have **: "?Mu \<in> unit_disc" "?Mv \<in> unit_disc"
using assms
using unit_disc_fix_iff
by blast+
have ***: "?Mu \<noteq> ?Mv"
using \<open>u \<noteq> v\<close>
by simp
have "poincare_distance u v = calc_poincare_distance u i1 v i2"
using poincare_distance_I[OF assms(2-3) \<open>u \<noteq> v\<close> *]
by auto
moreover
have "unit_circle_fix M"
using assms
by simp
hence ++: "ideal_points (poincare_line ?Mu ?Mv) = {?Mi1, ?Mi2}"
using \<open>u \<noteq> v\<close> assms *
by simp
have "poincare_distance ?Mu ?Mv = calc_poincare_distance ?Mu ?Mi1 ?Mv ?Mi2"
by (rule poincare_distance_I[OF ** *** ++])
moreover
have "calc_poincare_distance ?Mu ?Mi1 ?Mv ?Mi2 = calc_poincare_distance u i1 v i2"
using ideal_points_different[OF assms(2-3) \<open>u \<noteq> v\<close> *]
unfolding calc_poincare_distance_def
by (subst moebius_preserve_cross_ratio[symmetric], simp_all)
ultimately
show ?thesis
by simp
qed
text\<open>Knowing ideal points for x-axis, we can easily explicitly calculate distances.\<close>
lemma poincare_distance_x_axis_x_axis:
assumes "x \<in> unit_disc" and "y \<in> unit_disc" and "x \<in> circline_set x_axis" and "y \<in> circline_set x_axis"
shows "poincare_distance x y =
(let x' = to_complex x; y' = to_complex y
in abs (ln (Re (((1 + x') * (1 - y')) / ((1 - x') * (1 + y'))))))"
proof-
obtain x' y' where *: "x = of_complex x'" "y = of_complex y'"
using inf_or_of_complex[of x] inf_or_of_complex[of y] \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
by auto
have "cmod x' < 1" "cmod y' < 1"
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> *
by (metis unit_disc_iff_cmod_lt_1)+
hence **: "x' \<noteq> 1" "x' \<noteq> 1" "y' \<noteq> -1" "y' \<noteq> 1"
by auto
have "1 + y' \<noteq> 0"
using **
by (metis add.left_cancel add_neg_numeral_special(7))
show ?thesis
proof (cases "x = y")
case True
thus ?thesis
using assms(1-2)
- using unit_disc_iff_cmod_lt_1[of "to_complex x"] * ** `1 + y' \<noteq> 0`
+ using unit_disc_iff_cmod_lt_1[of "to_complex x"] * ** \<open>1 + y' \<noteq> 0\<close>
by auto
next
case False
hence "poincare_line x y = x_axis"
using poincare_line_x_axis[OF assms]
by simp
hence "ideal_points (poincare_line x y) = {of_complex (-1), of_complex 1}"
by simp
hence "poincare_distance x y = calc_poincare_distance x (of_complex (-1)) y (of_complex 1)"
using poincare_distance_I assms \<open>x \<noteq> y\<close>
by auto
also have "... = abs (ln (Re (((x' + 1) * (y' - 1)) / ((x' - 1) * (y' + 1)))))"
using * \<open>cmod x' < 1\<close> \<open>cmod y' < 1\<close>
by (simp, transfer, transfer, auto)
finally
show ?thesis
using *
by (metis (no_types, lifting) add.commute minus_diff_eq minus_divide_divide mult_minus_left mult_minus_right to_complex_of_complex)
qed
qed
lemma poincare_distance_zero_x_axis:
assumes "x \<in> unit_disc" and "x \<in> circline_set x_axis"
shows "poincare_distance 0\<^sub>h x = (let x' = to_complex x in abs (ln (Re ((1 - x') / (1 + x')))))"
using assms
using poincare_distance_x_axis_x_axis[of "0\<^sub>h" x]
by (simp add: Let_def)
lemma poincare_distance_zero:
assumes "x \<in> unit_disc"
shows "poincare_distance 0\<^sub>h x = (let x' = to_complex x in abs (ln (Re ((1 - cmod x') / (1 + cmod x')))))" (is "?P x")
proof (cases "x = 0\<^sub>h")
case True
thus ?thesis
by auto
next
case False
show ?thesis
proof (rule wlog_rotation_to_positive_x_axis)
show "x \<in> unit_disc" "x \<noteq> 0\<^sub>h" by fact+
next
fix \<phi> u
assume "u \<in> unit_disc" "u \<noteq> 0\<^sub>h" "?P (moebius_pt (moebius_rotation \<phi>) u)"
thus "?P u"
using unit_disc_fix_preserve_poincare_distance[of "moebius_rotation \<phi>" "0\<^sub>h" u]
by (cases "u = \<infinity>\<^sub>h") (simp_all add: Let_def)
next
fix x
assume "is_real x" "0 < Re x" "Re x < 1"
thus "?P (of_complex x)"
using poincare_distance_zero_x_axis[of "of_complex x"]
by simp (auto simp add: circline_set_x_axis cmod_eq_Re complex_is_Real_iff)
qed
qed
lemma poincare_distance_zero_opposite [simp]:
assumes "of_complex z \<in> unit_disc"
shows "poincare_distance 0\<^sub>h (of_complex (- z)) = poincare_distance 0\<^sub>h (of_complex z)"
proof-
have *: "of_complex (-z) \<in> unit_disc"
using assms
by auto
show ?thesis
using poincare_distance_zero[OF assms]
using poincare_distance_zero[OF *]
by simp
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Distance explicit formula\<close>
(* ------------------------------------------------------------------ *)
text\<open>Instead of the h-distance itself, very frequently its hyperbolic cosine is analyzed.\<close>
abbreviation "cosh_dist u v \<equiv> cosh (poincare_distance u v)"
lemma cosh_poincare_distance_cross_ratio_average:
assumes "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v" "ideal_points (poincare_line u v) = {i1, i2}"
shows "cosh_dist u v =
((Re_cross_ratio u i1 v i2) + (Re_cross_ratio v i1 u i2)) / 2"
proof-
let ?cr = "cross_ratio u i1 v i2"
let ?crRe = "Re (to_complex ?cr)"
have "?cr \<noteq> \<infinity>\<^sub>h" "is_real (to_complex ?cr)" "?crRe > 0"
using distance_cross_ratio_real_positive[OF assms(1-3)] assms(4)
by simp_all
then obtain cr where *: "cross_ratio u i1 v i2 = of_complex cr" "cr \<noteq> 0" "is_real cr" "Re cr > 0"
using inf_or_of_complex[of "cross_ratio u i1 v i2"]
by (smt to_complex_of_complex zero_complex.simps(1))
thus ?thesis
using *
using assms cross_ratio_commute_13[of v i1 u i2]
unfolding poincare_distance_I[OF assms] calc_poincare_distance_def cosh_def
by (cases "Re cr \<ge> 1")
(auto simp add: ln_div[of 0] exp_minus field_simps Re_divide power2_eq_square complex.expand)
qed
definition poincare_distance_formula' :: "complex \<Rightarrow> complex \<Rightarrow> real" where
[simp]: "poincare_distance_formula' u v = 1 + 2 * ((cmod (u - v))\<^sup>2 / ((1 - (cmod u)\<^sup>2) * (1 - (cmod v)\<^sup>2)))"
text\<open>Next we show that the following formula expresses h-distance between any two h-points (note
that the ideal points do not figure anymore).\<close>
definition poincare_distance_formula :: "complex \<Rightarrow> complex \<Rightarrow> real" where
[simp]: "poincare_distance_formula u v = arcosh (poincare_distance_formula' u v)"
lemma blaschke_preserve_distance_formula [simp]:
assumes "of_complex k \<in> unit_disc" "u \<in> unit_disc" "v \<in> unit_disc"
shows "poincare_distance_formula (to_complex (moebius_pt (blaschke k) u)) (to_complex (moebius_pt (blaschke k) v)) =
poincare_distance_formula (to_complex u) (to_complex v)"
proof (cases "k = 0")
case True
thus ?thesis
by simp
next
case False
obtain u' v' where *: "u' = to_complex u" "v' = to_complex v"
by auto
have "cmod u' < 1" "cmod v' < 1" "cmod k < 1"
using assms *
using inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
obtain nu du nv dv d kk ddu ddv where
**: "nu = u' - k" "du = 1 - cnj k *u'" "nv = v' - k" "dv = 1 - cnj k * v'"
"d = u' - v'" "ddu = 1 - u'*cnj u'" "ddv = 1 - v'*cnj v'" "kk = 1 - k*cnj k"
by auto
have d: "nu*dv - nv*du = d*kk"
by (subst **)+ (simp add: field_simps)
have ddu: "du*cnj du - nu*cnj nu = ddu*kk"
by (subst **)+ (simp add: field_simps)
have ddv: "dv*cnj dv - nv*cnj nv = ddv*kk"
by (subst **)+ (simp add: field_simps)
have "du \<noteq> 0"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "cmod (1 - cnj k * u') = 0"
using \<open>du = 1 - cnj k * u'\<close>
by auto
hence "cmod (cnj k * u') = 1"
by auto
hence "cmod k * cmod u' = 1"
by auto
thus False
using \<open>cmod k < 1\<close> \<open>cmod u' < 1\<close>
using mult_strict_mono[of "cmod k" 1 "cmod u'" 1]
by simp
qed
have "dv \<noteq> 0"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "cmod (1 - cnj k * v') = 0"
using \<open>dv = 1 - cnj k * v'\<close>
by auto
hence "cmod (cnj k * v') = 1"
by auto
hence "cmod k * cmod v' = 1"
by auto
thus False
using \<open>cmod k < 1\<close> \<open>cmod v' < 1\<close>
using mult_strict_mono[of "cmod k" 1 "cmod v'" 1]
by simp
qed
have "kk \<noteq> 0"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "cmod (1 - k * cnj k) = 0"
using \<open>kk = 1 - k * cnj k\<close>
by auto
hence "cmod (k * cnj k) = 1"
by auto
hence "cmod k * cmod k = 1"
by auto
thus False
using \<open>cmod k < 1\<close>
using mult_strict_mono[of "cmod k" 1 "cmod k" 1]
by simp
qed
note nz = \<open>du \<noteq> 0\<close> \<open>dv \<noteq> 0\<close> \<open>kk \<noteq> 0\<close>
have "nu / du - nv / dv = (nu*dv - nv*du) / (du * dv)"
using nz
by (simp add: field_simps)
hence "(cmod (nu/du - nv/dv))\<^sup>2 = cmod ((d*kk) / (du*dv) * (cnj ((d*kk) / (du*dv))))" (is "?lhs = _")
unfolding complex_mod_mult_cnj[symmetric]
by (subst (asm) d) simp
also have "... = cmod ((d*cnj d*kk*kk) / (du*cnj du*dv*cnj dv))"
by (simp add: field_simps)
finally have 1: "?lhs = cmod ((d*cnj d*kk*kk) / (du*cnj du*dv*cnj dv))"
.
have "(1 - ((cmod nu) / (cmod du))\<^sup>2)*(1 - ((cmod nv) / (cmod dv))\<^sup>2) =
(1 - cmod((nu * cnj nu) / (du * cnj du)))*(1 - cmod((nv * cnj nv) / (dv * cnj dv)))" (is "?rhs = _")
by (metis cmod_divide complex_mod_mult_cnj power_divide)
also have "... = cmod(((du*cnj du - nu*cnj nu) / (du * cnj du)) * ((dv*cnj dv - nv*cnj nv) / (dv * cnj dv)))"
proof-
have "u' \<noteq> 1 / cnj k" "v' \<noteq> 1 / cnj k"
using \<open>cmod u' < 1\<close> \<open>cmod v' < 1\<close> \<open>cmod k < 1\<close>
by (auto simp add: False)
moreover
have "cmod k \<noteq> 1"
using \<open>cmod k < 1\<close>
by linarith
ultimately
have "cmod (nu/du) < 1" "cmod (nv/dv) < 1"
using **(1-4)
using unit_disc_fix_discI[OF blaschke_unit_disc_fix[OF \<open>cmod k < 1\<close>] \<open>u \<in> unit_disc\<close>] \<open>u' = to_complex u\<close>
using unit_disc_fix_discI[OF blaschke_unit_disc_fix[OF \<open>cmod k < 1\<close>] \<open>v \<in> unit_disc\<close>] \<open>v' = to_complex v\<close>
using inf_or_of_complex[of u] \<open>u \<in> unit_disc\<close> inf_or_of_complex[of v] \<open>v \<in> unit_disc\<close>
using moebius_pt_blaschke[of k u'] using moebius_pt_blaschke[of k v']
by auto
hence "(cmod (nu/du))\<^sup>2 < 1" "(cmod (nv/dv))\<^sup>2 < 1"
by (simp_all add: cmod_def)
hence "cmod (nu * cnj nu / (du * cnj du)) < 1" "cmod (nv * cnj nv / (dv * cnj dv)) < 1"
by (metis complex_mod_mult_cnj norm_divide power_divide)+
moreover
have "is_real (nu * cnj nu / (du * cnj du))" "is_real (nv * cnj nv / (dv * cnj dv))"
using eq_cnj_iff_real[of "nu * cnj nu / (du * cnj du)"]
using eq_cnj_iff_real[of "nv * cnj nv / (dv * cnj dv)"]
by (auto simp add: mult.commute)
moreover
have "Re (nu * cnj nu / (du * cnj du)) \<ge> 0" "Re (nv * cnj nv / (dv * cnj dv)) \<ge> 0"
using \<open>du \<noteq> 0\<close> \<open>dv \<noteq> 0\<close>
unfolding complex_mult_cnj_cmod
by simp_all
ultimately
have "1 - cmod (nu * cnj nu / (du * cnj du)) = cmod (1 - nu * cnj nu / (du * cnj du))"
"1 - cmod (nv * cnj nv / (dv * cnj dv)) = cmod (1 - nv * cnj nv / (dv * cnj dv))"
by (simp_all add: cmod_def)
thus ?thesis
using nz
apply simp
apply (subst diff_divide_eq_iff, simp, simp)
apply (subst diff_divide_eq_iff, simp, simp)
done
qed
also have "... = cmod(((ddu * kk) / (du * cnj du)) * ((ddv * kk) / (dv * cnj dv)))"
by (subst ddu, subst ddv, simp)
also have "... = cmod((ddu*ddv*kk*kk) / (du*cnj du*dv*cnj dv))"
by (simp add: field_simps)
finally have 2: "?rhs = cmod((ddu*ddv*kk*kk) / (du*cnj du*dv*cnj dv))"
.
have "?lhs / ?rhs =
cmod ((d*cnj d*kk*kk) / (du*cnj du*dv*cnj dv)) / cmod((ddu*ddv*kk*kk) / (du*cnj du*dv*cnj dv))"
by (subst 1, subst 2, simp)
also have "... = cmod ((d*cnj d)/(ddu*ddv))"
using nz
by simp
also have "... = (cmod d)\<^sup>2 / ((1 - (cmod u')\<^sup>2)*(1 - (cmod v')\<^sup>2))"
proof-
have "(cmod u')\<^sup>2 < 1" "(cmod v')\<^sup>2 < 1"
using \<open>cmod u' < 1\<close> \<open>cmod v' < 1\<close>
by (simp_all add: cmod_def)
hence "cmod (1 - u' * cnj u') = 1 - (cmod u')\<^sup>2" "cmod (1 - v' * cnj v') = 1 - (cmod v')\<^sup>2"
by (auto simp add: cmod_eq_Re cmod_power2 power2_eq_square[symmetric])
thus ?thesis
using nz
apply (subst **)+
unfolding complex_mod_mult_cnj[symmetric]
by simp
qed
finally
have 3: "?lhs / ?rhs = (cmod d)\<^sup>2 / ((1 - (cmod u')\<^sup>2)*(1 - (cmod v')\<^sup>2))"
.
have "cmod k \<noteq> 1" "u' \<noteq> 1 / cnj k" "v' \<noteq> 1 / cnj k" "u \<noteq> \<infinity>\<^sub>h" "v \<noteq> \<infinity>\<^sub>h"
using \<open>cmod k < 1\<close> \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close> * \<open>k \<noteq> 0\<close> ** \<open>kk \<noteq> 0\<close> nz
by auto
thus ?thesis using assms
using * ** 3
using moebius_pt_blaschke[of k u']
using moebius_pt_blaschke[of k v']
by simp
qed
text \<open>To prove the equivalence between the h-distance definition and the distance formula, we shall
employ the without loss of generality principle. Therefore, we must show that the distance formula
is preserved by h-isometries.\<close>
text\<open>Rotation preserve @{term poincare_distance_formula}.\<close>
lemma rotation_preserve_distance_formula [simp]:
assumes "u \<in> unit_disc" "v \<in> unit_disc"
shows "poincare_distance_formula (to_complex (moebius_pt (moebius_rotation \<phi>) u)) (to_complex (moebius_pt (moebius_rotation \<phi>) v)) =
poincare_distance_formula (to_complex u) (to_complex v)"
using assms
using inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
text\<open>Unit disc fixing Möbius preserve @{term poincare_distance_formula}.\<close>
lemma unit_disc_fix_preserve_distance_formula [simp]:
assumes "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
shows "poincare_distance_formula (to_complex (moebius_pt M u)) (to_complex (moebius_pt M v)) =
poincare_distance_formula (to_complex u) (to_complex v)" (is "?P' u v M")
proof-
have "\<forall> u \<in> unit_disc. \<forall> v \<in> unit_disc. ?P' u v M" (is "?P M")
proof (rule wlog_unit_disc_fix[OF assms(1)])
fix k
assume "cmod k < 1"
hence "of_complex k \<in> unit_disc"
by simp
thus "?P (blaschke k)"
using blaschke_preserve_distance_formula
by simp
next
fix \<phi>
show "?P (moebius_rotation \<phi>)"
using rotation_preserve_distance_formula
by simp
next
fix M1 M2
assume *: "?P M1" and **: "?P M2" and u11: "unit_disc_fix M1" "unit_disc_fix M2"
thus "?P (M1 + M2)"
by (auto simp del: poincare_distance_formula_def)
qed
thus ?thesis
using assms
by simp
qed
text\<open>The equivalence between the two h-distance representations.\<close>
lemma poincare_distance_formula:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance u v = poincare_distance_formula (to_complex u) (to_complex v)" (is "?P u v")
proof (rule wlog_x_axis)
fix x
assume *: "is_real x" "0 \<le> Re x" "Re x < 1"
show "?P 0\<^sub>h (of_complex x)" (is "?lhs = ?rhs")
proof-
have "of_complex x \<in> unit_disc" "of_complex x \<in> circline_set x_axis" "cmod x < 1"
using * cmod_eq_Re
by (auto simp add: circline_set_x_axis)
hence "?lhs = \<bar>ln (Re ((1 - x) / (1 + x)))\<bar>"
using poincare_distance_zero_x_axis[of "of_complex x"]
by simp
moreover
have "?rhs = \<bar>ln (Re ((1 - x) / (1 + x)))\<bar>"
proof-
let ?x = "1 + 2 * (cmod x)\<^sup>2 / (1 - (cmod x)\<^sup>2)"
have "0 \<le> 2 * (cmod x)\<^sup>2 / (1 - (cmod x)\<^sup>2)"
by (smt \<open>cmod x < 1\<close> divide_nonneg_nonneg norm_ge_zero power_le_one zero_le_power2)
hence arcosh_real_gt: "1 \<le> ?x"
by auto
have "?rhs = arcosh ?x"
by simp
also have "... = ln ((1 + (cmod x)\<^sup>2) / (1 - (cmod x)\<^sup>2) + 2 * (cmod x) / (1 - (cmod x)\<^sup>2))"
proof-
have "1 - (cmod x)\<^sup>2 > 0"
using \<open>cmod x < 1\<close>
by (smt norm_not_less_zero one_power2 power2_eq_imp_eq power_mono)
hence 1: "?x = (1 + (cmod x)\<^sup>2) / (1 - (cmod x)\<^sup>2)"
by (simp add: field_simps)
have 2: "?x\<^sup>2 - 1 = (4 * (cmod x)\<^sup>2) / (1 - (cmod x)\<^sup>2)\<^sup>2"
using \<open>1 - (cmod x)\<^sup>2 > 0\<close>
apply (subst 1)
unfolding power_divide
by (subst divide_diff_eq_iff, simp, simp add: power2_eq_square field_simps)
show ?thesis
using \<open>1 - (cmod x)\<^sup>2 > 0\<close>
apply (subst arcosh_real_def[OF arcosh_real_gt])
apply (subst 2)
apply (subst 1)
apply (subst real_sqrt_divide)
apply (subst real_sqrt_mult)
apply simp
done
qed
also have "... = ln (((1 + (cmod x))\<^sup>2) / (1 - (cmod x)\<^sup>2))"
apply (subst add_divide_distrib[symmetric])
apply (simp add: field_simps power2_eq_square)
done
also have "... = ln ((1 + cmod x) / (1 - (cmod x)))"
using \<open>cmod x < 1\<close>
using square_diff_square_factored[of 1 "cmod x"]
by (simp add: power2_eq_square)
also have "... = \<bar>ln (Re ((1 - x) / (1 + x)))\<bar>"
proof-
have *: "Re ((1 - x) / (1 + x)) \<le> 1" "Re ((1 - x) / (1 + x)) > 0"
using \<open>is_real x\<close> \<open>Re x \<ge> 0\<close> \<open>Re x < 1\<close>
using complex_is_Real_iff
by auto
hence "\<bar>ln (Re ((1 - x) / (1 + x)))\<bar> = - ln (Re ((1 - x) / (1 + x)))"
by auto
hence "\<bar>ln (Re ((1 - x) / (1 + x)))\<bar> = ln (Re ((1 + x) / (1 - x)))"
using ln_div[of 1 "Re ((1 - x)/(1 + x))"] * \<open>is_real x\<close>
by (simp add: complex_is_Real_iff)
moreover
have "ln ((1 + cmod x) / (1 - cmod x)) = ln ((1 + Re x) / (1 - Re x))"
using \<open>Re x \<ge> 0\<close> \<open>is_real x\<close>
using cmod_eq_Re by auto
moreover
have "(1 + Re x) / (1 - Re x) = Re ((1 + x) / (1 - x))"
using \<open>is_real x\<close> \<open>Re x < 1\<close>
by (smt Re_divide_real eq_iff_diff_eq_0 minus_complex.simps one_complex.simps plus_complex.simps)
ultimately
show ?thesis
by simp
qed
finally
show ?thesis
.
qed
ultimately
show ?thesis
by simp
qed
next
fix M u v
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
assume "?P (moebius_pt M u) (moebius_pt M v)"
thus "?P u v"
using *(1-3)
by (simp del: poincare_distance_formula_def)
next
show "u \<in> unit_disc" "v \<in> unit_disc"
by fact+
qed
text\<open>Some additional properties proved easily using the distance formula.\<close>
text \<open>@{term poincare_distance} is symmetric.\<close>
lemma poincare_distance_sym:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance u v = poincare_distance v u"
using assms
using poincare_distance_formula[OF assms(1) assms(2)]
using poincare_distance_formula[OF assms(2) assms(1)]
by (simp add: mult.commute norm_minus_commute)
lemma poincare_distance_formula'_ge_1:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "1 \<le> poincare_distance_formula' (to_complex u) (to_complex v)"
using unit_disc_cmod_square_lt_1[OF assms(1)] unit_disc_cmod_square_lt_1[OF assms(2)]
by auto
text\<open>@{term poincare_distance} is non-negative.\<close>
lemma poincare_distance_ge0:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance u v \<ge> 0"
using poincare_distance_formula'_ge_1
unfolding poincare_distance_formula[OF assms(1) assms(2)]
unfolding poincare_distance_formula_def
unfolding poincare_distance_formula'_def
by (rule arcosh_ge_0, simp_all add: assms)
lemma cosh_dist:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "cosh_dist u v = poincare_distance_formula' (to_complex u) (to_complex v)"
using poincare_distance_formula[OF assms] poincare_distance_formula'_ge_1[OF assms]
by simp
text\<open>@{term poincare_distance} is zero only if the two points are equal.\<close>
lemma poincare_distance_eq_0_iff:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance u v = 0 \<longleftrightarrow> u = v"
using assms
apply auto
using poincare_distance_formula'_ge_1[OF assms]
using unit_disc_cmod_square_lt_1[OF assms(1)] unit_disc_cmod_square_lt_1[OF assms(2)]
unfolding poincare_distance_formula[OF assms(1) assms(2)]
unfolding poincare_distance_formula_def
unfolding poincare_distance_formula'_def
apply (subst (asm) arcosh_eq_0_iff)
apply assumption
apply (simp add: unit_disc_to_complex_inj)
done
text\<open>Conjugate preserve @{term poincare_distance_formula}.\<close>
lemma conjugate_preserve_poincare_distance [simp]:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance (conjugate u) (conjugate v) = poincare_distance u v"
proof-
obtain u' v' where *: "u = of_complex u'" "v = of_complex v'"
using assms inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
have **: "conjugate u \<in> unit_disc" "conjugate v \<in> unit_disc"
using * assms
by auto
show ?thesis
using *
using poincare_distance_formula[OF assms]
using poincare_distance_formula[OF **]
by (metis complex_cnj_diff complex_mod_cnj conjugate_of_complex poincare_distance_def poincare_distance_formula'_def poincare_distance_formula_def to_complex_of_complex)
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Existence and uniqueness of points with a given distance\<close>
(* ------------------------------------------------------------------ *)
lemma ex_x_axis_poincare_distance_negative':
fixes d :: real
assumes "d \<ge> 0"
shows "let z = (1 - exp d) / (1 + exp d)
in is_real z \<and> Re z \<le> 0 \<and> Re z > -1 \<and>
of_complex z \<in> unit_disc \<and> of_complex z \<in> circline_set x_axis \<and>
poincare_distance 0\<^sub>h (of_complex z) = d"
proof-
have "exp d \<ge> 1"
using assms
using one_le_exp_iff[of d, symmetric]
by blast
hence "1 + exp d \<noteq> 0"
by linarith
let ?z = "(1 - exp d) / (1 + exp d)"
have "?z \<le> 0"
using \<open>exp d \<ge> 1\<close>
by (simp add: divide_nonpos_nonneg)
moreover
have "?z > -1"
using exp_gt_zero[of d]
by (smt divide_less_eq_1_neg nonzero_minus_divide_right)
moreover
hence "abs ?z < 1"
using \<open>?z \<le> 0\<close>
by simp
hence "cmod ?z < 1"
by (metis norm_of_real)
hence "of_complex ?z \<in> unit_disc"
by simp
moreover
have "of_complex ?z \<in> circline_set x_axis"
unfolding circline_set_x_axis
by simp
moreover
have "(1 - ?z) / (1 + ?z) = exp d"
proof-
have "1 + ?z = 2 / (1 + exp d)"
using \<open>1 + exp d \<noteq> 0\<close>
by (subst add_divide_eq_iff, auto)
moreover
have "1 - ?z = 2 * exp d / (1 + exp d)"
using \<open>1 + exp d \<noteq> 0\<close>
by (subst diff_divide_eq_iff, auto)
ultimately
show ?thesis
using \<open>1 + exp d \<noteq> 0\<close>
by simp
qed
ultimately
show ?thesis
using poincare_distance_zero_x_axis[of "of_complex ?z"]
using \<open>d \<ge> 0\<close> \<open>exp d \<ge> 1\<close>
by simp (simp add: cmod_eq_Re)
qed
lemma ex_x_axis_poincare_distance_negative:
assumes "d \<ge> 0"
shows "\<exists> z. is_real z \<and> Re z \<le> 0 \<and> Re z > -1 \<and>
of_complex z \<in> unit_disc \<and> of_complex z \<in> circline_set x_axis \<and>
poincare_distance 0\<^sub>h (of_complex z) = d" (is "\<exists> z. ?P z")
using ex_x_axis_poincare_distance_negative'[OF assms]
unfolding Let_def
by blast
text\<open>For each real number $d$ there is exactly one point on the positive x-axis such that h-distance
between 0 and that point is $d$.\<close>
lemma unique_x_axis_poincare_distance_negative:
assumes "d \<ge> 0"
shows "\<exists>! z. is_real z \<and> Re z \<le> 0 \<and> Re z > -1 \<and>
poincare_distance 0\<^sub>h (of_complex z) = d" (is "\<exists>! z. ?P z")
proof-
let ?z = "(1 - exp d) / (1 + exp d)"
have "?P ?z"
using ex_x_axis_poincare_distance_negative'[OF assms]
unfolding Let_def
by blast
moreover
have "\<forall> z'. ?P z' \<longrightarrow> z' = ?z"
proof-
let ?g = "\<lambda> x'. \<bar>ln (Re ((1 - x') / (1 + x')))\<bar>"
let ?A = "{x. is_real x \<and> Re x > -1 \<and> Re x \<le> 0}"
have "inj_on (poincare_distance 0\<^sub>h \<circ> of_complex) ?A"
proof (rule comp_inj_on)
show "inj_on of_complex ?A"
using of_complex_inj
unfolding inj_on_def
by blast
next
show "inj_on (poincare_distance 0\<^sub>h) (of_complex ` ?A)" (is "inj_on ?f (of_complex ` ?A)")
proof (subst inj_on_cong)
have *: "of_complex ` ?A =
{z. z \<in> unit_disc \<and> z \<in> circline_set x_axis \<and> Re (to_complex z) \<le> 0}" (is "_ = ?B")
by (auto simp add: cmod_eq_Re circline_set_x_axis)
fix x
assume "x \<in> of_complex ` ?A"
hence "x \<in> ?B"
using *
by simp
thus "poincare_distance 0\<^sub>h x = (?g \<circ> to_complex) x"
using poincare_distance_zero_x_axis
by (simp add: Let_def)
next
have *: "to_complex ` of_complex ` ?A = ?A"
by (auto simp add: image_iff)
show "inj_on (?g \<circ> to_complex) (of_complex ` ?A)"
proof (rule comp_inj_on)
show "inj_on to_complex (of_complex ` ?A)"
unfolding inj_on_def
by auto
next
have "inj_on ?g ?A"
unfolding inj_on_def
proof(safe)
fix x y
assume hh: "is_real x" "is_real y" "- 1 < Re x" "Re x \<le> 0"
"- 1 < Re y" "Re y \<le> 0" "\<bar>ln (Re ((1 - x) / (1 + x)))\<bar> = \<bar>ln (Re ((1 - y) / (1 + y)))\<bar>"
have "is_real ((1 - x)/(1 + x))"
using \<open>is_real x\<close> div_reals[of "1-x" "1+x"]
by auto
have "is_real ((1 - y)/(1 + y))"
using \<open>is_real y\<close> div_reals[of "1-y" "1+y"]
by auto
have "Re (1 + x) > 0"
using \<open>- 1 < Re x\<close> by auto
hence "1 + x \<noteq> 0"
by force
have "Re (1 - x) \<ge> 0"
using \<open>Re x \<le> 0\<close> by auto
hence "Re ((1 - x)/(1 + x)) > 0"
using Re_divide_real \<open>0 < Re (1 + x)\<close> complex_eq_if_Re_eq hh(1) hh(4) by auto
have "Re(1 - x) \<ge> Re ( 1 + x)"
using hh by auto
hence "Re ((1 - x)/(1 + x)) \<ge> 1"
using \<open>Re (1 + x) > 0\<close> \<open>is_real ((1 - x)/(1 + x))\<close>
by (smt Re_divide_real arg_0_iff hh(1) le_divide_eq_1_pos one_complex.simps(2) plus_complex.simps(2))
have "Re (1 + y) > 0"
using \<open>- 1 < Re y\<close> by auto
hence "1 + y \<noteq> 0"
by force
have "Re (1 - y) \<ge> 0"
using \<open>Re y \<le> 0\<close> by auto
hence "Re ((1 - y)/(1 + y)) > 0"
using Re_divide_real \<open>0 < Re (1 + y)\<close> complex_eq_if_Re_eq hh by auto
have "Re(1 - y) \<ge> Re ( 1 + y)"
using hh by auto
hence "Re ((1 - y)/(1 + y)) \<ge> 1"
using \<open>Re (1 + y) > 0\<close> \<open>is_real ((1 - y)/(1 + y))\<close>
by (smt Re_divide_real arg_0_iff hh le_divide_eq_1_pos one_complex.simps(2) plus_complex.simps(2))
have "ln (Re ((1 - x) / (1 + x))) = ln (Re ((1 - y) / (1 + y)))"
using \<open>Re ((1 - y)/(1 + y)) \<ge> 1\<close> \<open>Re ((1 - x)/(1 + x)) \<ge> 1\<close> hh
by auto
hence "Re ((1 - x) / (1 + x)) = Re ((1 - y) / (1 + y))"
using \<open>Re ((1 - y)/(1 + y)) > 0\<close> \<open>Re ((1 - x)/(1 + x)) > 0\<close>
by auto
hence "(1 - x) / (1 + x) = (1 - y) / (1 + y)"
using \<open>is_real ((1 - y)/(1 + y))\<close> \<open>is_real ((1 - x)/(1 + x))\<close>
using complex_eq_if_Re_eq by blast
hence "(1 - x) * (1 + y) = (1 - y) * (1 + x)"
using \<open>1 + y \<noteq> 0\<close> \<open>1 + x \<noteq> 0\<close>
by (simp add:field_simps)
thus "x = y"
by (simp add:field_simps)
qed
thus "inj_on ?g (to_complex ` of_complex ` ?A)"
using *
by simp
qed
qed
qed
thus ?thesis
using \<open>?P ?z\<close>
unfolding inj_on_def
by auto
qed
ultimately
show ?thesis
by blast
qed
lemma ex_x_axis_poincare_distance_positive:
assumes "d \<ge> 0"
shows "\<exists> z. is_real z \<and> Re z \<ge> 0 \<and> Re z < 1 \<and>
of_complex z \<in> unit_disc \<and> of_complex z \<in> circline_set x_axis \<and>
poincare_distance 0\<^sub>h (of_complex z) = d" (is "\<exists> z. is_real z \<and> Re z \<ge> 0 \<and> Re z < 1 \<and> ?P z")
proof-
obtain z where *: "is_real z" "Re z \<le> 0" "Re z > -1" "?P z"
using ex_x_axis_poincare_distance_negative[OF assms]
by auto
hence **: "of_complex z \<in> unit_disc" "of_complex z \<in> circline_set x_axis"
by (auto simp add: cmod_eq_Re)
have "is_real (-z) \<and> Re (-z) \<ge> 0 \<and> Re (-z) < 1 \<and> ?P (-z)"
using * **
by (simp add: circline_set_x_axis)
thus ?thesis
by blast
qed
lemma unique_x_axis_poincare_distance_positive:
assumes "d \<ge> 0"
shows "\<exists>! z. is_real z \<and> Re z \<ge> 0 \<and> Re z < 1 \<and>
poincare_distance 0\<^sub>h (of_complex z) = d" (is "\<exists>! z. is_real z \<and> Re z \<ge> 0 \<and> Re z < 1 \<and> ?P z")
proof-
obtain z where *: "is_real z" "Re z \<le> 0" "Re z > -1" "?P z"
using unique_x_axis_poincare_distance_negative[OF assms]
by auto
hence **: "of_complex z \<in> unit_disc" "of_complex z \<in> circline_set x_axis"
by (auto simp add: cmod_eq_Re circline_set_x_axis)
show ?thesis
proof
show "is_real (-z) \<and> Re (-z) \<ge> 0 \<and> Re (-z) < 1 \<and> ?P (-z)"
using * **
by simp
next
fix z'
assume "is_real z' \<and> Re z' \<ge> 0 \<and> Re z' < 1 \<and> ?P z'"
hence "is_real (-z') \<and> Re (-z') \<le> 0 \<and> Re (-z') > -1 \<and> ?P (-z')"
by (auto simp add: circline_set_x_axis cmod_eq_Re)
hence "-z' = z"
using unique_x_axis_poincare_distance_negative[OF assms] *
by blast
thus "z' = -z"
by auto
qed
qed
text\<open>Equal distance implies that segments are isometric - this means that congruence could be
defined either by two segments having the same distance or by requiring existence of an isometry
that maps one segment to the other.\<close>
lemma poincare_distance_eq_ex_moebius:
assumes in_disc: "u \<in> unit_disc" and "v \<in> unit_disc" and "u' \<in> unit_disc" and "v' \<in> unit_disc"
assumes "poincare_distance u v = poincare_distance u' v'"
shows "\<exists> M. unit_disc_fix M \<and> moebius_pt M u = u' \<and> moebius_pt M v = v'" (is "?P' u v u' v'")
proof (cases "u = v")
case True
thus ?thesis
using assms poincare_distance_eq_0_iff[of u' v']
by (simp add: unit_disc_fix_transitive)
next
case False
have "\<forall> u' v'. u \<noteq> v \<and> u' \<in> unit_disc \<and> v' \<in> unit_disc \<and> poincare_distance u v = poincare_distance u' v' \<longrightarrow>
?P' u' v' u v" (is "?P u v")
proof (rule wlog_positive_x_axis[where P="?P"])
fix x
assume "is_real x" "0 < Re x" "Re x < 1"
hence "of_complex x \<in> unit_disc" "of_complex x \<in> circline_set x_axis"
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re)
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix u' v'
assume "0\<^sub>h \<noteq> of_complex x" and in_disc: "u' \<in> unit_disc" "v' \<in> unit_disc" and
"poincare_distance 0\<^sub>h (of_complex x) = poincare_distance u' v'"
hence "u' \<noteq> v'" "poincare_distance u' v' > 0"
using poincare_distance_eq_0_iff[of "0\<^sub>h" "of_complex x"] \<open>of_complex x \<in> unit_disc\<close>
using poincare_distance_ge0[of "0\<^sub>h" "of_complex x"]
by auto
then obtain M where M: "unit_disc_fix M" "moebius_pt M u' = 0\<^sub>h" "moebius_pt M v' \<in> positive_x_axis"
using ex_unit_disc_fix_to_zero_positive_x_axis[of u' v'] in_disc
by auto
then obtain Mv' where Mv': "moebius_pt M v' = of_complex Mv'"
using inf_or_of_complex[of "moebius_pt M v'"] in_disc unit_disc_fix_iff[of M]
by (metis image_eqI inf_notin_unit_disc)
have "moebius_pt M v' \<in> unit_disc"
using M(1) \<open>v' \<in> unit_disc\<close>
by auto
have "Re Mv' > 0" "is_real Mv'" "Re Mv' < 1"
using M Mv' of_complex_inj \<open>moebius_pt M v' \<in> unit_disc\<close>
unfolding positive_x_axis_def circline_set_x_axis
using cmod_eq_Re
by auto fastforce
have "poincare_distance 0\<^sub>h (moebius_pt M v') = poincare_distance u' v'"
using M(1)
using in_disc
by (subst M(2)[symmetric], simp)
have "Mv' = x"
using \<open>poincare_distance 0\<^sub>h (moebius_pt M v') = poincare_distance u' v'\<close> Mv'
using \<open>poincare_distance 0\<^sub>h (of_complex x) = poincare_distance u' v'\<close>
using unique_x_axis_poincare_distance_positive[of "poincare_distance u' v'"]
\<open>poincare_distance u' v' > 0\<close>
using \<open>Re Mv' > 0\<close> \<open>Re Mv' < 1\<close> \<open>is_real Mv'\<close>
using \<open>is_real x\<close> \<open>Re x > 0\<close> \<open>Re x < 1\<close>
unfolding positive_x_axis_def
by auto
thus "?P' u' v' 0\<^sub>h (of_complex x)"
using M Mv'
by auto
qed
next
show "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
by fact+
next
fix M u v
let ?Mu = "moebius_pt M u" and ?Mv = "moebius_pt M v"
assume 1: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
hence 2: "?Mu \<noteq> ?Mv" "?Mu \<in> unit_disc" "?Mv \<in> unit_disc"
by auto
assume 3: "?P (moebius_pt M u) (moebius_pt M v)"
show "?P u v"
proof safe
fix u' v'
assume 4: "u' \<in> unit_disc" "v' \<in> unit_disc" "poincare_distance u v = poincare_distance u' v'"
hence "poincare_distance ?Mu ?Mv = poincare_distance u v"
using 1
by simp
then obtain M' where 5: "unit_disc_fix M'" "moebius_pt M' u' = ?Mu" "moebius_pt M' v' = ?Mv"
using 2 3 4
by auto
let ?M = "(-M) + M'"
have "unit_disc_fix ?M \<and> moebius_pt ?M u' = u \<and> moebius_pt ?M v' = v"
using 5 \<open>unit_disc_fix M\<close>
using unit_disc_fix_moebius_comp[of "-M" "M'"]
using unit_disc_fix_moebius_inv[of M]
by simp
thus "\<exists>M. unit_disc_fix M \<and> moebius_pt M u' = u \<and> moebius_pt M v' = v"
by blast
qed
qed
then obtain M where "unit_disc_fix M \<and> moebius_pt M u' = u \<and> moebius_pt M v' = v"
using assms \<open>u \<noteq> v\<close>
by blast
hence "unit_disc_fix (-M) \<and> moebius_pt (-M) u = u' \<and> moebius_pt (-M) v = v'"
using unit_disc_fix_moebius_inv[of M]
by auto
thus ?thesis
by blast
qed
lemma unique_midpoint_x_axis:
assumes x: "is_real x" "-1 < Re x" "Re x < 1" and
y: "is_real y" "-1 < Re y" "Re y < 1" and
"x \<noteq> y"
shows "\<exists>! z. -1 < Re z \<and> Re z < 1 \<and> is_real z \<and> poincare_distance (of_complex z) (of_complex x) = poincare_distance (of_complex z) (of_complex y)" (is "\<exists>! z. ?R z (of_complex x) (of_complex y)")
proof-
let ?x = "of_complex x" and ?y = "of_complex y"
let ?P = "\<lambda> x y. \<exists>! z. ?R z x y"
have "\<forall> x. -1 < Re x \<and> Re x < 1 \<and> is_real x \<and> of_complex x \<noteq> ?y \<longrightarrow> ?P (of_complex x) ?y" (is "?Q (of_complex y)")
proof (rule wlog_real_zero)
show "?y \<in> unit_disc"
using y
by (simp add: cmod_eq_Re)
next
show "is_real (to_complex ?y)"
using y
by simp
next
show "?Q 0\<^sub>h"
proof (rule allI, rule impI, (erule conjE)+)
fix x
assume x: "-1 < Re x" "Re x < 1" "is_real x"
let ?x = "of_complex x"
assume "?x \<noteq> 0\<^sub>h"
hence "x \<noteq> 0"
by auto
hence "Re x \<noteq> 0"
using x
using complex_neq_0
by auto
have *: "\<forall> a. -1 < a \<and> a < 1 \<longrightarrow>
(poincare_distance (of_complex (cor a)) ?x = poincare_distance (of_complex (cor a)) 0\<^sub>h \<longleftrightarrow>
(Re x) * a * a - 2 * a + Re x = 0)"
proof (rule allI, rule impI)
fix a :: real
assume "-1 < a \<and> a < 1"
hence "of_complex (cor a) \<in> unit_disc"
by auto
moreover
have "(a - Re x)\<^sup>2 / ((1 - a\<^sup>2) * (1 - (Re x)\<^sup>2)) = a\<^sup>2 / (1 - a\<^sup>2) \<longleftrightarrow>
(Re x) * a * a - 2 * a + Re x = 0" (is "?lhs \<longleftrightarrow> ?rhs")
proof-
have "1 - a\<^sup>2 \<noteq> 0"
using \<open>-1 < a \<and> a < 1\<close>
by (metis cancel_comm_monoid_add_class.diff_cancel diff_eq_diff_less less_numeral_extra(4) power2_eq_1_iff right_minus_eq)
hence "?lhs \<longleftrightarrow> (a - Re x)\<^sup>2 / (1 - (Re x)\<^sup>2) = a\<^sup>2"
by (smt divide_cancel_right divide_divide_eq_left mult.commute)
also have "... \<longleftrightarrow> (a - Re x)\<^sup>2 = a\<^sup>2 * (1 - (Re x)\<^sup>2)"
proof-
have "1 - (Re x)\<^sup>2 \<noteq> 0"
using x
by (smt power2_eq_1_iff)
thus ?thesis
by (simp add: divide_eq_eq)
qed
also have "... \<longleftrightarrow> a\<^sup>2 * (Re x)\<^sup>2 - 2*a*Re x + (Re x)\<^sup>2 = 0"
by (simp add: power2_diff field_simps)
also have "... \<longleftrightarrow> Re x * (a\<^sup>2 * Re x - 2 * a + Re x) = 0"
by (simp add: power2_eq_square field_simps)
also have "... \<longleftrightarrow> ?rhs"
using \<open>Re x \<noteq> 0\<close>
by (simp add: mult.commute mult.left_commute power2_eq_square)
finally
show ?thesis
.
qed
moreover
have "arcosh (1 + 2 * ((a - Re x)\<^sup>2 / ((1 - a\<^sup>2) * (1 - (Re x)\<^sup>2)))) = arcosh (1 + 2 * a\<^sup>2 / (1 - a\<^sup>2)) \<longleftrightarrow> ?lhs"
using \<open>-1 < a \<and> a < 1\<close> x mult_left_cancel[of "2::real" "(a - Re x)\<^sup>2 / ((1 - a\<^sup>2) * (1 - (Re x)\<^sup>2))" "a\<^sup>2 / (1 - a\<^sup>2)"]
by (subst arcosh_eq_iff, simp_all add: square_le_1)
ultimately
show "poincare_distance (of_complex (cor a)) (of_complex x) = poincare_distance (of_complex (cor a)) 0\<^sub>h \<longleftrightarrow>
(Re x) * a * a - 2 * a + Re x = 0"
using x
by (auto simp add: poincare_distance_formula cmod_eq_Re)
qed
show "?P ?x 0\<^sub>h"
proof
let ?a = "(1 - sqrt(1 - (Re x)\<^sup>2)) / (Re x)"
let ?b = "(1 + sqrt(1 - (Re x)\<^sup>2)) / (Re x)"
have "is_real ?a"
by simp
moreover
have "1 - (Re x)\<^sup>2 > 0"
using x
by (smt power2_eq_1_iff square_le_1)
have "\<bar>?a\<bar> < 1"
proof (cases "Re x > 0")
case True
have "(1 - Re x)\<^sup>2 < 1 - (Re x)\<^sup>2"
using \<open>Re x > 0\<close> x
by (simp add: power2_eq_square field_simps)
hence "1 - Re x < sqrt (1 - (Re x)\<^sup>2)"
using real_less_rsqrt by fastforce
thus ?thesis
using \<open>1 - (Re x)\<^sup>2 > 0\<close> \<open>Re x > 0\<close>
by simp
next
case False
hence "Re x < 0"
using \<open>Re x \<noteq> 0\<close>
by simp
have "1 + Re x > 0"
using \<open>Re x > -1\<close>
by simp
hence "2*Re x + 2*Re x*Re x < 0"
using \<open>Re x < 0\<close>
by (metis comm_semiring_class.distrib mult.commute mult_2_right mult_less_0_iff one_add_one zero_less_double_add_iff_zero_less_single_add)
hence "(1 + Re x)\<^sup>2 < 1 - (Re x)\<^sup>2"
by (simp add: power2_eq_square field_simps)
hence "1 + Re x < sqrt (1 - (Re x)\<^sup>2)"
using \<open>1 - (Re x)\<^sup>2 > 0\<close>
using real_less_rsqrt by blast
thus ?thesis
using \<open>Re x < 0\<close>
by (simp add: field_simps)
qed
hence "-1 < ?a" "?a < 1"
by linarith+
moreover
have "(Re x) * ?a * ?a - 2 * ?a + Re x = 0"
using \<open>Re x \<noteq> 0\<close> \<open>1 - (Re x)\<^sup>2 > 0\<close>
by (simp add: field_simps power2_eq_square)
ultimately
show "-1 < Re (cor ?a) \<and> Re (cor ?a) < 1 \<and> is_real ?a \<and> poincare_distance (of_complex ?a) (of_complex x) = poincare_distance (of_complex ?a) 0\<^sub>h"
using *
by auto
fix z
assume **: "- 1 < Re z \<and> Re z < 1 \<and> is_real z \<and>
poincare_distance (of_complex z) (of_complex x) = poincare_distance (of_complex z) 0\<^sub>h"
hence "Re x * Re z * Re z - 2 * Re z + Re x = 0"
using *[rule_format, of "Re z"] x
by auto
moreover
have "sqrt (4 - 4 * Re x * Re x) = 2 * sqrt(1 - Re x * Re x)"
proof-
have "sqrt (4 - 4 * Re x * Re x) = sqrt(4 * (1 - Re x * Re x))"
by simp
thus ?thesis
by (simp only: real_sqrt_mult, simp)
qed
moreover
have "(2 - 2 * sqrt (1 - Re x * Re x)) / (2 * Re x) = ?a"
proof-
have "(2 - 2 * sqrt (1 - Re x * Re x)) / (2 * Re x) =
(2 * (1 - sqrt (1 - Re x * Re x))) / (2 * Re x)"
by simp
thus ?thesis
by (subst (asm) mult_divide_mult_cancel_left) (auto simp add: power2_eq_square)
qed
moreover
have "(2 + 2 * sqrt (1 - Re x * Re x)) / (2 * Re x) = ?b"
proof-
have "(2 + 2 * sqrt (1 - Re x * Re x)) / (2 * Re x) =
(2 * (1 + sqrt (1 - Re x * Re x))) / (2 * Re x)"
by simp
thus ?thesis
by (subst (asm) mult_divide_mult_cancel_left) (auto simp add: power2_eq_square)
qed
ultimately
have "Re z = ?a \<or> Re z = ?b"
using discriminant_nonneg[of "Re x" "-2" "Re x" "Re z"] discrim_def[of "Re x" "-2" "Re x"]
using \<open>Re x \<noteq> 0\<close> \<open>-1 < Re x\<close> \<open>Re x < 1\<close> \<open>1 - (Re x)\<^sup>2 > 0\<close>
by (auto simp add:power2_eq_square)
have "\<bar>?b\<bar> > 1"
proof (cases "Re x > 0")
case True
have "(Re x - 1)\<^sup>2 < 1 - (Re x)\<^sup>2"
using \<open>Re x > 0\<close> x
by (simp add: power2_eq_square field_simps)
hence "Re x - 1 < sqrt (1 - (Re x)\<^sup>2)"
using real_less_rsqrt
by simp
thus ?thesis
using \<open>1 - (Re x)\<^sup>2 > 0\<close> \<open>Re x > 0\<close>
by simp
next
case False
hence "Re x < 0"
using \<open>Re x \<noteq> 0\<close>
by simp
have "1 + Re x > 0"
using \<open>Re x > -1\<close>
by simp
hence "2*Re x + 2*Re x*Re x < 0"
using \<open>Re x < 0\<close>
by (metis comm_semiring_class.distrib mult.commute mult_2_right mult_less_0_iff one_add_one zero_less_double_add_iff_zero_less_single_add)
hence "1 - (Re x)\<^sup>2 > (- 1 - (Re x))\<^sup>2"
by (simp add: field_simps power2_eq_square)
hence "sqrt (1 - (Re x)\<^sup>2) > -1 - Re x"
using real_less_rsqrt
by simp
thus ?thesis
using \<open>Re x < 0\<close>
by (simp add: field_simps)
qed
hence "?b < -1 \<or> ?b > 1"
by auto
hence "Re z = ?a"
using \<open>Re z = ?a \<or> Re z = ?b\<close> **
by auto
thus "z = ?a"
using ** complex_of_real_Re
by fastforce
qed
qed
next
fix a u
let ?M = "moebius_pt (blaschke a)"
let ?Mu = "?M u"
assume "u \<in> unit_disc" "is_real a" "cmod a < 1"
assume *: "?Q ?Mu"
show "?Q u"
proof (rule allI, rule impI, (erule conjE)+)
fix x
assume x: "-1 < Re x" "Re x < 1" "is_real x" "of_complex x \<noteq> u"
let ?Mx = "?M (of_complex x)"
have "of_complex x \<in> unit_disc"
using x cmod_eq_Re
by auto
hence "?Mx \<in> unit_disc"
using \<open>is_real a\<close> \<open>cmod a < 1\<close> blaschke_unit_disc_fix[of a]
using unit_disc_fix_discI
by blast
hence "?Mx \<noteq> \<infinity>\<^sub>h"
by auto
moreover
have "of_complex x \<in> circline_set x_axis"
using x
by auto
hence "?Mx \<in> circline_set x_axis"
using blaschke_real_preserve_x_axis[OF \<open>is_real a\<close> \<open>cmod a < 1\<close>, of "of_complex x"]
by auto
hence "-1 < Re (to_complex ?Mx) \<and> Re (to_complex ?Mx) < 1 \<and> is_real (to_complex ?Mx)"
using \<open>?Mx \<noteq> \<infinity>\<^sub>h\<close> \<open>?Mx \<in> unit_disc\<close>
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re)
moreover
have "?Mx \<noteq> ?Mu"
using \<open>of_complex x \<noteq> u\<close>
by simp
ultimately
have "?P ?Mx ?Mu"
using *[rule_format, of "to_complex ?Mx"] \<open>?Mx \<noteq> \<infinity>\<^sub>h\<close>
by simp
then obtain Mz where
"?R Mz ?Mx ?Mu"
by blast
have "of_complex Mz \<in> unit_disc" "of_complex Mz \<in> circline_set x_axis"
using \<open>?R Mz ?Mx ?Mu\<close>
using cmod_eq_Re
by auto
let ?Minv = "- (blaschke a)"
let ?z = "moebius_pt ?Minv (of_complex Mz)"
have "?z \<in> unit_disc"
using \<open>of_complex Mz \<in> unit_disc\<close> \<open>cmod a < 1\<close>
by auto
moreover
have "?z \<in> circline_set x_axis"
using \<open>of_complex Mz \<in> circline_set x_axis\<close>
using blaschke_real_preserve_x_axis \<open>is_real a\<close> \<open>cmod a < 1\<close>
by fastforce
ultimately
have z1: "-1 < Re (to_complex ?z)" "Re (to_complex ?z) < 1" "is_real (to_complex ?z)"
using inf_or_of_complex[of "?z"]
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re)
have z2: "poincare_distance ?z (of_complex x) = poincare_distance ?z u"
using \<open>?R Mz ?Mx ?Mu\<close> \<open>cmod a < 1\<close> \<open>?z \<in> unit_disc\<close> \<open>of_complex x \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close>
by (metis blaschke_preserve_distance_formula blaschke_unit_disc_fix moebius_pt_comp_inv_right poincare_distance_formula uminus_moebius_def unit_disc_fix_discI unit_disc_iff_cmod_lt_1)
show "?P (of_complex x) u"
proof
show "?R (to_complex ?z) (of_complex x) u"
using z1 z2 \<open>?z \<in> unit_disc\<close> inf_or_of_complex[of ?z]
by auto
next
fix z'
assume "?R z' (of_complex x) u"
hence "of_complex z' \<in> unit_disc" "of_complex z' \<in> circline_set x_axis"
by (auto simp add: cmod_eq_Re)
let ?Mz' = "?M (of_complex z')"
have "?Mz' \<in> unit_disc" "?Mz' \<in> circline_set x_axis"
using \<open>of_complex z' \<in> unit_disc\<close> \<open>of_complex z' \<in> circline_set x_axis\<close> \<open>cmod a < 1\<close> \<open>is_real a\<close>
using blaschke_unit_disc_fix unit_disc_fix_discI
using blaschke_real_preserve_x_axis circline_set_x_axis
by blast+
hence "-1 < Re (to_complex ?Mz')" "Re (to_complex ?Mz') < 1" "is_real (to_complex ?Mz')"
unfolding circline_set_x_axis
by (auto simp add: cmod_eq_Re)
moreover
have "poincare_distance ?Mz' ?Mx = poincare_distance ?Mz' ?Mu"
using \<open>?R z' (of_complex x) u\<close>
using \<open>cmod a < 1\<close> \<open>of_complex x \<in> unit_disc\<close> \<open>of_complex z' \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close>
by auto
ultimately
have "?R (to_complex ?Mz') ?Mx ?Mu"
using \<open>?Mz' \<in> unit_disc\<close> inf_or_of_complex[of ?Mz']
by auto
hence "?Mz' = of_complex Mz"
using \<open>?P ?Mx ?Mu\<close> \<open>?R Mz ?Mx ?Mu\<close>
by (metis \<open>moebius_pt (blaschke a) (of_complex z') \<in> unit_disc\<close> \<open>of_complex Mz \<in> unit_disc\<close> to_complex_of_complex unit_disc_to_complex_inj)
thus "z' = to_complex ?z"
by (simp add: moebius_pt_invert)
qed
qed
qed
thus ?thesis
using assms
by (metis to_complex_of_complex)
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Triangle inequality\<close>
(* ------------------------------------------------------------------ *)
lemma poincare_distance_formula_zero_sum:
assumes "u \<in> unit_disc" and "v \<in> unit_disc"
shows "poincare_distance u 0\<^sub>h + poincare_distance 0\<^sub>h v =
(let u' = cmod (to_complex u); v' = cmod (to_complex v)
in arcosh (((1 + u'\<^sup>2) * (1 + v'\<^sup>2) + 4 * u' * v') / ((1 - u'\<^sup>2) * (1 - v'\<^sup>2))))"
proof-
obtain u' v' where uv: "u' = to_complex u" "v' = to_complex v"
by auto
have uv': "u = of_complex u'" "v = of_complex v'"
using uv assms inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
let ?u' = "cmod u'" and ?v' = "cmod v'"
have disc: "?u'\<^sup>2 < 1" "?v'\<^sup>2 < 1"
using unit_disc_cmod_square_lt_1[OF \<open>u \<in> unit_disc\<close>]
using unit_disc_cmod_square_lt_1[OF \<open>v \<in> unit_disc\<close>] uv
by auto
thm arcosh_add
have "arcosh (1 + 2 * ?u'\<^sup>2 / (1 - ?u'\<^sup>2)) + arcosh (1 + 2 * ?v'\<^sup>2 / (1 - ?v'\<^sup>2)) =
arcosh (((1 + ?u'\<^sup>2) * (1 + ?v'\<^sup>2) + 4 * ?u' * ?v') / ((1 - ?u'\<^sup>2) * (1 - ?v'\<^sup>2)))" (is "arcosh ?ll + arcosh ?rr = arcosh ?r")
proof (subst arcosh_add)
show "?ll \<ge> 1" "?rr \<ge> 1"
using disc
by auto
next
show "arcosh ((1 + 2 * ?u'\<^sup>2 / (1 - ?u'\<^sup>2)) * (1 + 2 * ?v'\<^sup>2 / (1 - ?v'\<^sup>2)) +
sqrt (((1 + 2 * ?u'\<^sup>2 / (1 - ?u'\<^sup>2))\<^sup>2 - 1) * ((1 + 2 * ?v'\<^sup>2 / (1 - ?v'\<^sup>2))\<^sup>2 - 1))) =
arcosh ?r" (is "arcosh ?l = _")
proof-
have "1 + 2 * ?u'\<^sup>2 / (1 - ?u'\<^sup>2) = (1 + ?u'\<^sup>2) / (1 - ?u'\<^sup>2)"
using disc
by (subst add_divide_eq_iff, simp_all)
moreover
have "1 + 2 * ?v'\<^sup>2 / (1 - ?v'\<^sup>2) = (1 + ?v'\<^sup>2) / (1 - ?v'\<^sup>2)"
using disc
by (subst add_divide_eq_iff, simp_all)
moreover
have "sqrt (((1 + 2 * ?u'\<^sup>2 / (1 - ?u'\<^sup>2))\<^sup>2 - 1) * ((1 + 2 * ?v'\<^sup>2 / (1 - ?v'\<^sup>2))\<^sup>2 - 1)) =
(4 * ?u' * ?v') / ((1 - ?u'\<^sup>2) * (1 - ?v'\<^sup>2))" (is "sqrt ?s = ?t")
proof-
have "?s = ?t\<^sup>2"
using disc
apply (subst add_divide_eq_iff, simp)+
apply (subst power_divide)+
apply simp
apply (subst divide_diff_eq_iff, simp)+
apply (simp add: power2_eq_square field_simps)
done
thus ?thesis
using disc
by simp
qed
ultimately
have "?l = ?r"
using disc
by simp (subst add_divide_distrib, simp)
thus ?thesis
by simp
qed
qed
thus ?thesis
using uv' assms
using poincare_distance_formula
by (simp add: Let_def)
qed
lemma poincare_distance_triangle_inequality:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "w \<in> unit_disc"
shows "poincare_distance u v + poincare_distance v w \<ge> poincare_distance u w" (is "?P' u v w")
proof-
have "\<forall> w. w \<in> unit_disc \<longrightarrow> ?P' u v w" (is "?P v u")
proof (rule wlog_x_axis[where P="?P"])
fix x
assume "is_real x" "0 \<le> Re x" "Re x < 1"
hence "of_complex x \<in> unit_disc"
by (simp add: cmod_eq_Re)
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix w
assume "w \<in> unit_disc"
then obtain w' where w: "w = of_complex w'"
using inf_or_of_complex[of w]
by auto
let ?x = "cmod x" and ?w = "cmod w'" and ?xw = "cmod (x - w')"
have disc: "?x\<^sup>2 < 1" "?w\<^sup>2 < 1"
using unit_disc_cmod_square_lt_1[OF \<open>of_complex x \<in> unit_disc\<close>]
using unit_disc_cmod_square_lt_1[OF \<open>w \<in> unit_disc\<close>] w
by auto
have "poincare_distance (of_complex x) 0\<^sub>h + poincare_distance 0\<^sub>h w =
arcosh (((1 + ?x\<^sup>2) * (1 + ?w\<^sup>2) + 4 * ?x * ?w) / ((1 - ?x\<^sup>2) * (1 - ?w\<^sup>2)))" (is "_ = arcosh ?r1")
using poincare_distance_formula_zero_sum[OF \<open>of_complex x \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>] w
by (simp add: Let_def)
moreover
have "poincare_distance (of_complex x) (of_complex w') =
arcosh (((1 - ?x\<^sup>2) * (1 - ?w\<^sup>2) + 2 * ?xw\<^sup>2) / ((1 - ?x\<^sup>2) * (1 - ?w\<^sup>2)))" (is "_ = arcosh ?r2")
using disc
using poincare_distance_formula[OF \<open>of_complex x \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close>] w
by (subst add_divide_distrib) simp
moreover
have *: "(1 - ?x\<^sup>2) * (1 - ?w\<^sup>2) + 2 * ?xw\<^sup>2 \<le> (1 + ?x\<^sup>2) * (1 + ?w\<^sup>2) + 4 * ?x * ?w"
proof-
have "(cmod (x - w'))\<^sup>2 \<le> (cmod x + cmod w')\<^sup>2"
using norm_triangle_ineq4[of x w']
by (simp add: power_mono)
thus ?thesis
by (simp add: field_simps power2_sum)
qed
have "arcosh ?r1 \<ge> arcosh ?r2"
proof (subst arcosh_mono)
show "?r1 \<ge> 1"
using disc
by (smt "*" le_divide_eq_1_pos mult_pos_pos zero_le_power2)
next
show "?r2 \<ge> 1"
using disc
by simp
next
show "?r1 \<ge> ?r2"
using disc
using *
by (subst divide_right_mono, simp_all)
qed
ultimately
show "poincare_distance (of_complex x) w \<le> poincare_distance (of_complex x) 0\<^sub>h + poincare_distance 0\<^sub>h w"
using \<open>of_complex x \<in> unit_disc\<close> \<open>w \<in> unit_disc\<close> w
using poincare_distance_formula
by simp
qed
next
show "v \<in> unit_disc" "u \<in> unit_disc"
by fact+
next
fix M u v
assume *: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
assume **: "?P (moebius_pt M u) (moebius_pt M v)"
show "?P u v"
proof safe
fix w
assume "w \<in> unit_disc"
thus "?P' v u w"
using * **[rule_format, of "moebius_pt M w"]
by simp
qed
qed
thus ?thesis
using assms
by auto
qed
end
\ No newline at end of file
diff --git a/thys/Poincare_Disc/Poincare_Lines.thy b/thys/Poincare_Disc/Poincare_Lines.thy
--- a/thys/Poincare_Disc/Poincare_Lines.thy
+++ b/thys/Poincare_Disc/Poincare_Lines.thy
@@ -1,1878 +1,1878 @@
(* ------------------------------------------------------------------ *)
section \<open>H-lines in the Poincar\'e model\<close>
(* ------------------------------------------------------------------ *)
theory Poincare_Lines
imports Complex_Geometry.Unit_Circle_Preserving_Moebius Complex_Geometry.Circlines_Angle
begin
(* ------------------------------------------------------------------ *)
subsection \<open>Definition and basic properties of h-lines\<close>
(* ------------------------------------------------------------------ *)
text \<open>H-lines in the Poincar\'e model are either line segments passing trough the origin or
segments (within the unit disc) of circles that are perpendicular to the unit circle. Algebraically
these are circlines that are represented by Hermitean matrices of
the form
$$H = \left(
\begin{array}{cc}
A & B\\
\overline{B} & A
\end{array}
\right),$$
for $A \in \mathbb{R}$, and $B \in \mathbb{C}$, and $|B|^2 > A^2$,
where the circline equation is the usual one: $z^*Hz = 0$, for homogenous coordinates $z$.\<close>
definition is_poincare_line_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "is_poincare_line_cmat H \<longleftrightarrow>
(let (A, B, C, D) = H
in hermitean (A, B, C, D) \<and> A = D \<and> (cmod B)\<^sup>2 > (cmod A)\<^sup>2)"
lift_definition is_poincare_line_clmat :: "circline_mat \<Rightarrow> bool" is is_poincare_line_cmat
done
text \<open>We introduce the predicate that checks if a given complex matrix is a matrix of a h-line in
the Poincar\'e model, and then by means of the lifting package lift it to the type of non-zero
Hermitean matrices, and then to circlines (that are equivalence classes of such matrices).\<close>
lift_definition is_poincare_line :: "circline \<Rightarrow> bool" is is_poincare_line_clmat
proof (transfer, transfer)
fix H1 H2 :: complex_mat
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" "hermitean H2 \<and> H2 \<noteq> mat_zero"
assume "circline_eq_cmat H1 H2"
thus "is_poincare_line_cmat H1 \<longleftrightarrow> is_poincare_line_cmat H2"
using hh
by (cases H1, cases H2) (auto simp add: power_mult_distrib)
qed
lemma is_poincare_line_mk_circline:
assumes "(A, B, C, D) \<in> hermitean_nonzero"
shows "is_poincare_line (mk_circline A B C D) \<longleftrightarrow> (cmod B)\<^sup>2 > (cmod A)\<^sup>2 \<and> A = D"
using assms
by (transfer, transfer, auto simp add: Let_def)
text\<open>Abstract characterisation of @{term is_poincare_line} predicate: H-lines in the Poincar\'e
model are real circlines (circlines with the negative determinant) perpendicular to the unit
circle.\<close>
lemma is_poincare_line_iff:
shows "is_poincare_line H \<longleftrightarrow> circline_type H = -1 \<and> perpendicular H unit_circle"
unfolding perpendicular_def
proof (simp, transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
have **: "is_real A" "is_real D" "C = cnj B"
using hh * hermitean_elems
by auto
hence "(Re A = Re D \<and> cmod A * cmod A < cmod B * cmod B) =
(Re A * Re D < Re B * Re B + Im B * Im B \<and> (Re D = Re A \<or> Re A * Re D = Re B * Re B + Im B * Im B))"
using *
by (smt cmod_power2 power2_eq_square zero_power2)+
thus "is_poincare_line_cmat H \<longleftrightarrow>
circline_type_cmat H = - 1 \<and> cos_angle_cmat (of_circline_cmat H) unit_circle_cmat = 0"
using * **
by (auto simp add: sgn_1_neg complex_eq_if_Re_eq cmod_square power2_eq_square simp del: pos_oriented_cmat_def)
qed
text\<open>The @{term x_axis} is an h-line.\<close>
lemma is_poincare_line_x_axis [simp]:
shows "is_poincare_line x_axis"
by (transfer, transfer) (auto simp add: hermitean_def mat_adj_def mat_cnj_def)
text\<open>The @{term unit_circle} is not an h-line.\<close>
lemma not_is_poincare_line_unit_circle [simp]:
shows "\<not> is_poincare_line unit_circle"
by (transfer, transfer, simp)
(* ------------------------------------------------------------------ *)
subsubsection \<open>Collinear points\<close>
(* ------------------------------------------------------------------ *)
text\<open>Points are collinear if they all belong to an h-line. \<close>
definition poincare_collinear :: "complex_homo set \<Rightarrow> bool" where
"poincare_collinear S \<longleftrightarrow> (\<exists> p. is_poincare_line p \<and> S \<subseteq> circline_set p)"
(* ------------------------------------------------------------------ *)
subsubsection \<open>H-lines and inversion\<close>
(* ------------------------------------------------------------------ *)
text\<open>Every h-line in the Poincar\'e model contains the inverse (wrt.~the unit circle) of each of its
points (note that at most one of them belongs to the unit disc).\<close>
lemma is_poincare_line_inverse_point:
assumes "is_poincare_line H" "u \<in> circline_set H"
shows "inversion u \<in> circline_set H"
using assms
unfolding is_poincare_line_iff circline_set_def perpendicular_def inversion_def
apply simp
proof (transfer, transfer)
fix u H
assume hh: "hermitean H \<and> H \<noteq> mat_zero" "u \<noteq> vec_zero" and
aa: "circline_type_cmat H = - 1 \<and> cos_angle_cmat (of_circline_cmat H) unit_circle_cmat = 0" "on_circline_cmat_cvec H u"
obtain A B C D u1 u2 where *: "H = (A, B, C, D)" "u = (u1, u2)"
by (cases H, cases u, auto)
have "is_real A" "is_real D" "C = cnj B"
using * hh hermitean_elems
by auto
moreover
have "A = D"
using aa(1) * \<open>is_real A\<close> \<open>is_real D\<close>
by (auto simp del: pos_oriented_cmat_def simp add: complex.expand split: if_split_asm)
thus "on_circline_cmat_cvec H (conjugate_cvec (reciprocal_cvec u))"
using aa(2) *
by (simp add: vec_cnj_def field_simps)
qed
text\<open>Every h-line in the Poincar\'e model and is invariant under unit circle inversion.\<close>
lemma circline_inversion_poincare_line:
assumes "is_poincare_line H"
shows "circline_inversion H = H"
proof-
obtain u v w where *: "u \<noteq> v" "v \<noteq> w" "u \<noteq> w" "{u, v, w} \<subseteq> circline_set H"
using assms is_poincare_line_iff[of H]
using circline_type_neg_card_gt3[of H]
by auto
hence "{inversion u, inversion v, inversion w} \<subseteq> circline_set (circline_inversion H)"
"{inversion u, inversion v, inversion w} \<subseteq> circline_set H"
using is_poincare_line_inverse_point[OF assms]
by auto
thus ?thesis
using * unique_circline_set[of "inversion u" "inversion v" "inversion w"]
by (metis insert_subset inversion_involution)
qed
(* ------------------------------------------------------------------ *)
subsubsection \<open>Classification of h-lines into Euclidean segments and circles\<close>
(* ------------------------------------------------------------------ *)
text\<open>If an h-line contains zero, than it also contains infinity (the inverse point of zero) and is by
definition an Euclidean line.\<close>
lemma is_poincare_line_trough_zero_trough_infty [simp]:
assumes "is_poincare_line l" and "0\<^sub>h \<in> circline_set l"
shows "\<infinity>\<^sub>h \<in> circline_set l"
using is_poincare_line_inverse_point[OF assms]
by simp
lemma is_poincare_line_trough_zero_is_line:
assumes "is_poincare_line l" and "0\<^sub>h \<in> circline_set l"
shows "is_line l"
using assms
using inf_in_circline_set is_poincare_line_trough_zero_trough_infty
by blast
text\<open>If an h-line does not contain zero, than it also does not contain infinity (the inverse point of
zero) and is by definition an Euclidean circle.\<close>
lemma is_poincare_line_not_trough_zero_not_trough_infty [simp]:
assumes "is_poincare_line l"
assumes "0\<^sub>h \<notin> circline_set l"
shows "\<infinity>\<^sub>h \<notin> circline_set l"
using assms
using is_poincare_line_inverse_point[OF assms(1), of "\<infinity>\<^sub>h"]
by auto
lemma is_poincare_line_not_trough_zero_is_circle:
assumes "is_poincare_line l" "0\<^sub>h \<notin> circline_set l"
shows "is_circle l"
using assms
using inf_in_circline_set is_poincare_line_not_trough_zero_not_trough_infty
by auto
(* ------------------------------------------------------------------ *)
subsubsection\<open>Points on h-line\<close>
(* ------------------------------------------------------------------ *)
text\<open>Each h-line in the Poincar\'e model contains at least two different points within the unit
disc.\<close>
text\<open>First we prove an auxiliary lemma.\<close>
lemma ex_is_poincare_line_points':
assumes i12: "i1 \<in> circline_set H \<inter> unit_circle_set"
"i2 \<in> circline_set H \<inter> unit_circle_set"
"i1 \<noteq> i2"
assumes a: "a \<in> circline_set H" "a \<notin> unit_circle_set"
shows "\<exists> b. b \<noteq> i1 \<and> b \<noteq> i2 \<and> b \<noteq> a \<and> b \<noteq> inversion a \<and> b \<in> circline_set H"
proof-
have "inversion a \<notin> unit_circle_set"
using \<open>a \<notin> unit_circle_set\<close>
unfolding unit_circle_set_def circline_set_def
by (metis inversion_id_iff_on_unit_circle inversion_involution mem_Collect_eq)
have "a \<noteq> inversion a"
using \<open>a \<notin> unit_circle_set\<close> inversion_id_iff_on_unit_circle[of a]
unfolding unit_circle_set_def circline_set_def
by auto
have "a \<noteq> i1" "a \<noteq> i2" "inversion a \<noteq> i1" "inversion a \<noteq> i2"
using assms \<open>inversion a \<notin> unit_circle_set\<close>
by auto
then obtain b where cr2: "cross_ratio b i1 a i2 = of_complex 2"
using \<open>i1 \<noteq> i2\<close>
using ex_cross_ratio[of i1 a i2]
by blast
have distinct_b: "b \<noteq> i1" "b \<noteq> i2" "b \<noteq> a"
using \<open>i1 \<noteq> i2\<close> \<open>a \<noteq> i1\<close> \<open>a \<noteq> i2\<close>
using ex1_cross_ratio[of i1 a i2]
using cross_ratio_0[of i1 a i2] cross_ratio_1[of i1 a i2] cross_ratio_inf[of i1 i2 a]
using cr2
by auto
hence "b \<in> circline_set H"
using assms four_points_on_circline_iff_cross_ratio_real[of b i1 a i2] cr2
using unique_circline_set[of i1 i2 a]
by auto
moreover
have "b \<noteq> inversion a"
proof (rule ccontr)
assume *: "\<not> ?thesis"
have "inversion i1 = i1" "inversion i2 = i2"
using i12
unfolding unit_circle_set_def
by auto
hence "cross_ratio (inversion a) i1 a i2 = cross_ratio a i1 (inversion a) i2"
using * cross_ratio_inversion[of i1 a i2 b] \<open>a \<noteq> i1\<close> \<open>a \<noteq> i2\<close> \<open>i1 \<noteq> i2\<close> \<open>b \<noteq> i1\<close>
using four_points_on_circline_iff_cross_ratio_real[of b i1 a i2]
using i12 distinct_b conjugate_id_iff[of "cross_ratio b i1 a i2"]
using i12 a \<open>b \<in> circline_set H\<close>
by auto
hence "cross_ratio (inversion a) i1 a i2 \<noteq> of_complex 2"
using cross_ratio_commute_13[of "inversion a" i1 a i2]
using reciprocal_id_iff
using of_complex_inj
by force
thus False
using * cr2
by simp
qed
ultimately
show ?thesis
using assms \<open>b \<noteq> i1\<close> \<open>b \<noteq> i2\<close> \<open>b \<noteq> a\<close>
by auto
qed
text\<open>Now we can prove the statement.\<close>
lemma ex_is_poincare_line_points:
assumes "is_poincare_line H"
shows "\<exists> u v. u \<in> unit_disc \<and> v \<in> unit_disc \<and> u \<noteq> v \<and> {u, v} \<subseteq> circline_set H"
proof-
obtain u v w where *: "u \<noteq> v" "v \<noteq> w" "u \<noteq> w" "{u, v, w} \<subseteq> circline_set H"
using assms is_poincare_line_iff[of H]
using circline_type_neg_card_gt3[of H]
by auto
have "\<not> {u, v, w} \<subseteq> unit_circle_set"
using unique_circline_set[of u v w] *
by (metis assms insert_subset not_is_poincare_line_unit_circle unit_circle_set_def)
hence "H \<noteq> unit_circle"
unfolding unit_circle_set_def
using *
by auto
show ?thesis
proof (cases "(u \<in> unit_disc \<and> v \<in> unit_disc) \<or>
(u \<in> unit_disc \<and> w \<in> unit_disc) \<or>
(v \<in> unit_disc \<and> w \<in> unit_disc)")
case True
thus ?thesis
using *
by auto
next
case False
have "\<exists> a b. a \<noteq> b \<and> a \<noteq> inversion b \<and> a \<in> circline_set H \<and> b \<in> circline_set H \<and> a \<notin> unit_circle_set \<and> b \<notin> unit_circle_set"
proof (cases "(u \<in> unit_circle_set \<and> v \<in> unit_circle_set) \<or>
(u \<in> unit_circle_set \<and> w \<in> unit_circle_set) \<or>
(v \<in> unit_circle_set \<and> w \<in> unit_circle_set)")
case True
then obtain i1 i2 a where *:
"i1 \<in> unit_circle_set \<inter> circline_set H" "i2 \<in> unit_circle_set \<inter> circline_set H"
"a \<in> circline_set H" "a \<notin> unit_circle_set"
"i1 \<noteq> i2" "i1 \<noteq> a" "i2 \<noteq> a"
using * \<open>\<not> {u, v, w} \<subseteq> unit_circle_set\<close>
by auto
then obtain b where "b \<in> circline_set H" "b \<noteq> i1" "b \<noteq> i2" "b \<noteq> a" "b \<noteq> inversion a"
using ex_is_poincare_line_points'[of i1 H i2 a]
by blast
hence "b \<notin> unit_circle_set"
using * \<open>H \<noteq> unit_circle\<close> unique_circline_set[of i1 i2 b]
unfolding unit_circle_set_def
by auto
thus ?thesis
using * \<open>b \<in> circline_set H\<close> \<open>b \<noteq> a\<close> \<open>b \<noteq> inversion a\<close>
by auto
next
case False
then obtain f g h where
*: "f \<noteq> g" "f \<in> circline_set H" "f \<notin> unit_circle_set"
"g \<in> circline_set H" "g \<notin> unit_circle_set"
"h \<in> circline_set H" "h \<noteq> f" "h \<noteq> g"
using *
by auto
show ?thesis
proof (cases "f = inversion g")
case False
thus ?thesis
using *
by auto
next
case True
show ?thesis
proof (cases "h \<in> unit_circle_set")
case False
thus ?thesis
using * \<open>f = inversion g\<close>
by auto
next
case True
obtain m where cr2: "cross_ratio m h f g = of_complex 2"
using ex_cross_ratio[of h f g] * \<open>f \<noteq> g\<close> \<open>h \<noteq> f\<close> \<open>h \<noteq> g\<close>
by auto
hence "m \<noteq> h" "m \<noteq> f" "m \<noteq> g"
using \<open>h \<noteq> f\<close> \<open>h \<noteq> g\<close> \<open>f \<noteq> g\<close>
using ex1_cross_ratio[of h f g]
using cross_ratio_0[of h f g] cross_ratio_1[of h f g] cross_ratio_inf[of h g f]
using cr2
by auto
hence "m \<in> circline_set H"
using four_points_on_circline_iff_cross_ratio_real[of m h f g] cr2
using \<open>h \<noteq> f\<close> \<open>h \<noteq> g\<close> \<open>f \<noteq> g\<close> *
using unique_circline_set[of h f g]
by auto
show ?thesis
proof (cases "m \<in> unit_circle_set")
case False
thus ?thesis
using \<open>m \<noteq> f\<close> \<open>m \<noteq> g\<close> \<open>f = inversion g\<close> * \<open>m \<in> circline_set H\<close>
by auto
next
case True
then obtain n where "n \<noteq> h" "n \<noteq> m" "n \<noteq> f" "n \<noteq> inversion f" "n \<in> circline_set H"
using ex_is_poincare_line_points'[of h H m f] * \<open>m \<in> circline_set H\<close> \<open>h \<in> unit_circle_set\<close> \<open>m \<noteq> h\<close>
by auto
hence "n \<notin> unit_circle_set"
using * \<open>H \<noteq> unit_circle\<close> unique_circline_set[of m n h]
using \<open>m \<noteq> h\<close> \<open>m \<in> unit_circle_set\<close> \<open>h \<in> unit_circle_set\<close> \<open>m \<in> circline_set H\<close>
unfolding unit_circle_set_def
by auto
thus ?thesis
using * \<open>n \<in> circline_set H\<close> \<open>n \<noteq> f\<close> \<open>n \<noteq> inversion f\<close>
by auto
qed
qed
qed
qed
then obtain a b where ab: "a \<noteq> b" "a \<noteq> inversion b" "a \<in> circline_set H" "b \<in> circline_set H" "a \<notin> unit_circle_set" "b \<notin> unit_circle_set"
by blast
have "\<forall> x. x \<in> circline_set H \<and> x \<notin> unit_circle_set \<longrightarrow> (\<exists> x'. x' \<in> circline_set H \<inter> unit_disc \<and> (x' = x \<or> x' = inversion x))"
proof safe
fix x
assume x: "x \<in> circline_set H" "x \<notin> unit_circle_set"
show "\<exists> x'. x' \<in> circline_set H \<inter> unit_disc \<and> (x' = x \<or> x' = inversion x)"
proof (cases "x \<in> unit_disc")
case True
thus ?thesis
using x
by auto
next
case False
hence "x \<in> unit_disc_compl"
using x in_on_out_univ[of "ounit_circle"]
unfolding unit_circle_set_def unit_disc_def unit_disc_compl_def
by auto
hence "inversion x \<in> unit_disc"
using inversion_unit_disc_compl
by blast
thus ?thesis
using is_poincare_line_inverse_point[OF assms, of x] x
by auto
qed
qed
then obtain a' b' where
*: "a' \<in> circline_set H" "a' \<in> unit_disc" "b' \<in> circline_set H" "b' \<in> unit_disc" and
**: "a' = a \<or> a' = inversion a" "b' = b \<or> b' = inversion b"
using ab
by blast
have "a' \<noteq> b'"
using \<open>a \<noteq> b\<close> \<open>a \<noteq> inversion b\<close> ** *
by (metis inversion_involution)
thus ?thesis
using *
by auto
qed
qed
(* ------------------------------------------------------------------ *)
subsubsection \<open>H-line uniqueness\<close>
(* ------------------------------------------------------------------ *)
text\<open>There is no more than one h-line that contains two different h-points (in the disc).\<close>
lemma unique_is_poincare_line:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes pl: "is_poincare_line l1" "is_poincare_line l2"
assumes on_l: "{u, v} \<subseteq> circline_set l1 \<inter> circline_set l2"
shows "l1 = l2"
proof-
have "u \<noteq> inversion u" "v \<noteq> inversion u"
using in_disc
using inversion_noteq_unit_disc[of u v]
using inversion_noteq_unit_disc[of u u]
by auto
thus ?thesis
using on_l
using unique_circline_set[of u "inversion u" "v"] \<open>u \<noteq> v\<close>
using is_poincare_line_inverse_point[of l1 u]
using is_poincare_line_inverse_point[of l2 u]
using pl
by auto
qed
text\<open>For the rest of our formalization it is often useful to consider points on h-lines that are not
within the unit disc. Many lemmas in the rest of this section will have such generalizations.\<close>
text\<open>There is no more than one h-line that contains two different and not mutually inverse points
(not necessary in the unit disc).\<close>
lemma unique_is_poincare_line_general:
assumes different: "u \<noteq> v" "u \<noteq> inversion v"
assumes pl: "is_poincare_line l1" "is_poincare_line l2"
assumes on_l: "{u, v} \<subseteq> circline_set l1 \<inter> circline_set l2"
shows "l1 = l2"
proof (cases "u \<noteq> inversion u")
case True
thus ?thesis
using unique_circline_set[of u "inversion u" "v"]
using assms
using is_poincare_line_inverse_point by force
next
case False
show ?thesis
proof (cases "v \<noteq> inversion v")
case True
thus ?thesis
using unique_circline_set[of u "inversion v" "v"]
using assms
using is_poincare_line_inverse_point by force
next
case False
have "on_circline unit_circle u" "on_circline unit_circle v"
- using `\<not> u \<noteq> inversion u` `\<not> v \<noteq> inversion v`
+ using \<open>\<not> u \<noteq> inversion u\<close> \<open>\<not> v \<noteq> inversion v\<close>
using inversion_id_iff_on_unit_circle
by fastforce+
thus ?thesis
- using pl on_l `u \<noteq> v`
+ using pl on_l \<open>u \<noteq> v\<close>
unfolding circline_set_def
apply simp
proof (transfer, transfer, safe)
fix u1 u2 v1 v2 A1 B1 C1 D1 A2 B2 C2 D2 :: complex
let ?u = "(u1, u2)" and ?v = "(v1, v2)" and ?H1 = "(A1, B1, C1, D1)" and ?H2 = "(A2, B2, C2, D2)"
assume *: "?u \<noteq> vec_zero" "?v \<noteq> vec_zero"
"on_circline_cmat_cvec unit_circle_cmat ?u" "on_circline_cmat_cvec unit_circle_cmat ?v"
"is_poincare_line_cmat ?H1" "is_poincare_line_cmat ?H2"
"hermitean ?H1" "?H1 \<noteq> mat_zero" "hermitean ?H2" "?H2 \<noteq> mat_zero"
"on_circline_cmat_cvec ?H1 ?u" "on_circline_cmat_cvec ?H1 ?v"
"on_circline_cmat_cvec ?H2 ?u" "on_circline_cmat_cvec ?H2 ?v"
"\<not> (u1, u2) \<approx>\<^sub>v (v1, v2)"
have **: "A1 = D1" "A2 = D2" "C1 = cnj B1" "C2 = cnj B2" "is_real A1" "is_real A2"
- using `is_poincare_line_cmat ?H1` `is_poincare_line_cmat ?H2`
- using `hermitean ?H1` `?H1 \<noteq> mat_zero` `hermitean ?H2` `?H2 \<noteq> mat_zero`
+ using \<open>is_poincare_line_cmat ?H1\<close> \<open>is_poincare_line_cmat ?H2\<close>
+ using \<open>hermitean ?H1\<close> \<open>?H1 \<noteq> mat_zero\<close> \<open>hermitean ?H2\<close> \<open>?H2 \<noteq> mat_zero\<close>
using hermitean_elems
by auto
have uv: "u1 \<noteq> 0" "u2 \<noteq> 0" "v1 \<noteq> 0" "v2 \<noteq> 0"
using *(1-4)
by (auto simp add: vec_cnj_def)
have u: "cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) = 1"
- using `on_circline_cmat_cvec unit_circle_cmat ?u` uv
+ using \<open>on_circline_cmat_cvec unit_circle_cmat ?u\<close> uv
apply (subst cor_add[symmetric])
apply (subst complex_mult_cnj[symmetric])
apply (simp add: vec_cnj_def mult.commute)
done
have v: "cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) = 1"
- using `on_circline_cmat_cvec unit_circle_cmat ?v` uv
+ using \<open>on_circline_cmat_cvec unit_circle_cmat ?v\<close> uv
apply (subst cor_add[symmetric])
apply (subst complex_mult_cnj[symmetric])
apply (simp add: vec_cnj_def mult.commute)
done
have
"A1 * (cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) + 1) + cor (Re B1) * cor(2 * Re (u1/u2)) + cor (Im B1) * cor(2 * Im (u1/u2)) = 0"
"A2 * (cor ((Re (u1/u2))\<^sup>2) + cor ((Im (u1/u2))\<^sup>2) + 1) + cor (Re B2) * cor(2 * Re (u1/u2)) + cor (Im B2) * cor(2 * Im (u1/u2)) = 0"
"A1 * (cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) + 1) + cor (Re B1) * cor(2 * Re (v1/v2)) + cor (Im B1) * cor(2 * Im (v1/v2)) = 0"
"A2 * (cor ((Re (v1/v2))\<^sup>2) + cor ((Im (v1/v2))\<^sup>2) + 1) + cor (Re B2) * cor(2 * Re (v1/v2)) + cor (Im B2) * cor(2 * Im (v1/v2)) = 0"
using circline_equation_quadratic_equation[of A1 "u1/u2" B1 D1 "Re (u1/u2)" "Im (u1 / u2)" "Re B1" "Im B1"]
using circline_equation_quadratic_equation[of A2 "u1/u2" B2 D2 "Re (u1/u2)" "Im (u1 / u2)" "Re B2" "Im B2"]
using circline_equation_quadratic_equation[of A1 "v1/v2" B1 D1 "Re (v1/v2)" "Im (v1 / v2)" "Re B1" "Im B1"]
using circline_equation_quadratic_equation[of A2 "v1/v2" B2 D2 "Re (v1/v2)" "Im (v1 / v2)" "Re B2" "Im B2"]
- using `on_circline_cmat_cvec ?H1 ?u` `on_circline_cmat_cvec ?H2 ?u`
- using `on_circline_cmat_cvec ?H1 ?v` `on_circline_cmat_cvec ?H2 ?v`
+ using \<open>on_circline_cmat_cvec ?H1 ?u\<close> \<open>on_circline_cmat_cvec ?H2 ?u\<close>
+ using \<open>on_circline_cmat_cvec ?H1 ?v\<close> \<open>on_circline_cmat_cvec ?H2 ?v\<close>
using ** uv
by (simp_all add: vec_cnj_def field_simps)
hence
"A1 + cor (Re B1) * cor(Re (u1/u2)) + cor (Im B1) * cor(Im (u1/u2)) = 0"
"A1 + cor (Re B1) * cor(Re (v1/v2)) + cor (Im B1) * cor(Im (v1/v2)) = 0"
"A2 + cor (Re B2) * cor(Re (u1/u2)) + cor (Im B2) * cor(Im (u1/u2)) = 0"
"A2 + cor (Re B2) * cor(Re (v1/v2)) + cor (Im B2) * cor(Im (v1/v2)) = 0"
using u v
by simp_all algebra+
hence
"cor (Re A1 + Re B1 * Re (u1/u2) + Im B1 * Im (u1/u2)) = 0"
"cor (Re A2 + Re B2 * Re (u1/u2) + Im B2 * Im (u1/u2)) = 0"
"cor (Re A1 + Re B1 * Re (v1/v2) + Im B1 * Im (v1/v2)) = 0"
"cor (Re A2 + Re B2 * Re (v1/v2) + Im B2 * Im (v1/v2)) = 0"
- using `is_real A1` `is_real A2`
+ using \<open>is_real A1\<close> \<open>is_real A2\<close>
by simp_all
hence
"Re A1 + Re B1 * Re (u1/u2) + Im B1 * Im (u1/u2) = 0"
"Re A1 + Re B1 * Re (v1/v2) + Im B1 * Im (v1/v2) = 0"
"Re A2 + Re B2 * Re (u1/u2) + Im B2 * Im (u1/u2) = 0"
"Re A2 + Re B2 * Re (v1/v2) + Im B2 * Im (v1/v2) = 0"
using of_real_eq_0_iff
by blast+
moreover
have "Re(u1/u2) \<noteq> Re(v1/v2) \<or> Im(u1/u2) \<noteq> Im(v1/v2)"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "u1/u2 = v1/v2"
using complex_eqI by blast
thus False
- using uv `\<not> (u1, u2) \<approx>\<^sub>v (v1, v2)`
+ using uv \<open>\<not> (u1, u2) \<approx>\<^sub>v (v1, v2)\<close>
using "*"(1) "*"(2) complex_cvec_eq_mix[OF *(1) *(2)]
by (auto simp add: field_simps)
qed
moreover
have "Re A1 \<noteq> 0 \<or> Re B1 \<noteq> 0 \<or> Im B1 \<noteq> 0"
- using `?H1 \<noteq> mat_zero` **
+ using \<open>?H1 \<noteq> mat_zero\<close> **
by (metis complex_cnj_zero complex_of_real_Re mat_zero_def of_real_0)
ultimately
obtain k where
k: "Re A2 = k * Re A1" "Re B2 = k * Re B1" "Im B2 = k * Im B1"
using linear_system_homogenous_3_2[of "\<lambda>x y z. 1 * x + Re (u1 / u2) * y + Im (u1 / u2) * z" 1 "Re (u1/u2)" "Im (u1/u2)"
"\<lambda>x y z. 1 * x + Re (v1 / v2) * y + Im (v1 / v2) * z" 1 "Re (v1/v2)" "Im (v1/v2)"
"Re A2" "Re B2" "Im B2" "Re A1" "Re B1" "Im B1"]
by (auto simp add: field_simps)
have "Re A2 \<noteq> 0 \<or> Re B2 \<noteq> 0 \<or> Im B2 \<noteq> 0"
- using `?H2 \<noteq> mat_zero` **
+ using \<open>?H2 \<noteq> mat_zero\<close> **
by (metis complex_cnj_zero complex_of_real_Re mat_zero_def of_real_0)
hence "k \<noteq> 0"
using k
by auto
show "circline_eq_cmat ?H1 ?H2"
- using ** k `k \<noteq> 0`
+ using ** k \<open>k \<noteq> 0\<close>
by (auto simp add: vec_cnj_def) (rule_tac x="k" in exI, auto simp add: complex.expand)
qed
qed
qed
text \<open>The only h-line that goes trough zero and a non-zero point on the x-axis is the x-axis.\<close>
lemma is_poincare_line_0_real_is_x_axis:
assumes "is_poincare_line l" "0\<^sub>h \<in> circline_set l"
"x \<in> circline_set l \<inter> circline_set x_axis" "x \<noteq> 0\<^sub>h" "x \<noteq> \<infinity>\<^sub>h"
shows "l = x_axis"
using assms
using is_poincare_line_trough_zero_trough_infty[OF assms(1-2)]
using unique_circline_set[of x "0\<^sub>h" "\<infinity>\<^sub>h"]
by auto
text \<open>The only h-line that goes trough zero and a non-zero point on the y-axis is the y-axis.\<close>
lemma is_poincare_line_0_imag_is_y_axis:
assumes "is_poincare_line l" "0\<^sub>h \<in> circline_set l"
"y \<in> circline_set l \<inter> circline_set y_axis" "y \<noteq> 0\<^sub>h" "y \<noteq> \<infinity>\<^sub>h"
shows "l = y_axis"
using assms
using is_poincare_line_trough_zero_trough_infty[OF assms(1-2)]
using unique_circline_set[of y "0\<^sub>h" "\<infinity>\<^sub>h"]
by auto
(* ------------------------------------------------------------------ *)
subsubsection\<open>H-isometries preserve h-lines\<close>
(* ------------------------------------------------------------------ *)
text\<open>\emph{H-isometries} are defined as homographies (actions of Möbius transformations) and
antihomographies (compositions of actions of Möbius transformations with conjugation) that fix the
unit disc (map it onto itself). They also map h-lines onto h-lines\<close>
text\<open>We prove a bit more general lemma that states that all Möbius transformations that fix the
unit circle (not necessary the unit disc) map h-lines onto h-lines\<close>
lemma unit_circle_fix_preserve_is_poincare_line [simp]:
assumes "unit_circle_fix M" "is_poincare_line H"
shows "is_poincare_line (moebius_circline M H)"
using assms
unfolding is_poincare_line_iff
proof (safe)
let ?H' = "moebius_ocircline M (of_circline H)"
let ?U' = "moebius_ocircline M ounit_circle"
assume ++: "unit_circle_fix M" "perpendicular H unit_circle"
have ounit: "ounit_circle = moebius_ocircline M ounit_circle \<or>
ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)"
using ++(1) unit_circle_fix_iff[of M]
by (simp add: inj_of_ocircline moebius_circline_ocircline)
show "perpendicular (moebius_circline M H) unit_circle"
proof (cases "pos_oriented ?H'")
case True
hence *: "of_circline (of_ocircline ?H') = ?H'"
using of_circline_of_ocircline_pos_oriented
by blast
from ounit show ?thesis
proof
assume **: "ounit_circle = moebius_ocircline M ounit_circle"
show ?thesis
using ++
unfolding perpendicular_def
by (simp, subst moebius_circline_ocircline, subst *, subst **) simp
next
assume **: "ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)"
show ?thesis
using ++
unfolding perpendicular_def
by (simp, subst moebius_circline_ocircline, subst *, subst **) simp
qed
next
case False
hence *: "of_circline (of_ocircline ?H') = opposite_ocircline ?H'"
by (metis of_circline_of_ocircline pos_oriented_of_circline)
from ounit show ?thesis
proof
assume **: "ounit_circle = moebius_ocircline M ounit_circle"
show ?thesis
using ++
unfolding perpendicular_def
by (simp, subst moebius_circline_ocircline, subst *, subst **) simp
next
assume **: "ounit_circle = moebius_ocircline M (opposite_ocircline ounit_circle)"
show ?thesis
using ++
unfolding perpendicular_def
by (simp, subst moebius_circline_ocircline, subst *, subst **) simp
qed
qed
qed simp
lemma unit_circle_fix_preserve_is_poincare_line_iff [simp]:
assumes "unit_circle_fix M"
shows "is_poincare_line (moebius_circline M H) \<longleftrightarrow> is_poincare_line H"
using assms
using unit_circle_fix_preserve_is_poincare_line[of M H]
using unit_circle_fix_preserve_is_poincare_line[of "moebius_inv M" "moebius_circline M H"]
by (auto simp del: unit_circle_fix_preserve_is_poincare_line)
text\<open>Since h-lines are preserved by transformations that fix the unit circle, so is collinearity.\<close>
lemma unit_disc_fix_preserve_poincare_collinear [simp]:
assumes "unit_circle_fix M" "poincare_collinear A"
shows "poincare_collinear (moebius_pt M ` A)"
using assms
unfolding poincare_collinear_def
by (auto, rule_tac x="moebius_circline M p" in exI, auto)
lemma unit_disc_fix_preserve_poincare_collinear_iff [simp]:
assumes "unit_circle_fix M"
shows "poincare_collinear (moebius_pt M ` A) \<longleftrightarrow> poincare_collinear A"
using assms
using unit_disc_fix_preserve_poincare_collinear[of M A]
using unit_disc_fix_preserve_poincare_collinear[of "moebius_inv M" "moebius_pt M ` A"]
by (auto simp del: unit_disc_fix_preserve_poincare_collinear)
lemma unit_disc_fix_preserve_poincare_collinear3 [simp]:
assumes "unit_disc_fix M"
shows "poincare_collinear {moebius_pt M u, moebius_pt M v, moebius_pt M w} \<longleftrightarrow>
poincare_collinear {u, v, w}"
using assms unit_disc_fix_preserve_poincare_collinear_iff[of M "{u, v, w}"]
by simp
text\<open>Conjugation is also an h-isometry and it preserves h-lines.\<close>
lemma is_poincare_line_conjugate_circline [simp]:
assumes "is_poincare_line H"
shows "is_poincare_line (conjugate_circline H)"
using assms
by (transfer, transfer, auto simp add: mat_cnj_def hermitean_def mat_adj_def)
lemma is_poincare_line_conjugate_circline_iff [simp]:
shows "is_poincare_line (conjugate_circline H) \<longleftrightarrow> is_poincare_line H"
using is_poincare_line_conjugate_circline[of "conjugate_circline H"]
by auto
text\<open>Since h-lines are preserved by conjugation, so is collinearity.\<close>
lemma conjugate_preserve_poincare_collinear [simp]:
assumes "poincare_collinear A"
shows "poincare_collinear (conjugate ` A)"
using assms
unfolding poincare_collinear_def
by auto (rule_tac x="conjugate_circline p" in exI, auto)
lemma conjugate_conjugate [simp]: "conjugate ` conjugate ` A = A"
by (auto simp add: image_iff)
lemma conjugate_preserve_poincare_collinear_iff [simp]:
shows "poincare_collinear (conjugate ` A) \<longleftrightarrow> poincare_collinear A"
using conjugate_preserve_poincare_collinear[of "A"]
using conjugate_preserve_poincare_collinear[of "conjugate ` A"]
by (auto simp del: conjugate_preserve_poincare_collinear)
(* ------------------------------------------------------------------ *)
subsubsection\<open>Mapping h-lines to x-axis\<close>
(* ------------------------------------------------------------------ *)
text\<open>Each h-line in the Poincar\'e model can be mapped onto the x-axis (by a unit-disc preserving
Möbius transformation).\<close>
lemma ex_unit_disc_fix_is_poincare_line_to_x_axis:
assumes "is_poincare_line l"
shows "\<exists> M. unit_disc_fix M \<and> moebius_circline M l = x_axis"
proof-
from assms obtain u v where "u \<noteq> v" "u \<in> unit_disc" "v \<in> unit_disc" and "{u, v} \<subseteq> circline_set l"
using ex_is_poincare_line_points
by blast
then obtain M where *: "unit_disc_fix M" "moebius_pt M u = 0\<^sub>h" "moebius_pt M v \<in> positive_x_axis"
using ex_unit_disc_fix_to_zero_positive_x_axis[of u v]
by auto
moreover
hence "{0\<^sub>h, moebius_pt M v} \<subseteq> circline_set x_axis"
unfolding positive_x_axis_def
by auto
moreover
have "moebius_pt M v \<noteq> 0\<^sub>h"
using \<open>u \<noteq> v\<close> *
by (metis moebius_pt_neq_I)
moreover
have "moebius_pt M v \<noteq> \<infinity>\<^sub>h"
using \<open>unit_disc_fix M\<close> \<open>v \<in> unit_disc\<close>
using unit_disc_fix_discI
by fastforce
ultimately
show ?thesis
using \<open>is_poincare_line l\<close> \<open>{u, v} \<subseteq> circline_set l\<close> \<open>unit_disc_fix M\<close>
using is_poincare_line_0_real_is_x_axis[of "moebius_circline M l" "moebius_pt M v"]
by (rule_tac x="M" in exI, force)
qed
text \<open>When proving facts about h-lines, without loss of generality it can be assumed that h-line is
the x-axis (if the property being proved is invariant under Möbius transformations that fix the
unit disc).\<close>
lemma wlog_line_x_axis:
assumes is_line: "is_poincare_line H"
assumes x_axis: "P x_axis"
assumes preserves: "\<And> M. \<lbrakk>unit_disc_fix M; P (moebius_circline M H)\<rbrakk> \<Longrightarrow> P H"
shows "P H"
using assms
using ex_unit_disc_fix_is_poincare_line_to_x_axis[of H]
by auto
(* ------------------------------------------------------------------ *)
subsection\<open>Construction of the h-line between the two given points\<close>
(* ------------------------------------------------------------------ *)
text\<open>Next we show how to construct the (unique) h-line between the two given points in the Poincar\'e model\<close>
text\<open>
Geometrically, h-line can be constructed by finding the inverse point of one of the two points and
by constructing the circle (or line) trough it and the two given points.
Algebraically, for two given points $u$ and $v$ in $\mathbb{C}$, the h-line matrix coefficients can
be $A = i\cdot(u\overline{v}-v\overline{u})$ and $B = i\cdot(v(|u|^2+1) - u(|v|^2+1))$.
We need to extend this to homogenous coordinates. There are several degenerate cases.
- If $\{z, w\} = \{0_h, \infty_h\}$ then there is no unique h-line (any line trough zero is an h-line).
- If z and w are mutually inverse, then the construction fails (both geometric and algebraic).
- If z and w are different points on the unit circle, then the standard construction fails (only geometric).
- None of this problematic cases occur when z and w are inside the unit disc.
We express the construction algebraically, and construct the Hermitean circline matrix for the two
points given in homogenous coordinates. It works correctly in all cases except when the two points
are the same or are mutually inverse.
\<close>
definition mk_poincare_line_cmat :: "real \<Rightarrow> complex \<Rightarrow> complex_mat" where
[simp]: "mk_poincare_line_cmat A B = (cor A, B, cnj B, cor A)"
lemma mk_poincare_line_cmat_zero_iff:
"mk_poincare_line_cmat A B = mat_zero \<longleftrightarrow> A = 0 \<and> B = 0"
by auto
lemma mk_poincare_line_cmat_hermitean
[simp]: "hermitean (mk_poincare_line_cmat A B)"
by simp
lemma mk_poincare_line_cmat_scale:
"cor k *\<^sub>s\<^sub>m mk_poincare_line_cmat A B = mk_poincare_line_cmat (k * A) (k * B)"
by simp
definition poincare_line_cvec_cmat :: "complex_vec \<Rightarrow> complex_vec \<Rightarrow> complex_mat" where
[simp]: "poincare_line_cvec_cmat z w =
(let (z1, z2) = z;
(w1, w2) = w;
nom = w1*cnj w2*(z1*cnj z1 + z2*cnj z2) - z1*cnj z2*(w1*cnj w1 + w2*cnj w2);
den = z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2
in if den \<noteq> 0 then
mk_poincare_line_cmat (Re(\<i>*den)) (\<i>*nom)
else if z1*cnj z2 \<noteq> 0 then
mk_poincare_line_cmat 0 (\<i>*z1*cnj z2)
else if w1*cnj w2 \<noteq> 0 then
mk_poincare_line_cmat 0 (\<i>*w1*cnj w2)
else
mk_poincare_line_cmat 0 \<i>)"
lemma poincare_line_cvec_cmat_AeqD:
assumes "poincare_line_cvec_cmat z w = (A, B, C, D)"
shows "A = D"
using assms
by (cases z, cases w) (auto split: if_split_asm)
lemma poincare_line_cvec_cmat_hermitean [simp]:
shows "hermitean (poincare_line_cvec_cmat z w)"
by (cases z, cases w) (auto split: if_split_asm simp del: mk_poincare_line_cmat_def)
lemma poincare_line_cvec_cmat_nonzero [simp]:
assumes "z \<noteq> vec_zero" "w \<noteq> vec_zero"
shows "poincare_line_cvec_cmat z w \<noteq> mat_zero"
proof-
obtain z1 z2 w1 w2 where *: "z = (z1, z2)" "w = (w1, w2)"
by (cases z, cases w, auto)
let ?den = "z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2"
show ?thesis
proof (cases "?den \<noteq> 0")
case True
have "is_real (\<i> * ?den)"
using eq_cnj_iff_real[of "\<i> *?den"]
by (simp add: field_simps)
hence "Re (\<i> * ?den) \<noteq> 0"
using \<open>?den \<noteq> 0\<close>
by (metis complex_i_not_zero complex_surj mult_eq_0_iff zero_complex.code)
thus ?thesis
using * \<open>?den \<noteq> 0\<close>
by (simp del: mk_poincare_line_cmat_def mat_zero_def add: mk_poincare_line_cmat_zero_iff)
next
case False
thus ?thesis
using *
by (simp del: mk_poincare_line_cmat_def mat_zero_def add: mk_poincare_line_cmat_zero_iff)
qed
qed
lift_definition poincare_line_hcoords_clmat :: "complex_homo_coords \<Rightarrow> complex_homo_coords \<Rightarrow> circline_mat" is poincare_line_cvec_cmat
using poincare_line_cvec_cmat_hermitean poincare_line_cvec_cmat_nonzero
by simp
lift_definition poincare_line :: "complex_homo \<Rightarrow> complex_homo \<Rightarrow> circline" is poincare_line_hcoords_clmat
proof transfer
fix za zb wa wb
assume "za \<noteq> vec_zero" "zb \<noteq> vec_zero" "wa \<noteq> vec_zero" "wb \<noteq> vec_zero"
assume "za \<approx>\<^sub>v zb" "wa \<approx>\<^sub>v wb"
obtain za1 za2 zb1 zb2 wa1 wa2 wb1 wb2 where
*: "(za1, za2) = za" "(zb1, zb2) = zb"
"(wa1, wa2) = wa" "(wb1, wb2) = wb"
by (cases za, cases zb, cases wa, cases wb, auto)
obtain kz kw where
**: "kz \<noteq> 0" "kw \<noteq> 0" "zb1 = kz * za1" "zb2 = kz * za2" "wb1 = kw * wa1" "wb2 = kw * wa2"
using \<open>za \<approx>\<^sub>v zb\<close> \<open>wa \<approx>\<^sub>v wb\<close> *[symmetric]
by auto
let ?nom = "\<lambda> z1 z2 w1 w2. w1*cnj w2*(z1*cnj z1 + z2*cnj z2) - z1*cnj z2*(w1*cnj w1 + w2*cnj w2)"
let ?den = "\<lambda> z1 z2 w1 w2. z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2"
show "circline_eq_cmat (poincare_line_cvec_cmat za wa)
(poincare_line_cvec_cmat zb wb)"
proof-
have "\<exists>k. k \<noteq> 0 \<and>
poincare_line_cvec_cmat (zb1, zb2) (wb1, wb2) = cor k *\<^sub>s\<^sub>m poincare_line_cvec_cmat (za1, za2) (wa1, wa2)"
proof (cases "?den za1 za2 wa1 wa2 \<noteq> 0")
case True
hence "?den zb1 zb2 wb1 wb2 \<noteq> 0"
using **
by (simp add: field_simps)
let ?k = "kz * cnj kz * kw * cnj kw"
have "?k \<noteq> 0"
using **
by simp
have "is_real ?k"
using eq_cnj_iff_real[of ?k]
by auto
have "cor (Re ?k) = ?k"
using \<open>is_real ?k\<close>
using complex_of_real_Re
by blast
have "Re ?k \<noteq> 0"
using \<open>?k \<noteq> 0\<close> \<open>cor (Re ?k) = ?k\<close>
by (metis of_real_0)
have arg1: "Re (\<i> * ?den zb1 zb2 wb1 wb2) = Re ?k * Re (\<i> * ?den za1 za2 wa1 wa2)"
apply (subst **)+
apply (subst Re_mult_real[symmetric, OF \<open>is_real ?k\<close>])
apply (rule arg_cong[where f=Re])
apply (simp add: field_simps)
done
have arg2: "\<i> * ?nom zb1 zb2 wb1 wb2 = ?k * \<i> * ?nom za1 za2 wa1 wa2"
using **
by (simp add: field_simps)
have "mk_poincare_line_cmat (Re (\<i>*?den zb1 zb2 wb1 wb2)) (\<i>*?nom zb1 zb2 wb1 wb2) =
cor (Re ?k) *\<^sub>s\<^sub>m mk_poincare_line_cmat (Re (\<i>*?den za1 za2 wa1 wa2)) (\<i>*?nom za1 za2 wa1 wa2)"
using \<open>cor (Re ?k) = ?k\<close> \<open>is_real ?k\<close>
apply (subst mk_poincare_line_cmat_scale)
apply (subst arg1, subst arg2)
apply (subst \<open>cor (Re ?k) = ?k\<close>)+
apply simp
done
thus ?thesis
using \<open>?den za1 za2 wa1 wa2 \<noteq> 0\<close> \<open>?den zb1 zb2 wb1 wb2 \<noteq> 0\<close>
using \<open>Re ?k \<noteq> 0\<close> \<open>cor (Re ?k) = ?k\<close>
by (rule_tac x="Re ?k" in exI, simp)
next
case False
hence "?den zb1 zb2 wb1 wb2 = 0"
using **
by (simp add: field_simps)
show ?thesis
proof (cases "za1*cnj za2 \<noteq> 0")
case True
hence "zb1*cnj zb2 \<noteq> 0"
using **
by (simp add: field_simps)
let ?k = "kz * cnj kz"
have "?k \<noteq> 0" "is_real ?k"
using **
using eq_cnj_iff_real[of ?k]
by auto
thus ?thesis
using \<open>za1 * cnj za2 \<noteq> 0\<close> \<open>zb1 * cnj zb2 \<noteq> 0\<close>
using \<open>\<not> (?den za1 za2 wa1 wa2 \<noteq> 0)\<close> \<open>?den zb1 zb2 wb1 wb2 = 0\<close> **
by (rule_tac x="Re (kz * cnj kz)" in exI, auto simp add: complex.expand)
next
case False
hence "zb1 * cnj zb2 = 0"
using **
by (simp add: field_simps)
show ?thesis
proof (cases "wa1 * cnj wa2 \<noteq> 0")
case True
hence "wb1*cnj wb2 \<noteq> 0"
using **
by (simp add: field_simps)
let ?k = "kw * cnj kw"
have "?k \<noteq> 0" "is_real ?k"
using **
using eq_cnj_iff_real[of ?k]
by auto
thus ?thesis
using \<open>\<not> (za1 * cnj za2 \<noteq> 0)\<close>
using \<open>wa1 * cnj wa2 \<noteq> 0\<close> \<open>wb1 * cnj wb2 \<noteq> 0\<close>
using \<open>\<not> (?den za1 za2 wa1 wa2 \<noteq> 0)\<close> \<open>?den zb1 zb2 wb1 wb2 = 0\<close> **
by (rule_tac x="Re (kw * cnj kw)" in exI)
(auto simp add: complex.expand)
next
case False
hence "wb1 * cnj wb2 = 0"
using **
by (simp add: field_simps)
thus ?thesis
using \<open>\<not> (za1 * cnj za2 \<noteq> 0)\<close> \<open>zb1 * cnj zb2 = 0\<close>
using \<open>\<not> (wa1 * cnj wa2 \<noteq> 0)\<close> \<open>wb1 * cnj wb2 = 0\<close>
using \<open>\<not> (?den za1 za2 wa1 wa2 \<noteq> 0)\<close> \<open>?den zb1 zb2 wb1 wb2 = 0\<close> **
by simp
qed
qed
qed
thus ?thesis
using *[symmetric]
by simp
qed
qed
subsubsection \<open>Correctness of the construction\<close>
text\<open>For finite points, our definition matches the classic algebraic definition for points in
$\mathbb{C}$ (given in ordinary, not homogenous coordinates).\<close>
lemma poincare_line_non_homogenous:
assumes "u \<noteq> \<infinity>\<^sub>h" "v \<noteq> \<infinity>\<^sub>h" "u \<noteq> v" "u \<noteq> inversion v"
shows "let u' = to_complex u; v' = to_complex v;
A = \<i> * (u' * cnj v' - v' * cnj u');
B = \<i> * (v' * ((cmod u')\<^sup>2 + 1) - u' * ((cmod v')\<^sup>2 + 1))
in poincare_line u v = mk_circline A B (cnj B) A"
using assms
unfolding unit_disc_def disc_def inversion_def
apply (simp add: Let_def)
proof (transfer, transfer, safe)
fix u1 u2 v1 v2
assume uv: "(u1, u2) \<noteq> vec_zero" "(v1, v2) \<noteq> vec_zero"
"\<not> (u1, u2) \<approx>\<^sub>v \<infinity>\<^sub>v" "\<not> (v1, v2) \<approx>\<^sub>v \<infinity>\<^sub>v"
"\<not> (u1, u2) \<approx>\<^sub>v (v1, v2)" "\<not> (u1, u2) \<approx>\<^sub>v conjugate_cvec (reciprocal_cvec (v1, v2))"
let ?u = "to_complex_cvec (u1, u2)" and ?v = "to_complex_cvec (v1, v2)"
let ?A = "\<i> * (?u * cnj ?v - ?v * cnj ?u)"
let ?B = "\<i> * (?v * ((cor (cmod ?u))\<^sup>2 + 1) - ?u * ((cor (cmod ?v))\<^sup>2 + 1))"
let ?C = "- (\<i> * (cnj ?v * ((cor (cmod ?u))\<^sup>2 + 1) - cnj ?u * ((cor (cmod ?v))\<^sup>2 + 1)))"
let ?D = ?A
let ?H = "(?A, ?B, ?C, ?D)"
let ?den = "u1 * cnj u2 * cnj v1 * v2 - v1 * cnj v2 * cnj u1 * u2"
have "u2 \<noteq> 0" "v2 \<noteq> 0"
using uv
using inf_cvec_z2_zero_iff
by blast+
have "\<not> (u1, u2) \<approx>\<^sub>v (cnj v2, cnj v1)"
using uv(6)
by (simp add: vec_cnj_def)
moreover
have "(cnj v2, cnj v1) \<noteq> vec_zero"
using uv(2)
by auto
ultimately
have *: "u1 * cnj v1 \<noteq> u2 * cnj v2" "u1 * v2 \<noteq> u2 * v1"
- using uv(5) uv(1) uv(2) `u2 \<noteq> 0` `v2 \<noteq> 0`
+ using uv(5) uv(1) uv(2) \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close>
using complex_cvec_eq_mix
by blast+
show "circline_eq_cmat (poincare_line_cvec_cmat (u1, u2) (v1, v2))
(mk_circline_cmat ?A ?B ?C ?D)"
proof (cases "?den \<noteq> 0")
case True
let ?nom = "v1 * cnj v2 * (u1 * cnj u1 + u2 * cnj u2) - u1 * cnj u2 * (v1 * cnj v1 + v2 * cnj v2)"
let ?H' = "mk_poincare_line_cmat (Re (\<i> * ?den)) (\<i> * ?nom)"
have "circline_eq_cmat ?H ?H'"
proof-
let ?k = "(u2 * cnj v2) * (v2 * cnj u2)"
have "is_real ?k"
using eq_cnj_iff_real
by fastforce
hence "cor (Re ?k) = ?k"
using complex_of_real_Re
by blast
have "Re (\<i> * ?den) = Re ?k * ?A"
proof-
have "?A = cnj ?A"
by (simp add: field_simps)
hence "is_real ?A"
using eq_cnj_iff_real
by fastforce
moreover
have "\<i> * ?den = cnj (\<i> * ?den)"
by (simp add: field_simps)
hence "is_real (\<i> * ?den)"
using eq_cnj_iff_real
by fastforce
hence "cor (Re (\<i> * ?den)) = \<i> * ?den"
using complex_of_real_Re
by blast
ultimately
show ?thesis
- using `cor (Re ?k) = ?k`
+ using \<open>cor (Re ?k) = ?k\<close>
by (simp add: field_simps)
qed
moreover
have "\<i> * ?nom = Re ?k * ?B"
- using `cor (Re ?k) = ?k` `u2 \<noteq> 0` `v2 \<noteq> 0` complex_mult_cnj_cmod[symmetric]
+ using \<open>cor (Re ?k) = ?k\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> complex_mult_cnj_cmod[symmetric]
by (auto simp add: field_simps)
moreover
have "?k \<noteq> 0"
- using `u2 \<noteq> 0` `v2 \<noteq> 0`
+ using \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close>
by simp
hence "Re ?k \<noteq> 0"
- using `is_real ?k`
+ using \<open>is_real ?k\<close>
by (metis \<open>cor (Re ?k) = ?k\<close> of_real_0)
ultimately
show ?thesis
by simp (rule_tac x="Re ?k" in exI, simp add: mult.commute)
qed
moreover
have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'"
- using `?den \<noteq> 0`
+ using \<open>?den \<noteq> 0\<close>
unfolding poincare_line_cvec_cmat_def
by (simp add: Let_def)
moreover
hence "hermitean ?H' \<and> ?H' \<noteq> mat_zero"
by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2))
hence "hermitean ?H \<and> ?H \<noteq> mat_zero"
- using `circline_eq_cmat ?H ?H'`
+ using \<open>circline_eq_cmat ?H ?H'\<close>
using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat
unfolding symp_def
by metis
hence "mk_circline_cmat ?A ?B ?C ?D = ?H"
by simp
ultimately
have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D)
(poincare_line_cvec_cmat (u1, u2) (v1, v2))"
by simp
thus ?thesis
using symp_circline_eq_cmat
unfolding symp_def
by blast
next
case False
let ?d = "v1 * (u1 * cnj u1 / (u2 * cnj u2) + 1) / v2 - u1 * (v1 * cnj v1 / (v2 * cnj v2) + 1) / u2"
let ?cd = "cnj v1 * (u1 * cnj u1 / (u2 * cnj u2) + 1) / cnj v2 - cnj u1 * (v1 * cnj v1 / (v2 * cnj v2) + 1) / cnj u2"
have "cnj ?d = ?cd"
by (simp add: mult.commute)
let ?d1 = "(v1 / v2) * (cnj u1 / cnj u2) - 1"
let ?d2 = "u1 / u2 - v1 / v2"
have **: "?d = ?d1 * ?d2"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close>
by(simp add: field_simps)
hence "?d \<noteq> 0"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` *
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> *
by auto (simp add: field_simps)+
have "is_real ?d1"
proof-
have "cnj ?d1 = ?d1"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` *
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> *
by (simp add: field_simps)
thus ?thesis
using eq_cnj_iff_real
by blast
qed
show ?thesis
proof (cases "u1 * cnj u2 \<noteq> 0")
case True
let ?nom = "u1 * cnj u2"
let ?H' = "mk_poincare_line_cmat 0 (\<i> * ?nom)"
have "circline_eq_cmat ?H ?H'"
proof-
let ?k = "(u1 * cnj u2) / ?d"
have "is_real ?k"
proof-
have "is_real ((u1 * cnj u2) / ?d2)"
proof-
let ?rhs = "(u2 * cnj u2) / (1 - (v1*u2)/(u1*v2))"
have 1: "(u1 * cnj u2) / ?d2 = ?rhs"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` * `u1 * cnj u2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> * \<open>u1 * cnj u2 \<noteq> 0\<close>
by (simp add: field_simps)
moreover
have "cnj ?rhs = ?rhs"
proof-
have "cnj (1 - v1 * u2 / (u1 * v2)) = 1 - v1 * u2 / (u1 * v2)"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` * `u1 * cnj u2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> * \<open>u1 * cnj u2 \<noteq> 0\<close>
by (simp add: field_simps)
moreover
have "cnj (u2 * cnj u2) = u2 * cnj u2"
by simp
ultimately
show ?thesis
by simp
qed
ultimately
show ?thesis
using eq_cnj_iff_real
by fastforce
qed
thus ?thesis
- using ** `is_real ?d1`
+ using ** \<open>is_real ?d1\<close>
by (metis complex_cnj_divide divide_divide_eq_left' eq_cnj_iff_real)
qed
have "?k \<noteq> 0"
- using `?d \<noteq> 0` `u1 * cnj u2 \<noteq> 0`
+ using \<open>?d \<noteq> 0\<close> \<open>u1 * cnj u2 \<noteq> 0\<close>
by simp
have "cnj ?k = ?k"
- using `is_real ?k`
+ using \<open>is_real ?k\<close>
using eq_cnj_iff_real by blast
have "Re ?k \<noteq> 0"
- using `?k \<noteq> 0` `is_real ?k`
+ using \<open>?k \<noteq> 0\<close> \<open>is_real ?k\<close>
by (metis complex.expand zero_complex.simps(1) zero_complex.simps(2))
have "u1 * cnj u2 = ?k * ?d"
- using `?d \<noteq> 0`
+ using \<open>?d \<noteq> 0\<close>
by simp
moreover
hence "cnj u1 * u2 = cnj ?k * cnj ?d"
by (metis complex_cnj_cnj complex_cnj_mult)
hence "cnj u1 * u2 = ?k * ?cd"
- using `cnj ?k = ?k` `cnj ?d = ?cd`
+ using \<open>cnj ?k = ?k\<close> \<open>cnj ?d = ?cd\<close>
by metis
ultimately
show ?thesis
- using `~ ?den \<noteq> 0` `u1 * cnj u2 \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` `Re ?k \<noteq> 0` `is_real ?k` `?d \<noteq> 0`
+ using \<open>~ ?den \<noteq> 0\<close> \<open>u1 * cnj u2 \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> \<open>Re ?k \<noteq> 0\<close> \<open>is_real ?k\<close> \<open>?d \<noteq> 0\<close>
using complex_mult_cnj_cmod[symmetric, of u1]
using complex_mult_cnj_cmod[symmetric, of v1]
using complex_mult_cnj_cmod[symmetric, of u2]
using complex_mult_cnj_cmod[symmetric, of v2]
apply (auto simp add: power_divide)
apply (rule_tac x="Re ?k" in exI)
apply simp
apply (simp add: field_simps)
done
qed
moreover
have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'"
- using `\<not> ?den \<noteq> 0` `u1 * cnj u2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u1 * cnj u2 \<noteq> 0\<close>
unfolding poincare_line_cvec_cmat_def
by (simp add: Let_def)
moreover
hence "hermitean ?H' \<and> ?H' \<noteq> mat_zero"
by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2))
hence "hermitean ?H \<and> ?H \<noteq> mat_zero"
- using `circline_eq_cmat ?H ?H'`
+ using \<open>circline_eq_cmat ?H ?H'\<close>
using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat
unfolding symp_def
by metis
hence "mk_circline_cmat ?A ?B ?C ?D = ?H"
by simp
ultimately
have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D)
(poincare_line_cvec_cmat (u1, u2) (v1, v2))"
by simp
thus ?thesis
using symp_circline_eq_cmat
unfolding symp_def
by blast
next
case False
show ?thesis
proof (cases "v1 * cnj v2 \<noteq> 0")
case True
let ?nom = "v1 * cnj v2"
let ?H' = "mk_poincare_line_cmat 0 (\<i> * ?nom)"
have "circline_eq_cmat ?H ?H'"
proof-
let ?k = "(v1 * cnj v2) / ?d"
have "is_real ?k"
proof-
have "is_real ((v1 * cnj v2) / ?d2)"
proof-
let ?rhs = "(v2 * cnj v2) / ((u1*v2)/(u2*v1) - 1)"
have 1: "(v1 * cnj v2) / ?d2 = ?rhs"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` * `v1 * cnj v2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> * \<open>v1 * cnj v2 \<noteq> 0\<close>
by (simp add: field_simps)
moreover
have "cnj ?rhs = ?rhs"
proof-
have "cnj (u1 * v2 / (u2 * v1) - 1) = u1 * v2 / (u2 * v1) - 1"
- using `\<not> ?den \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` * `v1 * cnj v2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> * \<open>v1 * cnj v2 \<noteq> 0\<close>
by (simp add: field_simps)
moreover
have "cnj (v2 * cnj v2) = v2 * cnj v2"
by simp
ultimately
show ?thesis
by simp
qed
ultimately
show ?thesis
using eq_cnj_iff_real
by fastforce
qed
thus ?thesis
- using ** `is_real ?d1`
+ using ** \<open>is_real ?d1\<close>
by (metis complex_cnj_divide divide_divide_eq_left' eq_cnj_iff_real)
qed
have "?k \<noteq> 0"
- using `?d \<noteq> 0` `v1 * cnj v2 \<noteq> 0`
+ using \<open>?d \<noteq> 0\<close> \<open>v1 * cnj v2 \<noteq> 0\<close>
by simp
have "cnj ?k = ?k"
- using `is_real ?k`
+ using \<open>is_real ?k\<close>
using eq_cnj_iff_real by blast
have "Re ?k \<noteq> 0"
- using `?k \<noteq> 0` `is_real ?k`
+ using \<open>?k \<noteq> 0\<close> \<open>is_real ?k\<close>
by (metis complex.expand zero_complex.simps(1) zero_complex.simps(2))
have "v1 * cnj v2 = ?k * ?d"
- using `?d \<noteq> 0`
+ using \<open>?d \<noteq> 0\<close>
by simp
moreover
hence "cnj v1 * v2 = cnj ?k * cnj ?d"
by (metis complex_cnj_cnj complex_cnj_mult)
hence "cnj v1 * v2 = ?k * ?cd"
- using `cnj ?k = ?k` `cnj ?d = ?cd`
+ using \<open>cnj ?k = ?k\<close> \<open>cnj ?d = ?cd\<close>
by metis
ultimately
show ?thesis
- using `~ ?den \<noteq> 0` `v1 * cnj v2 \<noteq> 0` `u2 \<noteq> 0` `v2 \<noteq> 0` `Re ?k \<noteq> 0` `is_real ?k` `?d \<noteq> 0`
+ using \<open>~ ?den \<noteq> 0\<close> \<open>v1 * cnj v2 \<noteq> 0\<close> \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close> \<open>Re ?k \<noteq> 0\<close> \<open>is_real ?k\<close> \<open>?d \<noteq> 0\<close>
using complex_mult_cnj_cmod[symmetric, of u1]
using complex_mult_cnj_cmod[symmetric, of v1]
using complex_mult_cnj_cmod[symmetric, of u2]
using complex_mult_cnj_cmod[symmetric, of v2]
apply (auto simp add: power_divide)
apply (rule_tac x="Re ?k" in exI)
apply simp
apply (simp add: field_simps)
done
qed
moreover
have "poincare_line_cvec_cmat (u1, u2) (v1, v2) = ?H'"
- using `\<not> ?den \<noteq> 0` `\<not> u1 * cnj u2 \<noteq> 0` `v1 * cnj v2 \<noteq> 0`
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>\<not> u1 * cnj u2 \<noteq> 0\<close> \<open>v1 * cnj v2 \<noteq> 0\<close>
unfolding poincare_line_cvec_cmat_def
by (simp add: Let_def)
moreover
hence "hermitean ?H' \<and> ?H' \<noteq> mat_zero"
by (metis mk_poincare_line_cmat_hermitean poincare_line_cvec_cmat_nonzero uv(1) uv(2))
hence "hermitean ?H \<and> ?H \<noteq> mat_zero"
- using `circline_eq_cmat ?H ?H'`
+ using \<open>circline_eq_cmat ?H ?H'\<close>
using circline_eq_cmat_hermitean_nonzero[of ?H' ?H] symp_circline_eq_cmat
unfolding symp_def
by metis
hence "mk_circline_cmat ?A ?B ?C ?D = ?H"
by simp
ultimately
have "circline_eq_cmat (mk_circline_cmat ?A ?B ?C ?D)
(poincare_line_cvec_cmat (u1, u2) (v1, v2))"
by simp
thus ?thesis
using symp_circline_eq_cmat
unfolding symp_def
by blast
next
case False
hence False
- using `\<not> ?den \<noteq> 0` `\<not> u1 * cnj u2 \<noteq> 0` uv
+ using \<open>\<not> ?den \<noteq> 0\<close> \<open>\<not> u1 * cnj u2 \<noteq> 0\<close> uv
by (simp add: \<open>u2 \<noteq> 0\<close> \<open>v2 \<noteq> 0\<close>)
thus ?thesis
by simp
qed
qed
qed
qed
text\<open>Our construction (in homogenous coordinates) always yields an h-line that contain two starting
points (this also holds for all degenerate cases except when points are the same).\<close>
lemma poincare_line [simp]:
assumes "z \<noteq> w"
shows "on_circline (poincare_line z w) z"
"on_circline (poincare_line z w) w"
proof-
have "on_circline (poincare_line z w) z \<and> on_circline (poincare_line z w) w"
using assms
proof (transfer, transfer)
fix z w
assume vz: "z \<noteq> vec_zero" "w \<noteq> vec_zero"
obtain z1 z2 w1 w2 where
zw: "(z1, z2) = z" "(w1, w2) = w"
by (cases z, cases w, auto)
let ?den = "z1*cnj z2*cnj w1*w2 - w1*cnj w2*cnj z1*z2"
have *: "cor (Re (\<i> * ?den)) = \<i> * ?den"
proof-
have "cnj ?den = -?den"
by auto
hence "is_imag ?den"
using eq_minus_cnj_iff_imag[of ?den]
by simp
thus ?thesis
using complex_of_real_Re[of "\<i> * ?den"]
by simp
qed
show "on_circline_cmat_cvec (poincare_line_cvec_cmat z w) z \<and>
on_circline_cmat_cvec (poincare_line_cvec_cmat z w) w"
unfolding poincare_line_cvec_cmat_def mk_poincare_line_cmat_def
apply (subst zw[symmetric])+
unfolding Let_def prod.case
apply (subst *)+
by (auto simp add: vec_cnj_def field_simps)
qed
thus "on_circline (poincare_line z w) z" "on_circline (poincare_line z w) w"
by auto
qed
lemma poincare_line_circline_set [simp]:
assumes "z \<noteq> w"
shows "z \<in> circline_set (poincare_line z w)"
"w \<in> circline_set (poincare_line z w)"
using assms
by (auto simp add: circline_set_def)
text\<open>When the points are different, the constructed line matrix always has a negative determinant\<close>
lemma poincare_line_type:
assumes "z \<noteq> w"
shows "circline_type (poincare_line z w) = -1"
proof-
have "\<exists> a b. a \<noteq> b \<and> {a, b} \<subseteq> circline_set (poincare_line z w)"
using poincare_line[of z w] assms
unfolding circline_set_def
by (rule_tac x=z in exI, rule_tac x=w in exI, simp)
thus ?thesis
using circline_type[of "poincare_line z w"]
using circline_type_pos_card_eq0[of "poincare_line z w"]
using circline_type_zero_card_eq1[of "poincare_line z w"]
by auto
qed
text\<open>The constructed line is an h-line in the Poincar\'e model (in all cases when the two points are
different)\<close>
lemma is_poincare_line_poincare_line [simp]:
assumes "z \<noteq> w"
shows "is_poincare_line (poincare_line z w)"
using poincare_line_type[of z w, OF assms]
proof (transfer, transfer)
fix z w
assume vz: "z \<noteq> vec_zero" "w \<noteq> vec_zero"
obtain A B C D where *: "poincare_line_cvec_cmat z w = (A, B, C, D)"
by (cases "poincare_line_cvec_cmat z w") auto
assume "circline_type_cmat (poincare_line_cvec_cmat z w) = - 1"
thus "is_poincare_line_cmat (poincare_line_cvec_cmat z w)"
using vz *
using poincare_line_cvec_cmat_hermitean[of z w]
using poincare_line_cvec_cmat_nonzero[of z w]
using poincare_line_cvec_cmat_AeqD[of z w A B C D]
using hermitean_elems[of A B C D]
using cmod_power2[of D] cmod_power2[of C]
unfolding is_poincare_line_cmat_def
by (simp del: poincare_line_cvec_cmat_def add: sgn_1_neg power2_eq_square)
qed
text \<open>When the points are different, the constructed h-line between two points also contains their inverses\<close>
lemma poincare_line_inversion:
assumes "z \<noteq> w"
shows "on_circline (poincare_line z w) (inversion z)"
"on_circline (poincare_line z w) (inversion w)"
using assms
using is_poincare_line_poincare_line[OF \<open>z \<noteq> w\<close>]
using is_poincare_line_inverse_point
unfolding circline_set_def
by auto
text \<open>When the points are different, the onstructed h-line between two points contains the inverse of its every point\<close>
lemma poincare_line_inversion_full:
assumes "u \<noteq> v"
assumes "on_circline (poincare_line u v) x"
shows "on_circline (poincare_line u v) (inversion x)"
using is_poincare_line_inverse_point[of "poincare_line u v" x]
- using is_poincare_line_poincare_line[OF `u \<noteq> v`] assms
+ using is_poincare_line_poincare_line[OF \<open>u \<noteq> v\<close>] assms
unfolding circline_set_def
by simp
subsubsection \<open>Existence of h-lines\<close>
text\<open>There is an h-line trough every point in the Poincar\'e model\<close>
lemma ex_poincare_line_one_point:
shows "\<exists> l. is_poincare_line l \<and> z \<in> circline_set l"
proof (cases "z = 0\<^sub>h")
case True
thus ?thesis
by (rule_tac x="x_axis" in exI) simp
next
case False
thus ?thesis
by (rule_tac x="poincare_line 0\<^sub>h z" in exI) auto
qed
lemma poincare_collinear_singleton [simp]:
assumes "u \<in> unit_disc"
shows "poincare_collinear {u}"
using assms
using ex_poincare_line_one_point[of u]
by (auto simp add: poincare_collinear_def)
text\<open>There is an h-line trough every two points in the Poincar\'e model\<close>
lemma ex_poincare_line_two_points:
assumes "z \<noteq> w"
shows "\<exists> l. is_poincare_line l \<and> z \<in> circline_set l \<and> w \<in> circline_set l"
using assms
by (rule_tac x="poincare_line z w" in exI, simp)
lemma poincare_collinear_doubleton [simp]:
assumes "u \<in> unit_disc" "v \<in> unit_disc"
shows "poincare_collinear {u, v}"
using assms
using ex_poincare_line_one_point[of u]
using ex_poincare_line_two_points[of u v]
by (cases "u = v") (simp_all add: poincare_collinear_def)
subsubsection \<open>Uniqueness of h-lines\<close>
text \<open>The only h-line between two points is the one obtained by the line-construction.\<close>
text \<open>First we show this only for two different points inside the disc.\<close>
lemma unique_poincare_line:
assumes in_disc: "u \<noteq> v" "u \<in> unit_disc" "v \<in> unit_disc"
assumes on_l: "u \<in> circline_set l" "v \<in> circline_set l" "is_poincare_line l"
shows "l = poincare_line u v"
using assms
using unique_is_poincare_line[of u v l "poincare_line u v"]
unfolding circline_set_def
by auto
text\<open>The assumption that the points are inside the disc can be relaxed.\<close>
lemma unique_poincare_line_general:
assumes in_disc: "u \<noteq> v" "u \<noteq> inversion v"
assumes on_l: "u \<in> circline_set l" "v \<in> circline_set l" "is_poincare_line l"
shows "l = poincare_line u v"
using assms
using unique_is_poincare_line_general[of u v l "poincare_line u v"]
unfolding circline_set_def
by auto
text\<open>The explicit line construction enables us to prove that there exists a unique h-line through any
given two h-points (uniqueness part was already shown earlier).\<close>
text \<open>First we show this only for two different points inside the disc.\<close>
lemma ex1_poincare_line:
assumes "u \<noteq> v" "u \<in> unit_disc" "v \<in> unit_disc"
shows "\<exists>! l. is_poincare_line l \<and> u \<in> circline_set l \<and> v \<in> circline_set l"
proof (rule ex1I)
let ?l = "poincare_line u v"
show "is_poincare_line ?l \<and> u \<in> circline_set ?l \<and> v \<in> circline_set ?l"
using assms
unfolding circline_set_def
by auto
next
fix l
assume "is_poincare_line l \<and> u \<in> circline_set l \<and> v \<in> circline_set l"
thus "l = poincare_line u v"
using unique_poincare_line assms
by auto
qed
text \<open>The assumption that the points are in the disc can be relaxed.\<close>
lemma ex1_poincare_line_general:
assumes "u \<noteq> v" "u \<noteq> inversion v"
shows "\<exists>! l. is_poincare_line l \<and> u \<in> circline_set l \<and> v \<in> circline_set l"
proof (rule ex1I)
let ?l = "poincare_line u v"
show "is_poincare_line ?l \<and> u \<in> circline_set ?l \<and> v \<in> circline_set ?l"
using assms
unfolding circline_set_def
by auto
next
fix l
assume "is_poincare_line l \<and> u \<in> circline_set l \<and> v \<in> circline_set l"
thus "l = poincare_line u v"
using unique_poincare_line_general assms
by auto
qed
subsubsection \<open>Some consequences of line uniqueness\<close>
text\<open>H-line $uv$ is the same as the h-line $vu$.\<close>
lemma poincare_line_sym:
assumes "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
shows "poincare_line u v = poincare_line v u"
using assms
using unique_poincare_line[of u v "poincare_line v u"]
by simp
lemma poincare_line_sym_general:
assumes "u \<noteq> v" "u \<noteq> inversion v"
shows "poincare_line u v = poincare_line v u"
using assms
using unique_poincare_line_general[of u v "poincare_line v u"]
by simp
text\<open>Each h-line is the h-line constructed out of its two arbitrary different points.\<close>
lemma ex_poincare_line_points:
assumes "is_poincare_line H"
shows "\<exists> u v. u \<in> unit_disc \<and> v \<in> unit_disc \<and> u \<noteq> v \<and> H = poincare_line u v"
using assms
using ex_is_poincare_line_points
using unique_poincare_line[where l=H]
by fastforce
text\<open>If an h-line contains two different points on x-axis/y-axis then it is the x-axis/y-axis.\<close>
lemma poincare_line_0_real_is_x_axis:
assumes "x \<in> circline_set x_axis" "x \<noteq> 0\<^sub>h" "x \<noteq> \<infinity>\<^sub>h"
shows "poincare_line 0\<^sub>h x = x_axis"
using assms
using is_poincare_line_0_real_is_x_axis[of "poincare_line 0\<^sub>h x" x]
by auto
lemma poincare_line_0_imag_is_y_axis:
assumes "y \<in> circline_set y_axis" "y \<noteq> 0\<^sub>h" "y \<noteq> \<infinity>\<^sub>h"
shows "poincare_line 0\<^sub>h y = y_axis"
using assms
using is_poincare_line_0_imag_is_y_axis[of "poincare_line 0\<^sub>h y" y]
by auto
lemma poincare_line_x_axis:
assumes "x \<in> unit_disc" "y \<in> unit_disc" "x \<in> circline_set x_axis" "y \<in> circline_set x_axis" "x \<noteq> y"
shows "poincare_line x y = x_axis"
using assms
using unique_poincare_line
by auto
lemma poincare_line_minus_one_one [simp]:
shows "poincare_line (of_complex (-1)) (of_complex 1) = x_axis"
proof-
have "0\<^sub>h \<in> circline_set (poincare_line (of_complex (-1)) (of_complex 1))"
unfolding circline_set_def
by simp (transfer, transfer, simp add: vec_cnj_def)
hence "poincare_line 0\<^sub>h (of_complex 1) = poincare_line (of_complex (-1)) (of_complex 1)"
by (metis is_poincare_line_poincare_line is_poincare_line_trough_zero_trough_infty not_zero_on_unit_circle of_complex_inj of_complex_one one_neq_neg_one one_on_unit_circle poincare_line_0_real_is_x_axis poincare_line_circline_set(2) reciprocal_involution reciprocal_one reciprocal_zero unique_circline_01inf')
thus ?thesis
using poincare_line_0_real_is_x_axis[of "of_complex 1"]
by auto
qed
subsubsection \<open>Transformations of constructed lines\<close>
text\<open>Unit dics preserving Möbius transformations preserve the h-line construction\<close>
lemma unit_disc_fix_preserve_poincare_line [simp]:
assumes "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
shows "poincare_line (moebius_pt M u) (moebius_pt M v) = moebius_circline M (poincare_line u v)"
proof (rule unique_poincare_line[symmetric])
show "moebius_pt M u \<noteq> moebius_pt M v"
using \<open>u \<noteq> v\<close>
by auto
next
show "moebius_pt M u \<in> circline_set (moebius_circline M (poincare_line u v))"
"moebius_pt M v \<in> circline_set (moebius_circline M (poincare_line u v))"
unfolding circline_set_def
using moebius_circline[of M "poincare_line u v"] \<open>u \<noteq> v\<close>
by auto
next
from assms(1) have "unit_circle_fix M"
by simp
thus "is_poincare_line (moebius_circline M (poincare_line u v))"
using unit_circle_fix_preserve_is_poincare_line assms
by auto
next
show "moebius_pt M u \<in> unit_disc" "moebius_pt M v \<in> unit_disc"
using assms(2-3) unit_disc_fix_iff[OF assms(1)]
by auto
qed
text\<open>Conjugate preserve the h-line construction\<close>
lemma conjugate_preserve_poincare_line [simp]:
assumes "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
shows "poincare_line (conjugate u) (conjugate v) = conjugate_circline (poincare_line u v)"
proof-
have "conjugate u \<noteq> conjugate v"
using \<open>u \<noteq> v\<close>
by (auto simp add: conjugate_inj)
moreover
have "conjugate u \<in> unit_disc" "conjugate v \<in> unit_disc"
using assms
by auto
moreover
have "conjugate u \<in> circline_set (conjugate_circline (poincare_line u v))"
"conjugate v \<in> circline_set (conjugate_circline (poincare_line u v))"
using \<open>u \<noteq> v\<close>
by simp_all
moreover
have "is_poincare_line (conjugate_circline (poincare_line u v))"
using is_poincare_line_poincare_line[OF \<open>u \<noteq> v\<close>]
by simp
ultimately
show ?thesis
using unique_poincare_line[of "conjugate u" "conjugate v" "conjugate_circline (poincare_line u v)"]
by simp
qed
subsubsection \<open>Collinear points and h-lines\<close>
lemma poincare_collinear3_poincare_line_general:
assumes "poincare_collinear {a, a1, a2}" "a1 \<noteq> a2" "a1 \<noteq> inversion a2"
shows "a \<in> circline_set (poincare_line a1 a2)"
using assms
using poincare_collinear_def unique_poincare_line_general
by auto
lemma poincare_line_poincare_collinear3_general:
assumes "a \<in> circline_set (poincare_line a1 a2)" "a1 \<noteq> a2"
shows "poincare_collinear {a, a1, a2}"
using assms
unfolding poincare_collinear_def
by (rule_tac x="poincare_line a1 a2" in exI, simp)
lemma poincare_collinear3_poincare_lines_equal_general:
assumes "poincare_collinear {a, a1, a2}" "a \<noteq> a1" "a \<noteq> a2" "a \<noteq> inversion a1" "a \<noteq> inversion a2"
shows "poincare_line a a1 = poincare_line a a2"
using assms
using unique_poincare_line_general[of a a2 "poincare_line a a1"]
by (simp add: insert_commute poincare_collinear3_poincare_line_general)
subsubsection \<open>Points collinear with @{term "0\<^sub>h"}\<close>
lemma poincare_collinear_zero_iff:
assumes "of_complex y' \<in> unit_disc" and "of_complex z' \<in> unit_disc" and
"y' \<noteq> z'" and "y' \<noteq> 0" and "z' \<noteq> 0"
shows "poincare_collinear {0\<^sub>h, of_complex y', of_complex z'} \<longleftrightarrow>
y'*cnj z' = cnj y'*z'" (is "?lhs \<longleftrightarrow> ?rhs")
proof-
have "of_complex y' \<noteq> of_complex z'"
using assms
using of_complex_inj
by blast
show ?thesis
proof
assume ?lhs
hence "0\<^sub>h \<in> circline_set (poincare_line (of_complex y') (of_complex z'))"
using unique_poincare_line[of "of_complex y'" "of_complex z'"]
using assms \<open>of_complex y' \<noteq> of_complex z'\<close>
unfolding poincare_collinear_def
by auto
moreover
let ?mix = "y' * cnj z' - cnj y' * z'"
have "is_real (\<i> * ?mix)"
using eq_cnj_iff_real[of ?mix]
by auto
hence "y' * cnj z' = cnj y' * z' \<longleftrightarrow> Re (\<i> * ?mix) = 0"
using complex.expand[of "\<i> * ?mix" 0]
by (metis complex_i_not_zero eq_iff_diff_eq_0 mult_eq_0_iff zero_complex.simps(1) zero_complex.simps(2))
ultimately
show ?rhs
using \<open>y' \<noteq> z'\<close> \<open>y' \<noteq> 0\<close> \<open>z' \<noteq> 0\<close>
unfolding circline_set_def
by simp (transfer, transfer, auto simp add: vec_cnj_def split: if_split_asm, metis Re_complex_of_real Re_mult_real Im_complex_of_real)
next
assume ?rhs
thus ?lhs
using assms \<open>of_complex y' \<noteq> of_complex z'\<close>
unfolding poincare_collinear_def
unfolding circline_set_def
apply (rule_tac x="poincare_line (of_complex y') (of_complex z')" in exI)
apply auto
apply (transfer, transfer, simp add: vec_cnj_def)
done
qed
qed
lemma poincare_collinear_zero_polar_form:
assumes "poincare_collinear {0\<^sub>h, of_complex x, of_complex y}" and
"x \<noteq> 0" and "y \<noteq> 0" and "of_complex x \<in> unit_disc" and "of_complex y \<in> unit_disc"
shows "\<exists> \<phi> rx ry. x = cor rx * cis \<phi> \<and> y = cor ry * cis \<phi> \<and> rx \<noteq> 0 \<and> ry \<noteq> 0"
proof-
from \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> obtain \<phi> \<phi>' rx ry where
polar: "x = cor rx * cis \<phi>" "y = cor ry * cis \<phi>'" and "\<phi> = arg x" "\<phi>' = arg y"
by (metis cmod_cis)
hence "rx \<noteq> 0" "ry \<noteq> 0"
using \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close>
by auto
have "of_complex y \<in> circline_set (poincare_line 0\<^sub>h (of_complex x))"
using assms
using unique_poincare_line[of "0\<^sub>h" "of_complex x"]
unfolding poincare_collinear_def
unfolding circline_set_def
using of_complex_zero_iff
by fastforce
hence "cnj x * y = x * cnj y"
using \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close>
unfolding circline_set_def
by simp (transfer, transfer, simp add: vec_cnj_def field_simps)
hence "cis(\<phi>' - \<phi>) = cis(\<phi> - \<phi>')"
using polar \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (simp add: cis_mult)
hence "sin (\<phi>' - \<phi>) = 0"
using cis_diff_cis_opposite[of "\<phi>' - \<phi>"]
by simp
then obtain k :: int where "\<phi>' - \<phi> = k * pi"
using sin_zero_iff_int2[of "\<phi>' - \<phi>"]
by auto
hence *: "\<phi>' = \<phi> + k * pi"
by simp
show ?thesis
proof (cases "even k")
case True
then obtain k' where "k = 2*k'"
using evenE by blast
hence "cis \<phi> = cis \<phi>'"
using * cos_periodic_int sin_periodic_int
by (simp add: cis.ctr field_simps)
thus ?thesis
using polar \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (rule_tac x=\<phi> in exI, rule_tac x=rx in exI, rule_tac x=ry in exI) simp
next
case False
then obtain k' where "k = 2*k' + 1"
using oddE by blast
hence "cis \<phi> = - cis \<phi>'"
using * cos_periodic_int sin_periodic_int
by (simp add: cis.ctr complex_minus field_simps)
thus ?thesis
using polar \<open>rx \<noteq> 0\<close> \<open>ry \<noteq> 0\<close>
by (rule_tac x=\<phi> in exI, rule_tac x=rx in exI, rule_tac x="-ry" in exI) simp
qed
qed
end
diff --git a/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy b/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy
--- a/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy
+++ b/thys/Poincare_Disc/Poincare_Lines_Axis_Intersections.thy
@@ -1,1235 +1,1235 @@
theory Poincare_Lines_Axis_Intersections
imports Poincare_Between
begin
(* ------------------------------------------------------------------ *)
section\<open>Intersection of h-lines with the x-axis in the Poincar\'e model\<close>
(* ------------------------------------------------------------------ *)
(* ---------------------------------------------------------------- *)
subsection\<open>Betweeness of x-axis intersection\<close>
(* ---------------------------------------------------------------- *)
text\<open>The intersection point of the h-line determined by points $u$ and $v$ and the x-axis is between
$u$ and $v$, then $u$ and $v$ are in the opposite half-planes (one must be in the upper, and the
other one in the lower half-plane).\<close>
lemma poincare_between_x_axis_intersection:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "z \<in> unit_disc" and "u \<noteq> v"
assumes "u \<notin> circline_set x_axis" and "v \<notin> circline_set x_axis"
assumes "z \<in> circline_set (poincare_line u v) \<inter> circline_set x_axis"
shows "poincare_between u z v \<longleftrightarrow> arg (to_complex u) * arg (to_complex v) < 0"
proof-
have "\<forall> u v. u \<in> unit_disc \<and> v \<in> unit_disc \<and> u \<noteq> v \<and>
u \<notin> circline_set x_axis \<and> v \<notin> circline_set x_axis \<and>
z \<in> circline_set (poincare_line u v) \<inter> circline_set x_axis \<longrightarrow>
(poincare_between u z v \<longleftrightarrow> arg (to_complex u) * arg (to_complex v) < 0)" (is "?P z")
proof (rule wlog_real_zero)
show "?P 0\<^sub>h"
proof ((rule allI)+, rule impI, (erule conjE)+)
fix u v
assume *: "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
"u \<notin> circline_set x_axis" "v \<notin> circline_set x_axis"
"0\<^sub>h \<in> circline_set (poincare_line u v) \<inter> circline_set x_axis"
obtain u' v' where uv: "u = of_complex u'" "v = of_complex v'"
using * inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
hence "u \<noteq> 0\<^sub>h" "v \<noteq> 0\<^sub>h" "u' \<noteq> 0" "v' \<noteq> 0"
using *
by auto
hence "arg u' \<noteq> 0" "arg v' \<noteq> 0"
using * arg_0_iff[of u'] arg_0_iff[of v']
unfolding circline_set_x_axis uv
by auto
have "poincare_collinear {0\<^sub>h, u, v}"
using *
unfolding poincare_collinear_def
by (rule_tac x="poincare_line u v" in exI, simp)
have "(\<exists>k<0. u' = cor k * v') \<longleftrightarrow> (arg u' * arg v' < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume "?lhs"
then obtain k where "k < 0" "u' = cor k * v'"
by auto
thus ?rhs
using arg_mult_real_negative[of k v'] arg_uminus_opposite_sign[of v']
using \<open>u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close> \<open>arg u' \<noteq> 0\<close> \<open>arg v' \<noteq> 0\<close>
by (auto simp add: mult_neg_pos mult_pos_neg)
next
assume ?rhs
obtain ru rv \<phi> where polar: "u' = cor ru * cis \<phi>" "v' = cor rv * cis \<phi>"
using \<open>poincare_collinear {0\<^sub>h, u, v}\<close> poincare_collinear_zero_polar_form[of u' v'] uv * \<open>u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close>
by auto
have "ru * rv < 0"
using polar \<open>?rhs\<close> \<open>u' \<noteq> 0\<close> \<open>v' \<noteq> 0\<close>
using arg_mult_real_negative[of "ru" "cis \<phi>"] arg_mult_real_positive[of "ru" "cis \<phi>"]
using arg_mult_real_negative[of "rv" "cis \<phi>"] arg_mult_real_positive[of "rv" "cis \<phi>"]
apply (cases "ru > 0")
apply (cases "rv > 0", simp, simp add: mult_pos_neg)
apply (cases "rv > 0", simp add: mult_neg_pos, simp)
done
thus "?lhs"
using polar
by (rule_tac x="ru / rv" in exI, auto simp add: divide_less_0_iff mult_less_0_iff)
qed
thus "poincare_between u 0\<^sub>h v = (arg (to_complex u) * arg (to_complex v) < 0)"
using poincare_between_u0v[of u v] * \<open>u \<noteq> 0\<^sub>h\<close> \<open>v \<noteq> 0\<^sub>h\<close> uv
by simp
qed
next
fix a z
assume 1: "is_real a" "cmod a < 1" "z \<in> unit_disc"
assume 2: "?P (moebius_pt (blaschke a) z)"
show "?P z"
proof ((rule allI)+, rule impI, (erule conjE)+)
fix u v
let ?M = "moebius_pt (blaschke a)"
let ?Mu = "?M u"
let ?Mv = "?M v"
assume *: "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v" "u \<notin> circline_set x_axis" "v \<notin> circline_set x_axis"
hence "u \<noteq> \<infinity>\<^sub>h" "v \<noteq> \<infinity>\<^sub>h"
by auto
have **: "\<And> x y :: real. x * y < 0 \<longleftrightarrow> sgn (x * y) < 0"
by simp
assume "z \<in> circline_set (poincare_line u v) \<inter> circline_set x_axis"
thus "poincare_between u z v = (arg (to_complex u) * arg (to_complex v) < 0)"
using * 1 2[rule_format, of ?Mu ?Mv] \<open>cmod a < 1\<close> \<open>is_real a\<close> blaschke_unit_disc_fix[of a]
using inversion_noteq_unit_disc[of "of_complex a" u] \<open>u \<noteq> \<infinity>\<^sub>h\<close>
using inversion_noteq_unit_disc[of "of_complex a" v] \<open>v \<noteq> \<infinity>\<^sub>h\<close>
apply auto
apply (subst (asm) **, subst **, subst (asm) sgn_mult, subst sgn_mult, simp)
apply (subst (asm) **, subst (asm) **, subst (asm) sgn_mult, subst (asm) sgn_mult, simp)
done
qed
next
show "z \<in> unit_disc" by fact
next
show "is_real (to_complex z)"
using assms inf_or_of_complex[of z]
by (auto simp add: circline_set_x_axis)
qed
thus ?thesis
using assms
by simp
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Check if an h-line intersects the x-axis\<close>
(* ------------------------------------------------------------------ *)
lemma x_axis_intersection_equation:
assumes
"H = mk_circline A B C D" and
"(A, B, C, D) \<in> hermitean_nonzero"
shows "of_complex z \<in> circline_set x_axis \<inter> circline_set H \<longleftrightarrow>
A*z\<^sup>2 + 2*Re B*z + D = 0 \<and> is_real z" (is "?lhs \<longleftrightarrow> ?rhs")
proof-
have "?lhs \<longleftrightarrow> A*z\<^sup>2 + (B + cnj B)*z + D = 0 \<and> z = cnj z"
using assms
using circline_equation_x_axis[of z]
using circline_equation[of H A B C D z]
using hermitean_elems
by (auto simp add: power2_eq_square field_simps)
thus ?thesis
using eq_cnj_iff_real[of z]
using hermitean_elems[of A B C D]
by (simp add: complex_add_cnj complex_eq_if_Re_eq)
qed
text \<open>Check if an h-line intersects x-axis within the unit disc - this could be generalized to
checking if an arbitrary circline intersects the x-axis, but we do not need that.\<close>
definition intersects_x_axis_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "intersects_x_axis_cmat H = (let (A, B, C, D) = H in A = 0 \<or> (Re B)\<^sup>2 > (Re A)\<^sup>2)"
lift_definition intersects_x_axis_clmat :: "circline_mat \<Rightarrow> bool" is intersects_x_axis_cmat
done
lift_definition intersects_x_axis :: "circline \<Rightarrow> bool" is intersects_x_axis_clmat
proof (transfer)
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" and "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where k: "k \<noteq> 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
show "intersects_x_axis_cmat H1 = intersects_x_axis_cmat H2"
proof-
have "k \<noteq> 0 \<Longrightarrow> (Re A1)\<^sup>2 < (Re B1)\<^sup>2 \<longleftrightarrow> (k * Re A1)\<^sup>2 < (k * Re B1)\<^sup>2"
by (smt mult_strict_left_mono power2_eq_square semiring_normalization_rules(13) zero_less_power2)
thus ?thesis
using * k
by auto
qed
qed
lemma intersects_x_axis_mk_circline:
assumes "is_real A" and "A \<noteq> 0 \<or> B \<noteq> 0"
shows "intersects_x_axis (mk_circline A B (cnj B) A) \<longleftrightarrow> A = 0 \<or> (Re B)\<^sup>2 > (Re A)\<^sup>2"
proof-
let ?H = "(A, B, (cnj B), A)"
have "hermitean ?H"
- using `is_real A`
+ using \<open>is_real A\<close>
unfolding hermitean_def mat_adj_def mat_cnj_def
using eq_cnj_iff_real
by auto
moreover
have "?H \<noteq> mat_zero"
using assms
by auto
ultimately
show ?thesis
by (transfer, transfer, auto simp add: Let_def)
qed
lemma intersects_x_axis_iff:
assumes "is_poincare_line H"
shows "(\<exists> x \<in> unit_disc. x \<in> circline_set H \<inter> circline_set x_axis) \<longleftrightarrow> intersects_x_axis H"
proof-
obtain Ac B C Dc where *: "H = mk_circline Ac B C Dc" "hermitean (Ac, B, C, Dc)" "(Ac, B, C, Dc) \<noteq> mat_zero"
using ex_mk_circline[of H]
by auto
hence "(cmod B)\<^sup>2 > (cmod Ac)\<^sup>2" "Ac = Dc"
using assms
using is_poincare_line_mk_circline
by auto
hence "H = mk_circline (Re Ac) B (cnj B) (Re Ac)" "hermitean (cor (Re Ac), B, (cnj B), cor (Re Ac))" "(cor (Re Ac), B, (cnj B), cor (Re Ac)) \<noteq> mat_zero"
using hermitean_elems[of Ac B C Dc] *
by auto
then obtain A where
*: "H = mk_circline (cor A) B (cnj B) (cor A)" "(cor A, B, (cnj B), cor A) \<in> hermitean_nonzero"
by auto
show ?thesis
proof (cases "A = 0")
case True
thus ?thesis
using *
using x_axis_intersection_equation[OF *(1-2), of 0]
using intersects_x_axis_mk_circline[of "cor A" B]
by auto
next
case False
show ?thesis
proof
assume "\<exists> x \<in> unit_disc. x \<in> circline_set H \<inter> circline_set x_axis"
then obtain x where **: "of_complex x \<in> unit_disc" "of_complex x \<in> circline_set H \<inter> circline_set x_axis"
by (metis inf_or_of_complex inf_notin_unit_disc)
hence "is_real x"
unfolding circline_set_x_axis
using of_complex_inj
by auto
hence eq: "A * (Re x)\<^sup>2 + 2 * Re B * Re x + A = 0"
using **
using x_axis_intersection_equation[OF *(1-2), of "Re x"]
by simp
hence "(2 * Re B)\<^sup>2 - 4 * A * A \<ge> 0"
using discriminant_iff[of A _ "2 * Re B" A]
using discrim_def[of A "2 * Re B" A] False
by auto
hence "(Re B)\<^sup>2 \<ge> (Re A)\<^sup>2"
by (simp add: power2_eq_square)
moreover
have "(Re B)\<^sup>2 \<noteq> (Re A)\<^sup>2"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "Re B = Re A \<or> Re B = - Re A"
using power2_eq_iff by blast
hence "A * (Re x)\<^sup>2 + A * 2* Re x + A = 0 \<or> A * (Re x)\<^sup>2 - A * 2 * Re x + A = 0"
using eq
by auto
hence "A * ((Re x)\<^sup>2 + 2* Re x + 1) = 0 \<or> A * ((Re x)\<^sup>2 - 2 * Re x + 1) = 0"
by (simp add: field_simps)
hence "(Re x)\<^sup>2 + 2 * Re x + 1 = 0 \<or> (Re x)\<^sup>2 - 2 * Re x + 1 = 0"
using \<open>A \<noteq> 0\<close>
by simp
hence "(Re x + 1)\<^sup>2 = 0 \<or> (Re x - 1)\<^sup>2 = 0"
by (simp add: power2_sum power2_diff field_simps)
hence "Re x = -1 \<or> Re x = 1"
by auto
thus False
using \<open>is_real x\<close> \<open>of_complex x \<in> unit_disc\<close>
by (auto simp add: cmod_eq_Re)
qed
ultimately
show "intersects_x_axis H"
using intersects_x_axis_mk_circline
using *
by auto
next
assume "intersects_x_axis H"
hence "(Re B)\<^sup>2 > (Re A)\<^sup>2"
using * False
using intersects_x_axis_mk_circline
by simp
hence discr: "(2 * Re B)\<^sup>2 - 4 * A * A > 0"
by (simp add: power2_eq_square)
then obtain x1 x2 where
eqs: "A * x1\<^sup>2 + 2 * Re B * x1 + A = 0" "A * x2\<^sup>2 + 2 * Re B * x2 + A = 0" "x1 \<noteq> x2"
using discriminant_pos_ex[OF \<open>A \<noteq> 0\<close>, of "2 * Re B" A]
using discrim_def[of A "2 * Re B" A]
by auto
hence "x1 * x2 = 1"
using viette2[OF \<open>A \<noteq> 0\<close>, of "2 * Re B" A x1 x2] discr \<open>A \<noteq> 0\<close>
by auto
have "abs x1 \<noteq> 1" "abs x2 \<noteq> 1"
using eqs discr \<open>x1 * x2 = 1\<close>
by (auto simp add: abs_if power2_eq_square)
hence "abs x1 < 1 \<or> abs x2 < 1"
using \<open>x1 * x2 = 1\<close>
by (smt mult_le_cancel_left1 mult_minus_right)
thus "\<exists>x \<in> unit_disc. x \<in> circline_set H \<inter> circline_set x_axis"
using x_axis_intersection_equation[OF *(1-2), of x1]
using x_axis_intersection_equation[OF *(1-2), of x2]
using eqs
by auto
qed
qed
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Check if a Poincar\'e line intersects the y-axis\<close>
(* ------------------------------------------------------------------ *)
definition intersects_y_axis_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "intersects_y_axis_cmat H = (let (A, B, C, D) = H in A = 0 \<or> (Im B)\<^sup>2 > (Re A)\<^sup>2)"
lift_definition intersects_y_axis_clmat :: "circline_mat \<Rightarrow> bool" is intersects_y_axis_cmat
done
lift_definition intersects_y_axis :: "circline \<Rightarrow> bool" is intersects_y_axis_clmat
proof (transfer)
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" and "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where k: "k \<noteq> 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
show "intersects_y_axis_cmat H1 = intersects_y_axis_cmat H2"
proof-
have "k \<noteq> 0 \<Longrightarrow> (Re A1)\<^sup>2 < (Im B1)\<^sup>2 \<longleftrightarrow> (k * Re A1)\<^sup>2 < (k * Im B1)\<^sup>2"
by (smt mult_strict_left_mono power2_eq_square semiring_normalization_rules(13) zero_less_power2)
thus ?thesis
using * k
by auto
qed
qed
lemma intersects_x_axis_intersects_y_axis [simp]:
shows "intersects_x_axis (moebius_circline (moebius_rotation (pi/2)) H) \<longleftrightarrow> intersects_y_axis H"
unfolding moebius_rotation_def moebius_similarity_def
by simp (transfer, transfer, auto simp add: mat_adj_def mat_cnj_def)
lemma intersects_y_axis_iff:
assumes "is_poincare_line H"
shows "(\<exists> y \<in> unit_disc. y \<in> circline_set H \<inter> circline_set y_axis) \<longleftrightarrow> intersects_y_axis H" (is "?lhs \<longleftrightarrow> ?rhs")
proof-
let ?R = "moebius_rotation (pi / 2)"
let ?H' = "moebius_circline ?R H"
have 1: "is_poincare_line ?H'"
using assms
using unit_circle_fix_preserve_is_poincare_line[OF _ assms, of ?R]
by simp
show ?thesis
proof
assume "?lhs"
then obtain y where "y \<in> unit_disc" "y \<in> circline_set H \<inter> circline_set y_axis"
by auto
hence "moebius_pt ?R y \<in> unit_disc \<and> moebius_pt ?R y \<in> circline_set ?H' \<inter> circline_set x_axis"
using rotation_pi_2_y_axis
by (metis Int_iff circline_set_moebius_circline_E moebius_circline_comp_inv_left moebius_pt_comp_inv_left unit_disc_fix_discI unit_disc_fix_rotation)
thus ?rhs
using intersects_x_axis_iff[OF 1]
using intersects_x_axis_intersects_y_axis[of H]
by auto
next
assume "intersects_y_axis H"
hence "intersects_x_axis ?H'"
using intersects_x_axis_intersects_y_axis[of H]
by simp
then obtain x where *: "x \<in> unit_disc" "x \<in> circline_set ?H' \<inter> circline_set x_axis"
using intersects_x_axis_iff[OF 1]
by auto
let ?y = "moebius_pt (-?R) x"
have "?y \<in> unit_disc \<and> ?y \<in> circline_set H \<inter> circline_set y_axis"
using * rotation_pi_2_y_axis[symmetric]
by (metis Int_iff circline_set_moebius_circline_E moebius_pt_comp_inv_left moebius_rotation_uminus uminus_moebius_def unit_disc_fix_discI unit_disc_fix_rotation)
thus ?lhs
by auto
qed
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Intersection point of a Poincar\'e line with the x-axis in the unit disc\<close>
(* ------------------------------------------------------------------ *)
definition calc_x_axis_intersection_cvec :: "complex \<Rightarrow> complex \<Rightarrow> complex_vec" where
[simp]: "calc_x_axis_intersection_cvec A B =
(let discr = (Re B)\<^sup>2 - (Re A)\<^sup>2 in
(-Re(B) + sgn (Re B) * sqrt(discr), A))"
(* intersection with the x-axis for poincare lines that are euclidean circles *)
definition calc_x_axis_intersection_cmat_cvec :: "complex_mat \<Rightarrow> complex_vec" where [simp]:
"calc_x_axis_intersection_cmat_cvec H =
(let (A, B, C, D) = H in
if A \<noteq> 0 then
calc_x_axis_intersection_cvec A B
else
(0, 1)
)"
lift_definition calc_x_axis_intersection_clmat_hcoords :: "circline_mat \<Rightarrow> complex_homo_coords" is calc_x_axis_intersection_cmat_cvec
by (auto split: if_split_asm)
lift_definition calc_x_axis_intersection :: "circline \<Rightarrow> complex_homo" is calc_x_axis_intersection_clmat_hcoords
proof transfer
fix H1 H2
assume *: "hermitean H1 \<and> H1 \<noteq> mat_zero" "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where hh: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where k: "k \<noteq> 0" "H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
have "calc_x_axis_intersection_cvec A1 B1 \<approx>\<^sub>v calc_x_axis_intersection_cvec A2 B2"
using hh k
apply simp
apply (rule_tac x="cor k" in exI)
apply auto
apply (simp add: sgn_mult power_mult_distrib)
apply (subst right_diff_distrib[symmetric])
apply (subst real_sqrt_mult)
apply (subst cor_mult)
by (simp add: real_sgn_eq right_diff_distrib)
thus "calc_x_axis_intersection_cmat_cvec H1 \<approx>\<^sub>v
calc_x_axis_intersection_cmat_cvec H2"
using hh k
by (auto simp del: calc_x_axis_intersection_cvec_def)
qed
lemma calc_x_axis_intersection_in_unit_disc:
assumes "is_poincare_line H" "intersects_x_axis H"
shows "calc_x_axis_intersection H \<in> unit_disc"
proof (cases "is_line H")
case True
thus ?thesis
using assms
unfolding unit_disc_def disc_def
by simp (transfer, transfer, auto simp add: vec_cnj_def)
next
case False
thus ?thesis
using assms
unfolding unit_disc_def disc_def
proof (simp, transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
then obtain A B D where *: "H = (A, B, cnj B, D)" "is_real A" "is_real D"
using hermitean_elems
by (cases H) blast
assume "is_poincare_line_cmat H"
hence *: "H = (A, B, cnj B, A)" "is_real A"
using *
by auto
assume "\<not> circline_A0_cmat H"
hence "A \<noteq> 0"
using *
by simp
assume "intersects_x_axis_cmat H"
hence "(Re B)\<^sup>2 > (Re A)\<^sup>2"
using * \<open>A \<noteq> 0\<close>
by (auto simp add: power2_eq_square complex.expand)
hence "Re B \<noteq> 0"
by auto
have "Re A \<noteq> 0"
using \<open>is_real A\<close> \<open>A \<noteq> 0\<close>
by (auto simp add: complex.expand)
have "sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) < sqrt((Re B)\<^sup>2)"
using \<open>Re A \<noteq> 0\<close>
by (subst real_sqrt_less_iff) auto
also have "... = sgn (Re B) * (Re B)"
by (smt mult_minus_right nonzero_eq_divide_eq real_sgn_eq real_sqrt_abs)
finally
have 1: "sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) < sgn (Re B) * (Re B)"
.
have 2: "(Re B)\<^sup>2 - (Re A)\<^sup>2 < sgn (Re B) * (Re B) * sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
using \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close>
using mult_strict_right_mono[OF 1, of "sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)"]
by simp
have 3: "(Re B)\<^sup>2 - 2*sgn (Re B)*Re B*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2) + (Re B)\<^sup>2 - (Re A)\<^sup>2 < (Re A)\<^sup>2"
using mult_strict_left_mono[OF 2, of 2]
by (simp add: field_simps)
have "(sgn (Re B))\<^sup>2 = 1"
using \<open>Re B \<noteq> 0\<close>
by (simp add: sgn_if)
hence "(-Re B + sgn (Re B) * sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 < (Re A)\<^sup>2"
using \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close> 3
by (simp add: power2_diff field_simps)
thus "in_ocircline_cmat_cvec unit_circle_cmat (calc_x_axis_intersection_cmat_cvec H)"
using * \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close>
by (auto simp add: vec_cnj_def power2_eq_square split: if_split_asm)
qed
qed
lemma calc_x_axis_intersection:
assumes "is_poincare_line H" and "intersects_x_axis H"
shows "calc_x_axis_intersection H \<in> circline_set H \<inter> circline_set x_axis"
proof (cases "is_line H")
case True
thus ?thesis
using assms
unfolding circline_set_def
by simp (transfer, transfer, auto simp add: vec_cnj_def)
next
case False
thus ?thesis
using assms
unfolding circline_set_def
proof (simp, transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
then obtain A B D where *: "H = (A, B, cnj B, D)" "is_real A" "is_real D"
using hermitean_elems
by (cases H) blast
assume "is_poincare_line_cmat H"
hence *: "H = (A, B, cnj B, A)" "is_real A"
using *
by auto
assume "\<not> circline_A0_cmat H"
hence "A \<noteq> 0"
using *
by auto
assume "intersects_x_axis_cmat H"
hence "(Re B)\<^sup>2 > (Re A)\<^sup>2"
using * \<open>A \<noteq> 0\<close>
by (auto simp add: power2_eq_square complex.expand)
hence "Re B \<noteq> 0"
by auto
show "on_circline_cmat_cvec H (calc_x_axis_intersection_cmat_cvec H) \<and>
on_circline_cmat_cvec x_axis_cmat (calc_x_axis_intersection_cmat_cvec H)" (is "?P1 \<and> ?P2")
proof
show "on_circline_cmat_cvec H (calc_x_axis_intersection_cmat_cvec H)"
proof (cases "circline_A0_cmat H")
case True
thus ?thesis
using * \<open>is_poincare_line_cmat H\<close> \<open>intersects_x_axis_cmat H\<close>
by (simp add: vec_cnj_def)
next
case False
let ?x = "calc_x_axis_intersection_cvec A B"
let ?nom = "fst ?x" and ?den = "snd ?x"
have x: "?x = (?nom, ?den)"
by simp
hence "on_circline_cmat_cvec H (calc_x_axis_intersection_cvec A B)"
proof (subst *, subst x, subst on_circline_cmat_cvec_circline_equation)
have "(sgn(Re B))\<^sup>2 = 1"
using \<open>Re B \<noteq> 0\<close> sgn_pos zero_less_power2 by fastforce
have "(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)\<^sup>2 - (Re A)\<^sup>2"
using \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close>
by simp
have "(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 =
(-(Re B))\<^sup>2 + (sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
by (simp add: power2_diff)
also have "... = (Re B)*(Re B) + (sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
by (simp add: power2_eq_square)
also have "... = (Re B)*(Re B) + (sgn(Re B))\<^sup>2*(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
by (simp add: power_mult_distrib)
also have "... = (Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
using \<open>(sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 = (Re B)\<^sup>2 - (Re A)\<^sup>2\<close> \<open>(sgn(Re B))\<^sup>2 = 1\<close>
by simp
finally have "(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 =
(Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)"
by simp
have "is_real ?nom" "is_real ?den"
using \<open>is_real A\<close>
by simp+
hence "cnj (?nom) = ?nom" "cnj (?den) = ?den"
by (simp add:eq_cnj_iff_real)+
hence "A*?nom*(cnj (?nom)) + B*?den*(cnj (?nom)) + (cnj B)*(cnj (?den))*?nom + A*?den*(cnj (?den))
= A*?nom*?nom + B*?den*?nom + (cnj B)*?den*?nom + A*?den*?den"
by auto
also have "... = A*?nom*?nom + (B + (cnj B))*?den*?nom + A*?den*?den"
by (simp add:field_simps)
also have "... = A*?nom*?nom + 2*(Re B)*?den*?nom + A*?den*?den"
by (simp add:complex_add_cnj)
also have "... = A*?nom\<^sup>2 + 2*(Re B)*?den*?nom + A*?den*?den"
by (simp add:power2_eq_square)
also have "... = A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2
+ 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A"
unfolding calc_x_axis_intersection_cvec_def
by auto
also have "... = A*((Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))
+ 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A"
using \<open>(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))\<^sup>2 =
(Re B)*(Re B) + (Re B)\<^sup>2 - (Re A)\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)\<close>
by simp
also have "... = A*((Re B)*(Re B) + (Re B)\<^sup>2 - A\<^sup>2 - 2*(Re B)*sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2))
+ 2*(Re B)*A*(-(Re B) + sgn(Re B)*sqrt((Re B)\<^sup>2 - (Re A)\<^sup>2)) + A*A*A"
using \<open>is_real A\<close>
by simp
also have "... = 0"
apply (simp add:field_simps)
by (simp add: power2_eq_square)
finally have "A*?nom*(cnj (?nom)) + B*?den*(cnj (?nom)) + (cnj B)*(cnj (?den))*?nom + A*?den*(cnj (?den)) = 0"
by simp
thus "circline_equation A B (cnj B) A ?nom ?den"
by simp
qed
thus ?thesis
using * \<open>is_poincare_line_cmat H\<close> \<open>intersects_x_axis_cmat H\<close>
by (simp add: vec_cnj_def)
qed
next
show "on_circline_cmat_cvec x_axis_cmat (calc_x_axis_intersection_cmat_cvec H)"
using * \<open>is_poincare_line_cmat H\<close> \<open>intersects_x_axis_cmat H\<close> \<open>is_real A\<close>
using eq_cnj_iff_real[of A]
by (simp add: vec_cnj_def)
qed
qed
qed
lemma unique_calc_x_axis_intersection:
assumes "is_poincare_line H" and "H \<noteq> x_axis"
assumes "x \<in> unit_disc" and "x \<in> circline_set H \<inter> circline_set x_axis"
shows "x = calc_x_axis_intersection H"
proof-
have *: "intersects_x_axis H"
using assms
using intersects_x_axis_iff[OF assms(1)]
by auto
show "x = calc_x_axis_intersection H"
using calc_x_axis_intersection[OF assms(1) *]
using calc_x_axis_intersection_in_unit_disc[OF assms(1) *]
using assms
using unique_is_poincare_line[of x "calc_x_axis_intersection H" H x_axis]
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Check if an h-line intersects the positive part of the x-axis\<close>
(* ------------------------------------------------------------------ *)
definition intersects_x_axis_positive_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "intersects_x_axis_positive_cmat H = (let (A, B, C, D) = H in Re A \<noteq> 0 \<and> Re B / Re A < -1)"
lift_definition intersects_x_axis_positive_clmat :: "circline_mat \<Rightarrow> bool" is intersects_x_axis_positive_cmat
done
lift_definition intersects_x_axis_positive :: "circline \<Rightarrow> bool" is intersects_x_axis_positive_clmat
proof (transfer)
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" and "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where "k \<noteq> 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
thus "intersects_x_axis_positive_cmat H1 = intersects_x_axis_positive_cmat H2"
using *
by simp
qed
lemma intersects_x_axis_positive_mk_circline:
assumes "is_real A" and "A \<noteq> 0 \<or> B \<noteq> 0"
shows "intersects_x_axis_positive (mk_circline A B (cnj B) A) \<longleftrightarrow> Re B / Re A < -1"
proof-
let ?H = "(A, B, (cnj B), A)"
have "hermitean ?H"
- using `is_real A`
+ using \<open>is_real A\<close>
unfolding hermitean_def mat_adj_def mat_cnj_def
using eq_cnj_iff_real
by auto
moreover
have "?H \<noteq> mat_zero"
using assms
by auto
ultimately
show ?thesis
by (transfer, transfer, auto simp add: Let_def)
qed
lemma intersects_x_axis_positive_intersects_x_axis [simp]:
assumes "intersects_x_axis_positive H"
shows "intersects_x_axis H"
proof-
have "\<And> a aa. \<lbrakk> Re a \<noteq> 0; Re aa / Re a < - 1; \<not> (Re a)\<^sup>2 < (Re aa)\<^sup>2 \<rbrakk> \<Longrightarrow> aa = 0 \<and> a = 0"
by (smt less_divide_eq_1_pos one_less_power pos2 power2_minus power_divide zero_less_power2)
thus ?thesis
using assms
apply transfer
apply transfer
apply (auto simp add: hermitean_def mat_adj_def mat_cnj_def)
done
qed
lemma add_less_abs_positive_iff:
fixes a b :: real
assumes "abs b < abs a"
shows "a + b > 0 \<longleftrightarrow> a > 0"
using assms
by auto
lemma calc_x_axis_intersection_positive_abs':
fixes A B :: real
assumes "B\<^sup>2 > A\<^sup>2" and "A \<noteq> 0"
shows "abs (sgn(B) * sqrt(B\<^sup>2 - A\<^sup>2) / A) < abs(-B/A)"
proof-
from assms have "B \<noteq> 0"
by auto
have "B\<^sup>2 - A\<^sup>2 < B\<^sup>2"
using \<open>A \<noteq> 0\<close>
by auto
hence "sqrt (B\<^sup>2 - A\<^sup>2) < abs B"
using real_sqrt_less_iff[of "B\<^sup>2 - A\<^sup>2" "B\<^sup>2"]
by simp
thus ?thesis
using assms \<open>B \<noteq> 0\<close>
by (simp add: abs_mult divide_strict_right_mono)
qed
lemma calc_intersect_x_axis_positive_lemma:
assumes "B\<^sup>2 > A\<^sup>2" and "A \<noteq> 0"
shows "(-B + sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A > 0 \<longleftrightarrow> -B/A > 1"
proof-
have "(-B + sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A = -B / A + (sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A"
using assms
by (simp add: field_simps)
moreover
have "-B / A + (sgn B * sqrt(B\<^sup>2 - A\<^sup>2)) / A > 0 \<longleftrightarrow> - B / A > 0"
using add_less_abs_positive_iff[OF calc_x_axis_intersection_positive_abs'[OF assms]]
by simp
moreover
hence "(B/A)\<^sup>2 > 1"
using assms
by (simp add: power_divide)
hence "B/A > 1 \<or> B/A < -1"
by (smt one_power2 pos2 power2_minus power_0 power_strict_decreasing zero_power2)
hence "-B / A > 0 \<longleftrightarrow> -B / A > 1"
by auto
ultimately
show ?thesis
using assms
by auto
qed
lemma intersects_x_axis_positive_iff':
assumes "is_poincare_line H"
shows "intersects_x_axis_positive H \<longleftrightarrow>
calc_x_axis_intersection H \<in> unit_disc \<and> calc_x_axis_intersection H \<in> circline_set H \<inter> positive_x_axis" (is "?lhs \<longleftrightarrow> ?rhs")
proof
let ?x = "calc_x_axis_intersection H"
assume ?lhs
hence "?x \<in> circline_set x_axis" "?x \<in> circline_set H" "?x \<in> unit_disc"
using calc_x_axis_intersection_in_unit_disc[OF assms] calc_x_axis_intersection[OF assms]
by auto
moreover
have "Re (to_complex ?x) > 0"
using \<open>?lhs\<close> assms
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
assume "intersects_x_axis_positive_cmat H"
hence **: "Re B / Re A < - 1" "Re A \<noteq> 0"
using *
by auto
have "(Re B)\<^sup>2 > (Re A)\<^sup>2"
using **
by (smt divide_less_eq_1_neg divide_minus_left less_divide_eq_1_pos real_sqrt_abs real_sqrt_less_iff right_inverse_eq)
have "is_real A" "A \<noteq> 0"
using hh hermitean_elems * \<open>Re A \<noteq> 0\<close> complex.expand[of A 0]
by auto
have "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
using \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close> \<open>is_real A\<close>
by (smt cmod_power2 power2_less_0 zero_power2)
have ***: "0 < (- Re B + sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) / Re A"
using calc_intersect_x_axis_positive_lemma[of "Re A" "Re B"] ** \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close>
by auto
assume "is_poincare_line_cmat H"
hence "A = D"
using * hh
by simp
have "Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) - cor (Re B)) / A) = (sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2) - Re B) / Re D"
using \<open>is_real A\<close> \<open>A = D\<close>
by (metis (no_types, lifting) Re_complex_of_real complex_of_real_Re of_real_diff of_real_divide of_real_mult)
thus "0 < Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H))"
using * hh ** *** \<open>(cmod B)\<^sup>2 > (cmod A)\<^sup>2\<close> \<open>(Re B)\<^sup>2 > (Re A)\<^sup>2\<close> \<open>A \<noteq> 0\<close> \<open>A = D\<close>
by simp
qed
ultimately
show ?rhs
unfolding positive_x_axis_def
by auto
next
let ?x = "calc_x_axis_intersection H"
assume ?rhs
hence "Re (to_complex ?x) > 0" "?x \<noteq> \<infinity>\<^sub>h" "?x \<in> circline_set x_axis" "?x \<in> unit_disc" "?x \<in> circline_set H"
unfolding positive_x_axis_def
by auto
hence "intersects_x_axis H"
using intersects_x_axis_iff[OF assms]
by auto
thus ?lhs
using \<open>Re (to_complex ?x) > 0\<close> assms
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
assume "0 < Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H))" "intersects_x_axis_cmat H" "is_poincare_line_cmat H"
hence **: "A \<noteq> 0" "0 < Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re A)\<^sup>2)) - cor (Re B)) / A)" "A = D" "is_real A" "(Re B)\<^sup>2 > (Re A)\<^sup>2"
using * hh hermitean_elems
by (auto split: if_split_asm)
have "Re A \<noteq> 0"
using complex.expand[of A 0] \<open>A \<noteq> 0\<close> \<open>is_real A\<close>
by auto
have "Re ((cor (sgn (Re B)) * cor (sqrt ((Re B)\<^sup>2 - (Re D)\<^sup>2)) - cor (Re B)) / D) = (sgn (Re B) * sqrt ((Re B)\<^sup>2 - (Re D)\<^sup>2) - Re B) / Re D"
using \<open>is_real A\<close> \<open>A = D\<close>
by (metis (no_types, lifting) Re_complex_of_real complex_of_real_Re of_real_diff of_real_divide of_real_mult)
thus "intersects_x_axis_positive_cmat H"
using * ** \<open>Re A \<noteq> 0\<close>
using calc_intersect_x_axis_positive_lemma[of "Re A" "Re B"]
by simp
qed
qed
lemma intersects_x_axis_positive_iff:
assumes "is_poincare_line H" and "H \<noteq> x_axis"
shows "intersects_x_axis_positive H \<longleftrightarrow>
(\<exists> x. x \<in> unit_disc \<and> x \<in> circline_set H \<inter> positive_x_axis)" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
thus ?rhs
using intersects_x_axis_positive_iff'[OF assms(1)]
by auto
next
assume ?rhs
then obtain x where "x \<in> unit_disc" "x \<in> circline_set H \<inter> positive_x_axis"
by auto
thus ?lhs
using unique_calc_x_axis_intersection[OF assms, of x]
using intersects_x_axis_positive_iff'[OF assms(1)]
unfolding positive_x_axis_def
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Check if an h-line intersects the positive part of the y-axis\<close>
(* ------------------------------------------------------------------ *)
definition intersects_y_axis_positive_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "intersects_y_axis_positive_cmat H = (let (A, B, C, D) = H in Re A \<noteq> 0 \<and> Im B / Re A < -1)"
lift_definition intersects_y_axis_positive_clmat :: "circline_mat \<Rightarrow> bool" is intersects_y_axis_positive_cmat
done
lift_definition intersects_y_axis_positive :: "circline \<Rightarrow> bool" is intersects_y_axis_positive_clmat
proof (transfer)
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" and "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where "k \<noteq> 0 \<and> H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
thus "intersects_y_axis_positive_cmat H1 = intersects_y_axis_positive_cmat H2"
using *
by simp
qed
lemma intersects_x_axis_positive_intersects_y_axis_positive [simp]:
shows "intersects_x_axis_positive (moebius_circline (moebius_rotation (-pi/2)) H) \<longleftrightarrow> intersects_y_axis_positive H"
using hermitean_elems
unfolding moebius_rotation_def moebius_similarity_def
by simp (transfer, transfer, auto simp add: mat_adj_def mat_cnj_def)
lemma intersects_y_axis_positive_iff:
assumes "is_poincare_line H" "H \<noteq> y_axis"
shows "(\<exists> y \<in> unit_disc. y \<in> circline_set H \<inter> positive_y_axis) \<longleftrightarrow> intersects_y_axis_positive H" (is "?lhs \<longleftrightarrow> ?rhs")
proof-
let ?R = "moebius_rotation (-pi / 2)"
let ?H' = "moebius_circline ?R H"
have 1: "is_poincare_line ?H'"
using assms
using unit_circle_fix_preserve_is_poincare_line[OF _ assms(1), of ?R]
by simp
have 2: "moebius_circline ?R H \<noteq> x_axis"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "H = moebius_circline (moebius_rotation (pi/2)) x_axis"
using moebius_circline_comp_inv_left[of ?R H]
by auto
thus False
using \<open>H \<noteq> y_axis\<close>
by auto
qed
show ?thesis
proof
assume "?lhs"
then obtain y where "y \<in> unit_disc" "y \<in> circline_set H \<inter> positive_y_axis"
by auto
hence "moebius_pt ?R y \<in> unit_disc" "moebius_pt ?R y \<in> circline_set ?H' \<inter> positive_x_axis"
using rotation_minus_pi_2_positive_y_axis
by auto
thus ?rhs
using intersects_x_axis_positive_iff[OF 1 2]
using intersects_x_axis_positive_intersects_y_axis_positive[of H]
by auto
next
assume "intersects_y_axis_positive H"
hence "intersects_x_axis_positive ?H'"
using intersects_x_axis_positive_intersects_y_axis_positive[of H]
by simp
then obtain x where *: "x \<in> unit_disc" "x \<in> circline_set ?H' \<inter> positive_x_axis"
using intersects_x_axis_positive_iff[OF 1 2]
by auto
let ?y = "moebius_pt (-?R) x"
have "?y \<in> unit_disc \<and> ?y \<in> circline_set H \<inter> positive_y_axis"
using * rotation_minus_pi_2_positive_y_axis[symmetric]
by (metis Int_iff circline_set_moebius_circline_E image_eqI moebius_pt_comp_inv_image_left moebius_rotation_uminus uminus_moebius_def unit_disc_fix_discI unit_disc_fix_rotation)
thus ?lhs
by auto
qed
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Position of the intersection point in the unit disc\<close>
(* ------------------------------------------------------------------ *)
text\<open>Check if the intersection point of one h-line with the x-axis is located more outward the edge
of the disc than the intersection point of another h-line.\<close>
definition outward_cmat :: "complex_mat \<Rightarrow> complex_mat \<Rightarrow> bool" where
[simp]: "outward_cmat H1 H2 = (let (A1, B1, C1, D1) = H1; (A2, B2, C2, D2) = H2
in -Re B1/Re A1 \<le> -Re B2/Re A2)"
lift_definition outward_clmat :: "circline_mat \<Rightarrow> circline_mat \<Rightarrow> bool" is outward_cmat
done
lift_definition outward :: "circline \<Rightarrow> circline \<Rightarrow> bool" is outward_clmat
apply transfer
apply simp
apply (case_tac circline_mat1, case_tac circline_mat2, case_tac circline_mat3, case_tac circline_mat4)
apply simp
apply (erule_tac exE)+
apply (erule_tac conjE)+
apply simp
done
lemma outward_mk_circline:
assumes "is_real A1" and "is_real A2" and "A1 \<noteq> 0 \<or> B1 \<noteq> 0" and "A2 \<noteq> 0 \<or> B2 \<noteq> 0"
shows "outward (mk_circline A1 B1 (cnj B1) A1) (mk_circline A2 B2 (cnj B2) A2) \<longleftrightarrow> - Re B1 / Re A1 \<le> - Re B2 / Re A2"
proof-
let ?H1 = "(A1, B1, (cnj B1), A1)"
let ?H2 = "(A2, B2, (cnj B2), A2)"
have "hermitean ?H1" "hermitean ?H2"
- using `is_real A1` `is_real A2`
+ using \<open>is_real A1\<close> \<open>is_real A2\<close>
unfolding hermitean_def mat_adj_def mat_cnj_def
using eq_cnj_iff_real
by auto
moreover
have "?H1 \<noteq> mat_zero" "?H2 \<noteq> mat_zero"
using assms
by auto
ultimately
show ?thesis
by (transfer, transfer, auto simp add: Let_def)
qed
lemma calc_x_axis_intersection_fun_mono:
fixes x1 x2 :: real
assumes "x1 > 1" and "x2 > x1"
shows "x1 - sqrt(x1\<^sup>2 - 1) > x2 - sqrt(x2\<^sup>2 - 1)"
using assms
proof-
have *: "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1) > 0"
using assms
by (smt one_less_power pos2 real_sqrt_gt_zero)
have "sqrt(x1\<^sup>2 - 1) < x1"
using real_sqrt_less_iff[of "x1\<^sup>2 - 1" "x1\<^sup>2"] \<open>x1 > 1\<close>
by auto
moreover
have "sqrt(x2\<^sup>2 - 1) < x2"
using real_sqrt_less_iff[of "x2\<^sup>2 - 1" "x2\<^sup>2"] \<open>x1 > 1\<close> \<open>x2 > x1\<close>
by auto
ultimately
have "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1) < x1 + x2"
by simp
hence "(x1 + x2) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) > 1"
using *
using less_divide_eq_1_pos[of "sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)" "x1 + x2"]
by simp
hence "(x2\<^sup>2 - x1\<^sup>2) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) > x2 - x1"
using \<open>x2 > x1\<close>
using mult_less_cancel_left_pos[of "x2 - x1" 1 "(x2 + x1) / (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1))"]
by (simp add: power2_eq_square field_simps)
moreover
have "(x2\<^sup>2 - x1\<^sup>2) = (sqrt(x1\<^sup>2 - 1) + sqrt(x2\<^sup>2 - 1)) * ((sqrt(x2\<^sup>2 - 1) - sqrt(x1\<^sup>2 - 1)))"
using \<open>x1 > 1\<close> \<open>x2 > x1\<close>
by (simp add: field_simps)
ultimately
have "sqrt(x2\<^sup>2 - 1) - sqrt(x1\<^sup>2 - 1) > x2 - x1"
using *
by simp
thus ?thesis
by simp
qed
lemma calc_x_axis_intersection_mono:
fixes a1 b1 a2 b2 :: real
assumes "-b1/a1 > 1" and "a1 \<noteq> 0" and "-b2/a2 \<ge> -b1/a1" and "a2 \<noteq> 0"
shows "(-b1 + sgn b1 * sqrt(b1\<^sup>2 - a1\<^sup>2)) / a1 \<ge> (-b2 + sgn b2 * sqrt(b2\<^sup>2 - a2\<^sup>2)) / a2" (is "?lhs \<ge> ?rhs")
proof-
have "?lhs = -b1/a1 - sqrt((-b1/a1)\<^sup>2 - 1)"
proof (cases "b1 > 0")
case True
hence "a1 < 0"
using assms
by (smt divide_neg_pos)
thus ?thesis
using \<open>b1 > 0\<close> \<open>a1 < 0\<close>
by (simp add: real_sqrt_divide field_simps)
next
case False
hence "b1 < 0"
using assms
by (cases "b1 = 0") auto
hence "a1 > 0"
using assms
by (smt divide_pos_neg)
thus ?thesis
using \<open>b1 < 0\<close> \<open>a1 > 0\<close>
by (simp add: real_sqrt_divide field_simps)
qed
moreover
have "?rhs = -b2/a2 - sqrt((-b2/a2)\<^sup>2 - 1)"
proof (cases "b2 > 0")
case True
hence "a2 < 0"
using assms
by (smt divide_neg_pos)
thus ?thesis
using \<open>b2 > 0\<close> \<open>a2 < 0\<close>
by (simp add: real_sqrt_divide field_simps)
next
case False
hence "b2 < 0"
using assms
by (cases "b2 = 0") auto
hence "a2 > 0"
using assms
by (smt divide_pos_neg)
thus ?thesis
using \<open>b2 < 0\<close> \<open>a2 > 0\<close>
by (simp add: real_sqrt_divide field_simps)
qed
ultimately
show ?thesis
using calc_x_axis_intersection_fun_mono[of "-b1/a1" "-b2/a2"]
using assms
by (cases "-b1/a1=-b2/a2", auto)
qed
lemma outward:
assumes "is_poincare_line H1" and "is_poincare_line H2"
assumes "intersects_x_axis_positive H1" and "intersects_x_axis_positive H2"
assumes "outward H1 H2"
shows "Re (to_complex (calc_x_axis_intersection H1)) \<ge> Re (to_complex (calc_x_axis_intersection H2))"
proof-
have "intersects_x_axis H1" "intersects_x_axis H2"
using assms
by auto
thus ?thesis
using assms
proof (transfer, transfer)
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
have "is_real A1" "is_real A2"
using hermitean_elems * hh
by auto
assume 1: "intersects_x_axis_positive_cmat H1" "intersects_x_axis_positive_cmat H2"
assume 2: "intersects_x_axis_cmat H1" "intersects_x_axis_cmat H2"
assume 3: "is_poincare_line_cmat H1" "is_poincare_line_cmat H2"
assume 4: "outward_cmat H1 H2"
have "A1 \<noteq> 0" "A2 \<noteq> 0"
using * \<open>is_real A1\<close> \<open>is_real A2\<close> 1 complex.expand[of A1 0] complex.expand[of A2 0]
by auto
hence "(sgn (Re B2) * sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2) - Re B2) / Re A2
\<le> (sgn (Re B1) * sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2) - Re B1) / Re A1"
using calc_x_axis_intersection_mono[of "Re B1" "Re A1" "Re B2" "Re A2"]
using 1 4 *
by simp
moreover
have "(sgn (Re B2) * sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2) - Re B2) / Re A2 =
Re ((cor (sgn (Re B2)) * cor (sqrt ((Re B2)\<^sup>2 - (Re A2)\<^sup>2)) - cor (Re B2)) / A2)"
using \<open>is_real A2\<close> \<open>A2 \<noteq> 0\<close>
by (simp add: Re_divide_real)
moreover
have "(sgn (Re B1) * sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2) - Re B1) / Re A1 =
Re ((cor (sgn (Re B1)) * cor (sqrt ((Re B1)\<^sup>2 - (Re A1)\<^sup>2)) - cor (Re B1)) / A1)"
using \<open>is_real A1\<close> \<open>A1 \<noteq> 0\<close>
by (simp add: Re_divide_real)
ultimately
show "Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H2))
\<le> Re (to_complex_cvec (calc_x_axis_intersection_cmat_cvec H1))"
using 2 3 \<open>A1 \<noteq> 0\<close> \<open>A2 \<noteq> 0\<close> * \<open>is_real A1\<close> \<open>is_real A2\<close>
by (simp del: is_poincare_line_cmat_def intersects_x_axis_cmat_def)
qed
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Ideal points and x-axis intersection\<close>
(* ------------------------------------------------------------------ *)
lemma ideal_points_intersects_x_axis:
assumes "is_poincare_line H" and "ideal_points H = {i1, i2}" and "H \<noteq> x_axis"
shows "intersects_x_axis H \<longleftrightarrow> Im (to_complex i1) * Im (to_complex i2) < 0"
using assms
proof-
have "i1 \<noteq> i2"
using assms(1) assms(2) ex_poincare_line_points ideal_points_different(1)
by blast
have "calc_ideal_points H = {i1, i2}"
using assms
using ideal_points_unique
by auto
have "\<forall> i1 \<in> calc_ideal_points H.
\<forall> i2 \<in> calc_ideal_points H.
is_poincare_line H \<and> H \<noteq> x_axis \<and> i1 \<noteq> i2 \<longrightarrow> (Im (to_complex i1) * Im (to_complex i2) < 0 \<longleftrightarrow> intersects_x_axis H)"
proof (transfer, transfer, (rule ballI)+, rule impI, (erule conjE)+, case_tac H, case_tac i1, case_tac i2)
fix i11 i12 i21 i22 A B C D H i1 i2
assume H: "H = (A, B, C, D)" "hermitean H" "H \<noteq> mat_zero"
assume line: "is_poincare_line_cmat H"
assume i1: "i1 = (i11, i12)" "i1 \<in> calc_ideal_points_cmat_cvec H"
assume i2: "i2 = (i21, i22)" "i2 \<in> calc_ideal_points_cmat_cvec H"
assume different: "\<not> i1 \<approx>\<^sub>v i2"
assume not_x_axis: "\<not> circline_eq_cmat H x_axis_cmat"
have "is_real A" "is_real D" "C = cnj B"
using H hermitean_elems
by auto
have "(cmod A)\<^sup>2 < (cmod B)\<^sup>2" "A = D"
using line H
by auto
let ?discr = "sqrt ((cmod B)\<^sup>2 - (Re D)\<^sup>2)"
let ?den = "(cmod B)\<^sup>2"
let ?i1 = "B * (- D - \<i> * ?discr)"
let ?i2 = "B * (- D + \<i> * ?discr)"
have "i11 = ?i1 \<or> i11 = ?i2" "i12 = ?den"
"i21 = ?i1 \<or> i21 = ?i2" "i22 = ?den"
using i1 i2 H line
by (auto split: if_split_asm)
hence i: "i11 = ?i1 \<and> i21 = ?i2 \<or> i11 = ?i2 \<and> i21 = ?i1"
- using `\<not> i1 \<approx>\<^sub>v i2` i1 i2
+ using \<open>\<not> i1 \<approx>\<^sub>v i2\<close> i1 i2
by auto
have "Im (i11 / i12) * Im (i21 / i22) = Im (?i1 / ?den) * Im (?i2 / ?den)"
- using i `i12 = ?den` `i22 = ?den`
+ using i \<open>i12 = ?den\<close> \<open>i22 = ?den\<close>
by auto
also have "... = Im (?i1) * Im (?i2) / ?den\<^sup>2"
by simp
also have "... = (Im B * (Im B * (Re D * Re D)) - Re B * (Re B * ((cmod B)\<^sup>2 - (Re D)\<^sup>2))) / cmod B ^ 4"
- using `(cmod B)\<^sup>2 > (cmod A)\<^sup>2` `A = D`
- using `is_real D` cmod_eq_Re[of A]
+ using \<open>(cmod B)\<^sup>2 > (cmod A)\<^sup>2\<close> \<open>A = D\<close>
+ using \<open>is_real D\<close> cmod_eq_Re[of A]
by (auto simp add: field_simps)
also have "... = ((Im B)\<^sup>2 * (Re D)\<^sup>2 - (Re B)\<^sup>2 * ((Re B)\<^sup>2 + (Im B)\<^sup>2 - (Re D)\<^sup>2)) / cmod B ^ 4"
proof-
have "cmod B * cmod B = Re B * Re B + Im B * Im B"
by (metis cmod_power2 power2_eq_square)
thus ?thesis
by (simp add: power2_eq_square)
qed
also have "... = (((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2)) / cmod B ^ 4"
by (simp add: power2_eq_square field_simps)
finally have Im_product: "Im (i11 / i12) * Im (i21 / i22) = ((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) / cmod B ^ 4"
.
show "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0 \<longleftrightarrow> intersects_x_axis_cmat H"
proof safe
assume opposite: "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0"
show "intersects_x_axis_cmat H"
proof-
have "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) / cmod B ^ 4 < 0"
using Im_product opposite i1 i2
by simp
hence "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) < 0"
by (simp add: divide_less_0_iff)
hence "(Re D)\<^sup>2 < (Re B)\<^sup>2"
by (simp add: mult_less_0_iff not_sum_power2_lt_zero)
thus ?thesis
- using H `A = D` `is_real D`
+ using H \<open>A = D\<close> \<open>is_real D\<close>
by auto
qed
next
have *: "(\<forall>k. k * Im B = 1 \<longrightarrow> k = 0) \<longrightarrow> Im B = 0"
apply (safe, erule_tac x="1 / Im B" in allE)
using divide_cancel_left by fastforce
assume "intersects_x_axis_cmat H"
hence "Re D = 0 \<or> (Re D)\<^sup>2 < (Re B)\<^sup>2"
- using H `A = D`
+ using H \<open>A = D\<close>
by auto
hence "(Re D)\<^sup>2 < (Re B)\<^sup>2"
- using `is_real D` line H `C = cnj B`
+ using \<open>is_real D\<close> line H \<open>C = cnj B\<close>
using not_x_axis *
by (auto simp add: complex_eq_iff)
hence "((Re D)\<^sup>2 - (Re B)\<^sup>2) * ((Re B)\<^sup>2 + (Im B)\<^sup>2) < 0"
by (metis add_cancel_left_left diff_less_eq mult_eq_0_iff mult_less_0_iff power2_eq_square power2_less_0 sum_squares_gt_zero_iff)
thus "Im (to_complex_cvec i1) * Im (to_complex_cvec i2) < 0"
using Im_product i1 i2
using divide_eq_0_iff divide_less_0_iff prod.simps(2) to_complex_cvec_def zero_complex.simps(1) zero_less_norm_iff
by fastforce
qed
qed
thus ?thesis
- using assms `calc_ideal_points H = {i1, i2}` `i1 \<noteq> i2`
+ using assms \<open>calc_ideal_points H = {i1, i2}\<close> \<open>i1 \<noteq> i2\<close>
by auto
qed
end
diff --git a/thys/Poincare_Disc/Poincare_Lines_Ideal_Points.thy b/thys/Poincare_Disc/Poincare_Lines_Ideal_Points.thy
--- a/thys/Poincare_Disc/Poincare_Lines_Ideal_Points.thy
+++ b/thys/Poincare_Disc/Poincare_Lines_Ideal_Points.thy
@@ -1,683 +1,683 @@
theory Poincare_Lines_Ideal_Points
imports Poincare_Lines
begin
(* ------------------------------------------------------------------ *)
subsection\<open>Ideal points of h-lines\<close>
(* ------------------------------------------------------------------ *)
(* TODO: Introduce ideal points for the oriented circline -
it would be a list, not a set of two points *)
text\<open>\emph{Ideal points} of an h-line are points where the h-line intersects the unit disc.\<close>
(* ------------------------------------------------------------------ *)
subsubsection \<open>Calculation of ideal points\<close>
(* ------------------------------------------------------------------ *)
text \<open>We decided to define ideal points constructively, i.e., we calculate the coordinates of ideal
points for a given h-line explicitly. Namely, if the h-line is determined by $A$ and $B$, the two
intersection points are $$\frac{B}{|B|^2}\left(-A \pm i\cdot \sqrt{|B|^2 - A^2}\right).$$\<close>
definition calc_ideal_point1_cvec :: "complex \<Rightarrow> complex \<Rightarrow> complex_vec" where
[simp]: "calc_ideal_point1_cvec A B =
(let discr = Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2) in
(B*(-A - \<i>*sqrt(discr)), (cmod B)\<^sup>2))"
definition calc_ideal_point2_cvec :: "complex \<Rightarrow> complex \<Rightarrow> complex_vec" where
[simp]: "calc_ideal_point2_cvec A B =
(let discr = Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2) in
(B*(-A + \<i>*sqrt(discr)), (cmod B)\<^sup>2))"
definition calc_ideal_points_cmat_cvec :: "complex_mat \<Rightarrow> complex_vec set" where
[simp]: "calc_ideal_points_cmat_cvec H =
(if is_poincare_line_cmat H then
let (A, B, C, D) = H
in {calc_ideal_point1_cvec A B, calc_ideal_point2_cvec A B}
else
{(-1, 1), (1, 1)})"
lift_definition calc_ideal_points_clmat_hcoords :: "circline_mat \<Rightarrow> complex_homo_coords set" is calc_ideal_points_cmat_cvec
by (auto simp add: Let_def split: if_split_asm)
lift_definition calc_ideal_points :: "circline \<Rightarrow> complex_homo set" is calc_ideal_points_clmat_hcoords
proof transfer
fix H1 H2
assume hh: "hermitean H1 \<and> H1 \<noteq> mat_zero" "hermitean H2 \<and> H2 \<noteq> mat_zero"
obtain A1 B1 C1 D1 A2 B2 C2 D2 where *: "H1 = (A1, B1, C1, D1)" "H2 = (A2, B2, C2, D2)"
by (cases H1, cases H2, auto)
assume "circline_eq_cmat H1 H2"
then obtain k where k: "k \<noteq> 0" "H2 = cor k *\<^sub>s\<^sub>m H1"
by auto
thus "rel_set (\<approx>\<^sub>v) (calc_ideal_points_cmat_cvec H1) (calc_ideal_points_cmat_cvec H2)"
proof (cases "is_poincare_line_cmat H1")
case True
hence "is_poincare_line_cmat H2"
using k * hermitean_mult_real[of H1 k] hh
by (auto simp add: power2_eq_square)
have **: "sqrt (\<bar>k\<bar> * cmod B1 * (\<bar>k\<bar> * cmod B1) - k * Re D1 * (k * Re D1)) =
\<bar>k\<bar> * sqrt(cmod B1 * cmod B1 - Re D1 * Re D1)"
proof-
have "\<bar>k\<bar> * cmod B1 * (\<bar>k\<bar> * cmod B1) - k * Re D1 * (k * Re D1) =
k\<^sup>2 * (cmod B1 * cmod B1 - Re D1 * Re D1)"
by (simp add: power2_eq_square field_simps)
thus ?thesis
by (simp add: real_sqrt_mult)
qed
show ?thesis
using \<open>is_poincare_line_cmat H1\<close> \<open>is_poincare_line_cmat H2\<close>
using * k
apply (simp add: Let_def)
apply safe
apply (simp add: power2_eq_square rel_set_def)
apply safe
apply (cases "k > 0")
apply (rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (erule notE, rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (cases "k > 0")
apply (erule notE, rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (cases "k > 0")
apply (rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (erule notE, rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (cases "k > 0")
apply (erule notE, rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
apply (rule_tac x="(cor k)\<^sup>2" in exI)
apply (subst **)
apply (simp add: power2_eq_square field_simps)
done
next
case False
hence "\<not> is_poincare_line_cmat H2"
using k * hermitean_mult_real[of H1 k] hh
by (auto simp add: power2_eq_square)
have "rel_set (\<approx>\<^sub>v) {(- 1, 1), (1, 1)} {(- 1, 1), (1, 1)}"
by (simp add: rel_set_def)
thus ?thesis
using \<open>\<not> is_poincare_line_cmat H1\<close> \<open>\<not> is_poincare_line_cmat H2\<close>
using *
apply (simp add: Let_def)
apply safe
done
qed
qed
text \<open>Correctness of the calculation\<close>
text\<open>We show that for every h-line its two calculated ideal points are different and are on the
intersection of that line and the unit circle.\<close>
text \<open>Calculated ideal points are on the unit circle\<close>
lemma calc_ideal_point_1_unit:
assumes "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
assumes "(z1, z2) = calc_ideal_point1_cvec A B"
shows "z1 * cnj z1 = z2 * cnj z2"
proof-
let ?discr = "Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2)"
have "?discr > 0"
using assms
by (simp add: cmod_power2)
have "(B*(-A - \<i>*sqrt(?discr))) * cnj (B*(-A - \<i>*sqrt(?discr))) = (B * cnj B) * (A\<^sup>2 + cor (abs ?discr))"
using \<open>is_real A\<close> eq_cnj_iff_real[of A]
by (simp add: field_simps power2_eq_square)
also have "... = (B * cnj B) * (cmod B)\<^sup>2"
using \<open>?discr > 0\<close>
using assms
using complex_of_real_Re[of "(cmod B)\<^sup>2 - (Re A)\<^sup>2"] complex_of_real_Re[of A] \<open>is_real A\<close>
by (simp add: power2_eq_square)
also have "... = (cmod B)\<^sup>2 * cnj ((cmod B)\<^sup>2)"
using complex_cnj_complex_of_real complex_mult_cnj_cmod
by presburger
finally show ?thesis
using assms
by simp
qed
lemma calc_ideal_point_2_unit:
assumes "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
assumes "(z1, z2) = calc_ideal_point2_cvec A B"
shows "z1 * cnj z1 = z2 * cnj z2"
proof-
let ?discr = "Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2)"
have "?discr > 0"
using assms
by (simp add: cmod_power2)
have "(B*(-A + \<i>*sqrt(?discr))) * cnj (B*(-A + \<i>*sqrt(?discr))) = (B * cnj B) * (A\<^sup>2 + cor (abs ?discr))"
using \<open>is_real A\<close> eq_cnj_iff_real[of A]
by (simp add: field_simps power2_eq_square)
also have "... = (B * cnj B) * (cmod B)\<^sup>2"
using \<open>?discr > 0\<close>
using assms
using complex_of_real_Re[of "(cmod B)\<^sup>2 - (Re A)\<^sup>2"] complex_of_real_Re[of A] \<open>is_real A\<close>
by (simp add: power2_eq_square)
also have "... = (cmod B)\<^sup>2 * cnj ((cmod B)\<^sup>2)"
using complex_cnj_complex_of_real complex_mult_cnj_cmod
by presburger
finally show ?thesis
using assms
by simp
qed
lemma calc_ideal_points_on_unit_circle:
shows "\<forall> z \<in> calc_ideal_points H. z \<in> circline_set unit_circle"
unfolding circline_set_def
apply simp
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
have "\<forall> (z1, z2) \<in> calc_ideal_points_cmat_cvec H. z1 * cnj z1 = z2 * cnj z2"
using hermitean_elems[of A B C D]
unfolding calc_ideal_points_cmat_cvec_def
using calc_ideal_point_1_unit[of A B]
using calc_ideal_point_2_unit[of A B]
using hh *
apply (cases "calc_ideal_point1_cvec A B", cases "calc_ideal_point2_cvec A B")
apply (auto simp add: Let_def simp del: calc_ideal_point1_cvec_def calc_ideal_point2_cvec_def)
done
thus "Ball (calc_ideal_points_cmat_cvec H) (on_circline_cmat_cvec unit_circle_cmat)"
using on_circline_cmat_cvec_unit
by (auto simp del: on_circline_cmat_cvec_def calc_ideal_points_cmat_cvec_def)
qed
text \<open>Calculated ideal points are on the h-line\<close>
lemma calc_ideal_point1_sq:
assumes "(z1, z2) = calc_ideal_point1_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "z1 * cnj z1 + z2 * cnj z2 = 2 * (B * cnj B)\<^sup>2"
proof-
let ?discr = "Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2)"
have "?discr > 0"
using assms
by (simp add: cmod_power2)
have "z1 * cnj z1 = (B * cnj B) * (-A + \<i>*sqrt(?discr))*(-A - \<i>*sqrt(?discr))"
using assms eq_cnj_iff_real[of A]
by (simp)
also have "... = (B * cnj B) * (A\<^sup>2 + ?discr)"
using complex_of_real_Re[of A] \<open>is_real A\<close> \<open>?discr > 0\<close>
by (simp add: power2_eq_square field_simps)
finally
have "z1 * cnj z1 = (B * cnj B)\<^sup>2"
using complex_of_real_Re[of "(cmod B)\<^sup>2 - (Re A)\<^sup>2"] complex_of_real_Re[of A] \<open>is_real A\<close>
using complex_mult_cnj_cmod[of B]
by (simp add: power2_eq_square)
moreover
have "z2 * cnj z2 = (B * cnj B)\<^sup>2"
using assms
by simp
ultimately
show ?thesis
by simp
qed
lemma calc_ideal_point2_sq:
assumes "(z1, z2) = calc_ideal_point2_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "z1 * cnj z1 + z2 * cnj z2 = 2 * (B * cnj B)\<^sup>2"
proof-
let ?discr = "Re ((cmod B)\<^sup>2 - (Re A)\<^sup>2)"
have "?discr > 0"
using assms
by (simp add: cmod_power2)
have "z1 * cnj z1 = (B * cnj B) * (-A + \<i>*sqrt(?discr))*(-A - \<i>*sqrt(?discr))"
using assms eq_cnj_iff_real[of A]
by simp
also have "... = (B * cnj B) * (A\<^sup>2 + ?discr)"
using complex_of_real_Re[of A] \<open>is_real A\<close> \<open>?discr > 0\<close>
by (simp add: power2_eq_square field_simps)
finally
have "z1 * cnj z1 = (B * cnj B)\<^sup>2"
using complex_of_real_Re[of "(cmod B)\<^sup>2 - (Re A)\<^sup>2"] complex_of_real_Re[of A] \<open>is_real A\<close>
using complex_mult_cnj_cmod[of B]
by (simp add: power2_eq_square)
moreover
have "z2 * cnj z2 = (B * cnj B)\<^sup>2"
using assms
by simp
ultimately
show ?thesis
by simp
qed
lemma calc_ideal_point1_mix:
assumes "(z1, z2) = calc_ideal_point1_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "B * cnj z1 * z2 + cnj B * z1 * cnj z2 = - 2 * A * (B * cnj B)\<^sup>2 "
proof-
have "B*cnj z1 + cnj B*z1 = -2*A*B*cnj B"
using assms eq_cnj_iff_real[of A]
by (simp, simp add: field_simps)
moreover
have "cnj z2 = z2"
using assms
by simp
hence "B*cnj z1*z2 + cnj B*z1*cnj z2 = (B*cnj z1 + cnj B*z1)*z2"
by (simp add: field_simps)
ultimately
have "B*cnj z1*z2 + cnj B*z1*cnj z2 = -2*A*(B* cnj B)*z2"
by simp
also have "\<dots> = -2*A*(B * cnj B)\<^sup>2"
using assms
using complex_mult_cnj_cmod[of B]
by (simp add: power2_eq_square)
finally
show ?thesis
.
qed
lemma calc_ideal_point2_mix:
assumes "(z1, z2) = calc_ideal_point2_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "B * cnj z1 * z2 + cnj B * z1 * cnj z2 = - 2 * A * (B * cnj B)\<^sup>2 "
proof-
have "B*cnj z1 + cnj B*z1 = -2*A*B*cnj B"
using assms eq_cnj_iff_real[of A]
by (simp, simp add: field_simps)
moreover
have "cnj z2 = z2"
using assms
by simp
hence "B*cnj z1*z2 + cnj B*z1*cnj z2 = (B*cnj z1 + cnj B*z1)*z2"
by (simp add: field_simps)
ultimately
have "B*cnj z1*z2 + cnj B*z1*cnj z2 = -2*A*(B* cnj B)*z2"
by simp
also have "\<dots> = -2*A*(B * cnj B)\<^sup>2"
using assms
using complex_mult_cnj_cmod[of B]
by (simp add: power2_eq_square)
finally
show ?thesis
.
qed
lemma calc_ideal_point1_on_circline:
assumes "(z1, z2) = calc_ideal_point1_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "A*z1*cnj z1 + B*cnj z1*z2 + cnj B*z1*cnj z2 + A*z2*cnj z2 = 0" (is "?lhs = 0")
proof-
have "?lhs = A * (z1 * cnj z1 + z2 * cnj z2) + (B * cnj z1 * z2 + cnj B * z1 * cnj z2)"
by (simp add: field_simps)
also have "... = 2*A*(B*cnj B)\<^sup>2 + (-2*A*(B*cnj B)\<^sup>2)"
using calc_ideal_point1_sq[OF assms]
using calc_ideal_point1_mix[OF assms]
by simp
finally
show ?thesis
by simp
qed
lemma calc_ideal_point2_on_circline:
assumes "(z1, z2) = calc_ideal_point2_cvec A B" "is_real A" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
shows "A*z1*cnj z1 + B*cnj z1*z2 + cnj B*z1*cnj z2 + A*z2*cnj z2 = 0" (is "?lhs = 0")
proof-
have "?lhs = A * (z1 * cnj z1 + z2 * cnj z2) + (B * cnj z1 * z2 + cnj B * z1 * cnj z2)"
by (simp add: field_simps)
also have "... = 2*A*(B*cnj B)\<^sup>2 + (-2*A*(B*cnj B)\<^sup>2)"
using calc_ideal_point2_sq[OF assms]
using calc_ideal_point2_mix[OF assms]
by simp
finally
show ?thesis
by simp
qed
lemma calc_ideal_points_on_circline:
assumes "is_poincare_line H"
shows "\<forall> z \<in> calc_ideal_points H. z \<in> circline_set H"
using assms
unfolding circline_set_def
apply simp
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
obtain z11 z12 z21 z22 where **: "(z11, z12) = calc_ideal_point1_cvec A B" "(z21, z22) = calc_ideal_point2_cvec A B"
by (cases "calc_ideal_point1_cvec A B", cases "calc_ideal_point2_cvec A B") auto
assume "is_poincare_line_cmat H"
hence "\<forall> (z1, z2) \<in> calc_ideal_points_cmat_cvec H. A*z1*cnj z1 + B*cnj z1*z2 + C*z1*cnj z2 + D*z2*cnj z2 = 0"
using * ** hh
using hermitean_elems[of A B C D]
using calc_ideal_point1_on_circline[of z11 z12 A B]
using calc_ideal_point2_on_circline[of z21 z22 A B]
by (auto simp del: calc_ideal_point1_cvec_def calc_ideal_point2_cvec_def)
thus "Ball (calc_ideal_points_cmat_cvec H) (on_circline_cmat_cvec H)"
using on_circline_cmat_cvec_circline_equation *
by (auto simp del: on_circline_cmat_cvec_def calc_ideal_points_cmat_cvec_def simp add: field_simps)
qed
text \<open>Calculated ideal points of an h-line are different\<close>
lemma calc_ideal_points_cvec_different [simp]:
assumes "(cmod B)\<^sup>2 > (cmod A)\<^sup>2" "is_real A"
shows "\<not> (calc_ideal_point1_cvec A B \<approx>\<^sub>v calc_ideal_point2_cvec A B)"
using assms
by (auto) (auto simp add: cmod_def)
lemma calc_ideal_points_different:
assumes "is_poincare_line H"
shows "\<exists> i1 \<in> (calc_ideal_points H). \<exists> i2 \<in> (calc_ideal_points H). i1 \<noteq> i2"
using assms
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero" "is_poincare_line_cmat H"
obtain A B C D where *: "H = (A, B, C, D)"
by (cases H, auto)
hence "is_real A" using hh hermitean_elems by auto
thus "\<exists>i1\<in>calc_ideal_points_cmat_cvec H. \<exists>i2\<in>calc_ideal_points_cmat_cvec H. \<not> i1 \<approx>\<^sub>v i2"
using * hh calc_ideal_points_cvec_different[of A B]
apply (rule_tac x="calc_ideal_point1_cvec A B" in bexI)
apply (rule_tac x="calc_ideal_point2_cvec A B" in bexI)
by auto
qed
lemma two_calc_ideal_points [simp]:
assumes "is_poincare_line H"
shows "card (calc_ideal_points H) = 2"
proof-
have "\<exists> x \<in> calc_ideal_points H. \<exists> y \<in> calc_ideal_points H. \<forall> z \<in> calc_ideal_points H. z = x \<or> z = y"
by (transfer, transfer, case_tac H, simp del: calc_ideal_point1_cvec_def calc_ideal_point2_cvec_def)
then obtain x y where *: "calc_ideal_points H = {x, y}"
by auto
moreover
have "x \<noteq> y"
using calc_ideal_points_different[OF assms] *
by auto
ultimately
show ?thesis
by auto
qed
subsubsection \<open>Ideal points\<close>
text \<open>Next we give a genuine definition of ideal points -- these are the intersections of the h-line with the unit circle\<close>
definition ideal_points :: "circline \<Rightarrow> complex_homo set" where
"ideal_points H = circline_intersection H unit_circle"
text \<open>Ideal points are on the unit circle and on the h-line\<close>
lemma ideal_points_on_unit_circle:
shows "\<forall> z \<in> ideal_points H. z \<in> circline_set unit_circle"
unfolding ideal_points_def circline_intersection_def circline_set_def
by simp
lemma ideal_points_on_circline:
shows "\<forall> z \<in> ideal_points H. z \<in> circline_set H"
unfolding ideal_points_def circline_intersection_def circline_set_def
by simp
text \<open>For each h-line there are exactly two ideal points\<close>
lemma two_ideal_points:
assumes "is_poincare_line H"
shows "card (ideal_points H) = 2"
proof-
have "H \<noteq> unit_circle"
using assms not_is_poincare_line_unit_circle
by auto
let ?int = "circline_intersection H unit_circle"
obtain i1 i2 where "i1 \<in> ?int" "i2 \<in> ?int" "i1 \<noteq> i2"
using calc_ideal_points_on_circline[OF assms]
using calc_ideal_points_on_unit_circle[of H]
using calc_ideal_points_different[OF assms]
unfolding circline_intersection_def circline_set_def
by auto
thus ?thesis
unfolding ideal_points_def
using circline_intersection_at_most_2_points[OF \<open>H \<noteq> unit_circle\<close>]
using card_geq_2_iff_contains_2_elems[of ?int]
by auto
qed
text \<open>They are exactly the two points that our calculation finds\<close>
lemma ideal_points_unique:
assumes "is_poincare_line H"
shows "ideal_points H = calc_ideal_points H"
proof-
have "calc_ideal_points H \<subseteq> ideal_points H"
using calc_ideal_points_on_circline[OF assms]
using calc_ideal_points_on_unit_circle[of H]
unfolding ideal_points_def circline_intersection_def circline_set_def
by auto
moreover
have "H \<noteq> unit_circle"
using not_is_poincare_line_unit_circle assms
by auto
hence "finite (ideal_points H)"
using circline_intersection_at_most_2_points[of H unit_circle]
unfolding ideal_points_def
by auto
ultimately
show ?thesis
using card_subset_eq[of "ideal_points H" "calc_ideal_points H"]
using two_calc_ideal_points[OF assms]
using two_ideal_points[OF assms]
by auto
qed
text \<open>For each h-line we can obtain two different ideal points\<close>
lemma obtain_ideal_points:
assumes "is_poincare_line H"
obtains i1 i2 where "i1 \<noteq> i2" "ideal_points H = {i1, i2}"
using two_ideal_points[OF assms] card_eq_2_iff_doubleton[of "ideal_points H"]
by blast
text \<open>Ideal points of each h-line constructed from two points in the disc are different than those two points\<close>
lemma ideal_points_different:
assumes "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes "ideal_points (poincare_line u v) = {i1, i2}"
shows "i1 \<noteq> i2" "u \<noteq> i1" "u \<noteq> i2" "v \<noteq> i1" "v \<noteq> i2"
proof-
have "i1 \<in> ocircline_set ounit_circle" "i2 \<in> ocircline_set ounit_circle"
using assms(3) assms(4) ideal_points_on_unit_circle is_poincare_line_poincare_line
by fastforce+
thus "u \<noteq> i1" "u \<noteq> i2" "v \<noteq> i1" "v \<noteq> i2"
using assms(1-2)
using disc_inter_ocircline_set[of ounit_circle]
unfolding unit_disc_def
by auto
show "i1 \<noteq> i2"
using assms
by (metis doubleton_eq_iff is_poincare_line_poincare_line obtain_ideal_points)
qed
text \<open>H-line is uniquely determined by its ideal points\<close>
lemma ideal_points_line_unique:
assumes "is_poincare_line H" "ideal_points H = {i1, i2}"
shows "H = poincare_line i1 i2"
by (smt assms(1) assms(2) calc_ideal_points_on_unit_circle circline_set_def ex_poincare_line_points ideal_points_different(1) ideal_points_on_circline ideal_points_unique insertI1 insert_commute inversion_unit_circle mem_Collect_eq unique_poincare_line_general)
text \<open>Ideal points of some special h-lines\<close>
text\<open>Ideal points of @{term x_axis}\<close>
lemma ideal_points_x_axis
[simp]: "ideal_points x_axis = {of_complex (-1), of_complex 1}"
proof (subst ideal_points_unique, simp)
have "calc_ideal_points_clmat_hcoords x_axis_clmat = {of_complex_hcoords (- 1), of_complex_hcoords 1}"
by transfer auto
thus "calc_ideal_points x_axis = {of_complex (- 1), of_complex 1}"
by (simp add: calc_ideal_points.abs_eq of_complex.abs_eq x_axis_def)
qed
text \<open>Ideal points are proportional vectors only if h-line is a line segment passing trough zero\<close>
lemma ideal_points_proportional:
assumes "is_poincare_line H" "ideal_points H = {i1, i2}" "to_complex i1 = cor k * to_complex i2"
shows "0\<^sub>h \<in> circline_set H"
proof-
have "i1 \<noteq> i2"
- using `ideal_points H = {i1, i2}`
- using `is_poincare_line H` ex_poincare_line_points ideal_points_different(1) by blast
+ using \<open>ideal_points H = {i1, i2}\<close>
+ using \<open>is_poincare_line H\<close> ex_poincare_line_points ideal_points_different(1) by blast
have "i1 \<in> circline_set unit_circle" "i2 \<in> circline_set unit_circle"
using assms calc_ideal_points_on_unit_circle ideal_points_unique
by blast+
hence "cmod (cor k) = 1"
- using `to_complex i1 = cor k * to_complex i2`
+ using \<open>to_complex i1 = cor k * to_complex i2\<close>
by (metis (mono_tags, lifting) circline_set_unit_circle imageE mem_Collect_eq mult.right_neutral norm_mult to_complex_of_complex unit_circle_set_def)
hence "k = -1"
- using `to_complex i1 = cor k * to_complex i2` `i1 \<noteq> i2`
+ using \<open>to_complex i1 = cor k * to_complex i2\<close> \<open>i1 \<noteq> i2\<close>
using \<open>i1 \<in> circline_set unit_circle\<close> \<open>i2 \<in> circline_set unit_circle\<close>
by (metis (no_types, lifting) circline_set_unit_circle complex_cnj_complex_of_real complex_mult_cnj_cmod cor_neg_one imageE mult_cancel_right2 norm_one of_real_eq_iff square_eq_1_iff to_complex_of_complex)
have "\<forall> i1 \<in> calc_ideal_points H. \<forall> i2 \<in> calc_ideal_points H. is_poincare_line H \<and> i1 \<noteq> i2 \<and> to_complex i1 = - to_complex i2 \<longrightarrow>
0\<^sub>h \<in> circline_set H"
unfolding circline_set_def
proof (simp, transfer, transfer, safe)
fix A B C D i11 i12 i21 i22 k
assume H:"hermitean (A, B, C, D)" "(A, B, C, D) \<noteq> mat_zero"
assume line: "is_poincare_line_cmat (A, B, C, D)"
assume i1: "(i11, i12) \<in> calc_ideal_points_cmat_cvec (A, B, C, D)"
assume i2:"(i21, i22) \<in> calc_ideal_points_cmat_cvec (A, B, C, D)"
assume "\<not> (i11, i12) \<approx>\<^sub>v (i21, i22)"
assume opposite: "to_complex_cvec (i11, i12) = - to_complex_cvec (i21, i22)"
let ?discr = "sqrt ((cmod B)\<^sup>2 - (Re D)\<^sup>2)"
let ?den = "(cmod B)\<^sup>2"
let ?i1 = "B * (- D - \<i> * ?discr)"
let ?i2 = "B * (- D + \<i> * ?discr)"
have "i11 = ?i1 \<or> i11 = ?i2" "i12 = ?den"
"i21 = ?i1 \<or> i21 = ?i2" "i22 = ?den"
using i1 i2 H line
by (auto split: if_split_asm)
hence i: "i11 = ?i1 \<and> i21 = ?i2 \<or> i11 = ?i2 \<and> i21 = ?i1"
- using `\<not> (i11, i12) \<approx>\<^sub>v (i21, i22)`
+ using \<open>\<not> (i11, i12) \<approx>\<^sub>v (i21, i22)\<close>
by auto
have "?den \<noteq> 0"
using line
by auto
hence "i11 = - i21"
- using opposite `i12 = ?den` `i22 = ?den`
+ using opposite \<open>i12 = ?den\<close> \<open>i22 = ?den\<close>
by (simp add: nonzero_neg_divide_eq_eq2)
hence "?i1 = - ?i2"
using i
by (metis add.inverse_inverse)
hence "D = 0"
- using `?den \<noteq> 0`
+ using \<open>?den \<noteq> 0\<close>
by (simp add: field_simps)
thus "on_circline_cmat_cvec (A, B, C, D) 0\<^sub>v"
by (simp add: vec_cnj_def)
qed
thus ?thesis
- using assms `k = -1`
+ using assms \<open>k = -1\<close>
using calc_ideal_points_different ideal_points_unique
by fastforce
qed
text \<open>Transformations of ideal points\<close>
text \<open>Möbius transformations that fix the unit disc when acting on h-lines map their ideal points to ideal points.\<close>
lemma ideal_points_moebius_circline [simp]:
assumes "unit_circle_fix M" "is_poincare_line H"
shows "ideal_points (moebius_circline M H) = (moebius_pt M) ` (ideal_points H)" (is "?I' = ?M ` ?I")
proof-
obtain i1 i2 where *: "i1 \<noteq> i2" "?I = {i1, i2}"
using assms(2)
by (rule obtain_ideal_points)
let ?Mi1 = "?M i1" and ?Mi2 = "?M i2"
have "?Mi1 \<in> ?M ` (circline_set H)"
"?Mi2 \<in> ?M ` (circline_set H)"
"?Mi1 \<in> ?M ` (circline_set unit_circle)"
"?Mi2 \<in> ?M ` (circline_set unit_circle)"
using *
unfolding ideal_points_def circline_intersection_def circline_set_def
by blast+
hence "?Mi1 \<in> ?I'"
"?Mi2 \<in> ?I'"
using unit_circle_fix_iff[of M] assms
unfolding ideal_points_def circline_intersection_def circline_set_def
by (metis mem_Collect_eq moebius_circline)+
moreover
have "?Mi1 \<noteq> ?Mi2"
using bij_moebius_pt[of M] *
using moebius_pt_invert by blast
moreover
have "is_poincare_line (moebius_circline M H)"
using assms unit_circle_fix_preserve_is_poincare_line
by simp
ultimately
have "?I' = {?Mi1, ?Mi2}"
using two_ideal_points[of "moebius_circline M H"]
using card_eq_2_doubleton[of ?I' ?Mi1 ?Mi2]
by simp
thus ?thesis
using *(2)
by auto
qed
lemma ideal_points_poincare_line_moebius [simp]:
assumes "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes "ideal_points (poincare_line u v) = {i1, i2}"
shows "ideal_points (poincare_line (moebius_pt M u) (moebius_pt M v)) = {moebius_pt M i1, moebius_pt M i2}"
using assms
by auto
text \<open>Conjugation also maps ideal points to ideal points\<close>
lemma ideal_points_conjugate [simp]:
assumes "is_poincare_line H"
shows "ideal_points (conjugate_circline H) = conjugate ` (ideal_points H)" (is "?I' = ?M ` ?I")
proof-
obtain i1 i2 where *: "i1 \<noteq> i2" "?I = {i1, i2}"
using assms
by (rule obtain_ideal_points)
let ?Mi1 = "?M i1" and ?Mi2 = "?M i2"
have "?Mi1 \<in> ?M ` (circline_set H)"
"?Mi2 \<in> ?M ` (circline_set H)"
"?Mi1 \<in> ?M ` (circline_set unit_circle)"
"?Mi2 \<in> ?M ` (circline_set unit_circle)"
using *
unfolding ideal_points_def circline_intersection_def circline_set_def
by blast+
hence "?Mi1 \<in> ?I'"
"?Mi2 \<in> ?I'"
unfolding ideal_points_def circline_intersection_def circline_set_def
using circline_set_conjugate_circline circline_set_def conjugate_unit_circle_set
by blast+
moreover
have "?Mi1 \<noteq> ?Mi2"
using \<open>i1 \<noteq> i2\<close>
by (auto simp add: conjugate_inj)
moreover
have "is_poincare_line (conjugate_circline H)"
using assms
by simp
ultimately
have "?I' = {?Mi1, ?Mi2}"
using two_ideal_points[of "conjugate_circline H"]
using card_eq_2_doubleton[of ?I' ?Mi1 ?Mi2]
by simp
thus ?thesis
using *(2)
by auto
qed
lemma ideal_points_poincare_line_conjugate [simp]:
assumes"u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes "ideal_points (poincare_line u v) = {i1, i2}"
shows "ideal_points (poincare_line (conjugate u) (conjugate v)) = {conjugate i1, conjugate i2}"
using assms
by auto
end
diff --git a/thys/Poincare_Disc/Poincare_Tarski.thy b/thys/Poincare_Disc/Poincare_Tarski.thy
--- a/thys/Poincare_Disc/Poincare_Tarski.thy
+++ b/thys/Poincare_Disc/Poincare_Tarski.thy
@@ -1,3045 +1,3045 @@
section \<open>Poincar\'e model satisfies Tarski axioms\<close>
theory Poincare_Tarski
imports Poincare Poincare_Lines_Axis_Intersections Tarski
begin
(* ------------------------------------------------------------------ *)
subsection\<open>Pasch axiom\<close>
(* ------------------------------------------------------------------ *)
lemma Pasch_fun_mono:
fixes r1 r2 :: real
assumes "0 < r1" and "r1 \<le> r2" and "r2 < 1"
shows "r1 + 1/r1 \<ge> r2 + 1/r2"
proof (cases "r1 = r2")
case True
thus ?thesis
by simp
next
case False
hence "r2 - r1 > 0"
using assms
by simp
have "r1 * r2 < 1"
using assms
by (smt mult_le_cancel_left1)
hence "1 / (r1 * r2) > 1"
using assms
by simp
hence "(r2 - r1) / (r1 * r2) > (r2 - r1)"
using \<open>r2 - r1 > 0\<close>
using mult_less_cancel_left_pos[of "r2 - r1" 1 "1 / (r1 * r2)"]
by simp
hence "1 / r1 - 1 / r2 > r2 - r1"
using assms
by (simp add: field_simps)
thus ?thesis
by simp
qed
text\<open>Pasch axiom, non-degenerative case.\<close>
lemma Pasch_nondeg:
assumes "x \<in> unit_disc" and "y \<in> unit_disc" and "z \<in> unit_disc" and "u \<in> unit_disc" and "v \<in> unit_disc"
assumes "distinct [x, y, z, u, v]"
assumes "\<not> poincare_collinear {x, y, z}"
assumes "poincare_between x u z" and "poincare_between y v z"
shows "\<exists> a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between x a v"
proof-
have "\<forall> y z u. distinct [x, y, z, u, v] \<and> \<not> poincare_collinear {x, y, z} \<and> y \<in> unit_disc \<and> z \<in> unit_disc \<and> u \<in> unit_disc \<and>
poincare_between x u z \<and> poincare_between y v z \<longrightarrow> (\<exists> a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between x a v)" (is "?P x v")
proof (rule wlog_positive_x_axis[where P="?P"])
fix v
assume v: "is_real v" "0 < Re v" "Re v < 1"
hence "of_complex v \<in> unit_disc"
by (auto simp add: cmod_eq_Re)
show "?P 0\<^sub>h (of_complex v)"
proof safe
fix y z u
assume distinct: "distinct [0\<^sub>h, y, z, u, of_complex v]"
assume in_disc: "y \<in> unit_disc" "z \<in> unit_disc" "u \<in> unit_disc"
then obtain y' z' u'
where *: "y = of_complex y'" "z = of_complex z'" "u = of_complex u'"
using inf_or_of_complex inf_notin_unit_disc
by metis
have "y' \<noteq> 0" "z' \<noteq> 0" "u' \<noteq> 0" "v \<noteq> 0" "y' \<noteq> z'" "y' \<noteq> u'" "z' \<noteq> u'" "y \<noteq> z" "y \<noteq> u" "z \<noteq> u"
using of_complex_inj distinct *
by auto
note distinct = distinct this
assume "\<not> poincare_collinear {0\<^sub>h, y, z}"
hence nondeg_yz: "y'*cnj z' \<noteq> cnj y' * z'"
using * poincare_collinear_zero_iff[of y' z'] in_disc distinct
by auto
assume "poincare_between 0\<^sub>h u z"
hence "arg u' = arg z'" "cmod u' \<le> cmod z'"
using * poincare_between_0uv[of u z] distinct in_disc
by auto
then obtain \<phi> ru rz where
uz_polar: "u' = cor ru * cis \<phi>" "z' = cor rz * cis \<phi>" "0 < ru" "ru \<le> rz" "0 < rz" and
"\<phi> = arg u'" "\<phi> = arg z'"
using * \<open>u' \<noteq> 0\<close> \<open>z' \<noteq> 0\<close>
by (smt cmod_cis norm_le_zero_iff)
obtain \<theta> ry where
y_polar: "y' = cor ry * cis \<theta>" "ry > 0" and "\<theta> = arg y'"
using \<open>y' \<noteq> 0\<close>
by (smt cmod_cis norm_le_zero_iff)
from in_disc * \<open>u' = cor ru * cis \<phi>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>y' = cor ry * cis \<theta>\<close>
have "ru < 1" "rz < 1" "ry < 1"
by simp_all
note polar = this y_polar uz_polar
have nondeg: "cis \<theta> * cis (- \<phi>) \<noteq> cis (- \<theta>) * cis \<phi>"
using nondeg_yz polar
by simp
let ?yz = "poincare_line y z"
let ?v = "calc_x_axis_intersection ?yz"
assume "poincare_between y (of_complex v) z"
hence "of_complex v \<in> circline_set ?yz"
using in_disc \<open>of_complex v \<in> unit_disc\<close>
using distinct poincare_between_poincare_collinear[of y "of_complex v" z]
using unique_poincare_line[of y z]
by (auto simp add: poincare_collinear_def)
moreover
have "of_complex v \<in> circline_set x_axis"
using \<open>is_real v\<close>
unfolding circline_set_x_axis
by auto
moreover
have "?yz \<noteq> x_axis"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "{0\<^sub>h, y, z} \<subseteq> circline_set (poincare_line y z)"
unfolding circline_set_def
using distinct poincare_line[of y z]
by auto
hence "poincare_collinear {0\<^sub>h, y, z}"
unfolding poincare_collinear_def
using distinct
by force
thus False
using \<open>\<not> poincare_collinear {0\<^sub>h, y, z}\<close>
by simp
qed
ultimately
have "?v = of_complex v" "intersects_x_axis ?yz"
using unique_calc_x_axis_intersection[of "poincare_line y z" "of_complex v"]
using intersects_x_axis_iff[of ?yz]
using distinct \<open>of_complex v \<in> unit_disc\<close>
by (metis IntI is_poincare_line_poincare_line)+
have "intersects_x_axis_positive ?yz"
using \<open>Re v > 0\<close> \<open>of_complex v \<in> unit_disc\<close>
using \<open>of_complex v \<in> circline_set ?yz\<close> \<open>of_complex v \<in> circline_set x_axis\<close>
using intersects_x_axis_positive_iff[of ?yz] \<open>y \<noteq> z\<close> \<open>?yz \<noteq> x_axis\<close>
unfolding positive_x_axis_def
by force
have "y \<notin> circline_set x_axis"
proof (rule ccontr)
assume "\<not> ?thesis"
moreover
hence "poincare_line y (of_complex v) = x_axis"
using distinct \<open>of_complex v \<in> circline_set x_axis\<close>
using in_disc \<open>of_complex v \<in> unit_disc\<close>
using unique_poincare_line[of y "of_complex v" x_axis]
by simp
moreover
have "z \<in> circline_set (poincare_line y (of_complex v))"
using \<open>of_complex v \<in> circline_set ?yz\<close>
using unique_poincare_line[of y "of_complex v" "poincare_line y z"]
using in_disc \<open>of_complex v \<in> unit_disc\<close> distinct
using poincare_line[of y z]
unfolding circline_set_def
by (metis distinct_length_2_or_more is_poincare_line_poincare_line mem_Collect_eq)
ultimately
have "y \<in> circline_set x_axis" "z \<in> circline_set x_axis"
by auto
hence "poincare_collinear {0\<^sub>h, y, z}"
unfolding poincare_collinear_def
by force
thus False
using \<open>\<not> poincare_collinear {0\<^sub>h, y, z}\<close>
by simp
qed
moreover
have "z \<notin> circline_set x_axis"
proof (rule ccontr)
assume "\<not> ?thesis"
moreover
hence "poincare_line z (of_complex v) = x_axis"
using distinct \<open>of_complex v \<in> circline_set x_axis\<close>
using in_disc \<open>of_complex v \<in> unit_disc\<close>
using unique_poincare_line[of z "of_complex v" x_axis]
by simp
moreover
have "y \<in> circline_set (poincare_line z (of_complex v))"
using \<open>of_complex v \<in> circline_set ?yz\<close>
using unique_poincare_line[of z "of_complex v" "poincare_line y z"]
using in_disc \<open>of_complex v \<in> unit_disc\<close> distinct
using poincare_line[of y z]
unfolding circline_set_def
by (metis distinct_length_2_or_more is_poincare_line_poincare_line mem_Collect_eq)
ultimately
have "y \<in> circline_set x_axis" "z \<in> circline_set x_axis"
by auto
hence "poincare_collinear {0\<^sub>h, y, z}"
unfolding poincare_collinear_def
by force
thus False
using \<open>\<not> poincare_collinear {0\<^sub>h, y, z}\<close>
by simp
qed
ultimately
have "\<phi> * \<theta> < 0"
using \<open>poincare_between y (of_complex v) z\<close>
using poincare_between_x_axis_intersection[of y z "of_complex v"]
using in_disc \<open>of_complex v \<in> unit_disc\<close> distinct
using \<open>of_complex v \<in> circline_set ?yz\<close> \<open>of_complex v \<in> circline_set x_axis\<close>
using \<open>\<phi> = arg z'\<close> \<open>\<theta> = arg y'\<close> *
by (simp add: field_simps)
have "\<phi> \<noteq> pi" "\<phi> \<noteq> 0"
using \<open>z \<notin> circline_set x_axis\<close> * polar cis_pi
unfolding circline_set_x_axis
by auto
have "\<theta> \<noteq> pi" "\<theta> \<noteq> 0"
using \<open>y \<notin> circline_set x_axis\<close> * polar cis_pi
unfolding circline_set_x_axis
by auto
have phi_sin: "\<phi> > 0 \<longleftrightarrow> sin \<phi> > 0" "\<phi> < 0 \<longleftrightarrow> sin \<phi> < 0"
using \<open>\<phi> = arg z'\<close> \<open>\<phi> \<noteq> 0\<close> \<open>\<phi> \<noteq> pi\<close>
using arg_bounded[of z']
by (smt sin_gt_zero sin_le_zero sin_pi_minus sin_0_iff_canon sin_ge_zero)+
have theta_sin: "\<theta> > 0 \<longleftrightarrow> sin \<theta> > 0" "\<theta> < 0 \<longleftrightarrow> sin \<theta> < 0"
using \<open>\<theta> = arg y'\<close> \<open>\<theta> \<noteq> 0\<close> \<open>\<theta> \<noteq> pi\<close>
using arg_bounded[of y']
by (smt sin_gt_zero sin_le_zero sin_pi_minus sin_0_iff_canon sin_ge_zero)+
have "sin \<phi> * sin \<theta> < 0"
using \<open>\<phi> * \<theta> < 0\<close> phi_sin theta_sin
by (simp add: mult_less_0_iff)
have "sin (\<phi> - \<theta>) \<noteq> 0"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "sin (\<phi> - \<theta>) = 0"
by simp
have "- 2 * pi < \<phi> - \<theta>" "\<phi> - \<theta> < 2 * pi"
using \<open>\<phi> = arg z'\<close> \<open>\<theta> = arg y'\<close> arg_bounded[of z'] arg_bounded[of y'] \<open>\<phi> \<noteq> pi\<close> \<open>\<theta> \<noteq> pi\<close>
by auto
hence "\<phi> - \<theta> = -pi \<or> \<phi> - \<theta> = 0 \<or> \<phi> - \<theta> = pi"
using \<open>sin (\<phi> - \<theta>) = 0\<close>
by (smt sin_0_iff_canon sin_periodic_pi2)
moreover
{
assume "\<phi> - \<theta> = - pi"
hence "\<phi> = \<theta> - pi"
by simp
hence False
using nondeg_yz
using \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>rz > 0\<close> \<open>ry > 0\<close>
by auto
}
moreover
{
assume "\<phi> - \<theta> = 0"
hence "\<phi> = \<theta>"
by simp
hence False
using \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>rz > 0\<close> \<open>ry > 0\<close>
using nondeg_yz
by auto
}
moreover
{
assume "\<phi> - \<theta> = pi"
hence "\<phi> = \<theta> + pi"
by simp
hence False
using \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>rz > 0\<close> \<open>ry > 0\<close>
using nondeg_yz
by auto
}
ultimately
show False
by auto
qed
have "u \<notin> circline_set x_axis"
proof-
have "\<not> is_real u'"
using * polar in_disc
using \<open>\<phi> \<noteq> 0\<close> \<open>\<phi> = arg u'\<close> \<open>\<phi> \<noteq> pi\<close> phi_sin(1) phi_sin(2)
by (metis is_real_arg2)
moreover
have "u \<noteq> \<infinity>\<^sub>h"
using in_disc
by auto
ultimately
show ?thesis
using * of_complex_inj[of u']
unfolding circline_set_x_axis
by auto
qed
let ?yu = "poincare_line y u"
have nondeg_yu: "y' * cnj u' \<noteq> cnj u' * u'"
using nondeg_yz polar \<open>ru > 0\<close> \<open>rz > 0\<close> distinct
by auto
{
(* derive results simultaneously for both u and z *)
fix r :: real
assume "r > 0"
have den: "cor ry * cis \<theta> * cnj 1 * cnj (cor r * cis \<phi>) * 1 - cor r * cis \<phi> * cnj 1 * cnj (cor ry * cis \<theta>) * 1 \<noteq> 0"
using \<open>0 < r\<close> \<open>0 < ry\<close> nondeg
by auto
let ?A = "2 * r * ry * sin(\<phi> - \<theta>)"
let ?B = "\<i> * (r * cis \<phi> * (1 + ry\<^sup>2) - ry * cis \<theta> * (1 + r\<^sup>2))"
let ?ReB = "ry * (1 + r\<^sup>2) * sin \<theta> - r * (1 + ry\<^sup>2) * sin \<phi>"
have "Re (\<i> * (r * cis (-\<phi>) * ry * cis (\<theta>) - ry * cis (-\<theta>) * r * cis (\<phi>))) = ?A"
by (simp add: sin_diff field_simps)
moreover
have "cor ry * cis (- \<theta>) * (cor ry * cis \<theta>) = ry\<^sup>2" "cor r * cis (- \<phi>) * (cor r * cis \<phi>) = r\<^sup>2"
by (metis cis_inverse cis_neq_zero divide_complex_def cor_squared nonzero_mult_div_cancel_right power2_eq_square semiring_normalization_rules(15))+
ultimately
have 1: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor r * cis \<phi>)) = (?A, ?B, cnj ?B, ?A)"
using den
unfolding poincare_line_cvec_cmat_def of_complex_cvec_def Let_def prod.case
by (simp add: field_simps)
have 2: "is_real ?A"
by simp
let ?mix = "cis \<theta> * cis (- \<phi>) - cis (- \<theta>) * cis \<phi>"
have "is_imag ?mix"
using eq_minus_cnj_iff_imag[of ?mix]
by simp
hence "Im ?mix \<noteq> 0"
using nondeg
using complex.expand[of ?mix 0]
by auto
hence 3: "Re ?A \<noteq> 0"
using \<open>r > 0\<close> \<open>ry > 0\<close>
by (simp add: sin_diff field_simps)
have "?A \<noteq> 0"
using 2 3
by auto
hence 4: "cor ?A \<noteq> 0"
using 2 3
by (metis zero_complex.simps(1))
have 5: "?ReB / ?A = (sin \<theta>) / (2 * sin(\<phi> - \<theta>)) * (1/r + r) - (sin \<phi>) / (2 * sin (\<phi> - \<theta>)) * (1/ry + ry)"
using \<open>ry > 0\<close> \<open>r > 0\<close>
apply (subst diff_divide_distrib)
apply (subst add_frac_num, simp)
apply (subst add_frac_num, simp)
apply (simp add: power2_eq_square mult.commute)
apply (simp add: field_simps)
done
have "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor r * cis \<phi>)) = (?A, ?B, cnj ?B, ?A) \<and>
is_real ?A \<and> Re ?A \<noteq> 0 \<and> ?A \<noteq> 0 \<and> cor ?A \<noteq> 0 \<and>
Re ?B = ?ReB \<and>
?ReB / ?A = (sin \<theta>) / (2 * sin(\<phi> - \<theta>)) * (1/r + r) - (sin \<phi>) / (2 * sin (\<phi> - \<theta>)) * (1/ry + ry)"
using 1 2 3 4 5
by auto
}
note ** = this
let ?Ayz = "2 * rz * ry * sin (\<phi> - \<theta>)"
let ?Byz = "\<i> * (rz * cis \<phi> * (1 + ry\<^sup>2) - ry * cis \<theta> * (1 + rz\<^sup>2))"
let ?ReByz = "ry * (1 + rz\<^sup>2) * sin \<theta> - rz * (1 + ry\<^sup>2) * sin \<phi>"
let ?Kz = "(sin \<theta>) / (2 * sin(\<phi> - \<theta>)) * (1/rz + rz) - (sin \<phi>) / (2 * sin (\<phi> - \<theta>)) * (1/ry + ry)"
have yz: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor rz * cis \<phi>)) = (?Ayz, ?Byz, cnj ?Byz, ?Ayz)"
"is_real ?Ayz" "Re ?Ayz \<noteq> 0" "?Ayz \<noteq> 0" "cor ?Ayz \<noteq> 0" "Re ?Byz = ?ReByz" and Kz: "?ReByz / ?Ayz = ?Kz"
using **[OF \<open>0 < rz\<close>]
by auto
let ?Ayu = "2 * ru * ry * sin (\<phi> - \<theta>)"
let ?Byu = "\<i> * (ru * cis \<phi> * (1 + ry\<^sup>2) - ry * cis \<theta> * (1 + ru\<^sup>2))"
let ?ReByu = "ry * (1 + ru\<^sup>2) * sin \<theta> - ru * (1 + ry\<^sup>2) * sin \<phi>"
let ?Ku = "(sin \<theta>) / (2 * sin(\<phi> - \<theta>)) * (1/ru + ru) - (sin \<phi>) / (2 * sin (\<phi> - \<theta>)) * (1/ry + ry)"
have yu: "poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor ru * cis \<phi>)) = (?Ayu, ?Byu, cnj ?Byu, ?Ayu)"
"is_real ?Ayu" "Re ?Ayu \<noteq> 0" "?Ayu \<noteq> 0" "cor ?Ayu \<noteq> 0" "Re ?Byu = ?ReByu" and Ku: "?ReByu / ?Ayu = ?Ku"
using **[OF \<open>0 < ru\<close>]
by auto
have "?Ayz \<noteq> 0"
using \<open>sin (\<phi> - \<theta>) \<noteq> 0\<close> \<open>ry > 0\<close> \<open>rz > 0\<close>
by auto
have "Re ?Byz / ?Ayz < -1"
using \<open>intersects_x_axis_positive ?yz\<close>
* \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>u' = cor ru * cis \<phi>\<close>
apply simp
apply (transfer fixing: ry rz ru \<theta> \<phi>)
apply (transfer fixing: ry rz ru \<theta> \<phi>)
proof-
assume "intersects_x_axis_positive_cmat (poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor rz * cis \<phi>)))"
thus "(ry * sin \<theta> * (1 + rz\<^sup>2) - rz * sin \<phi> * (1 + ry\<^sup>2)) / (2 * rz * ry * sin (\<phi> - \<theta>)) < - 1"
using yz
by simp
qed
have "?ReByz / ?Ayz \<ge> ?ReByu / ?Ayu"
proof (cases "sin \<phi> > 0")
case True
hence "sin \<theta> < 0"
using \<open>sin \<phi> * sin \<theta> < 0\<close>
by (smt mult_nonneg_nonneg)
have "?ReByz < 0"
proof-
have "ry * (1 + rz\<^sup>2) * sin \<theta> < 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
using \<open>sin \<theta> < 0\<close>
by (smt mult_pos_neg mult_pos_pos zero_less_power)
moreover
have "rz * (1 + ry\<^sup>2) * sin \<phi> > 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
using \<open>sin \<phi> > 0\<close>
by (smt mult_pos_neg mult_pos_pos zero_less_power)
ultimately
show ?thesis
by simp
qed
have "?Ayz > 0"
using \<open>Re ?Byz / ?Ayz < -1\<close> \<open>Re ?Byz = ?ReByz\<close> \<open>?ReByz < 0\<close>
by (smt divide_less_0_iff)
hence "sin (\<phi> - \<theta>) > 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
by (smt mult_pos_pos zero_less_mult_pos)
have "1 / ru + ru \<ge> 1 / rz + rz"
using Pasch_fun_mono[of ru rz] \<open>0 < ru\<close> \<open>ru \<le> rz\<close> \<open>rz < 1\<close>
by simp
hence "sin \<theta> * (1 / ru + ru) \<le> sin \<theta> * (1 / rz + rz)"
using \<open>sin \<theta> < 0\<close>
by auto
thus ?thesis
using \<open>ru > 0\<close> \<open>rz > 0\<close> \<open>ru \<le> rz\<close> \<open>rz < 1\<close> \<open>?Ayz > 0\<close> \<open>sin (\<phi> - \<theta>) > 0\<close>
using divide_right_mono[of "sin \<theta> * (1 / ru + ru)" "sin \<theta> * (1 / rz + rz)" "2 * sin (\<phi> - \<theta>)"]
by (subst Kz, subst Ku) simp
next
assume "\<not> sin \<phi> > 0"
hence "sin \<phi> < 0"
using \<open>sin \<phi> * sin \<theta> < 0\<close>
by (cases "sin \<phi> = 0", simp_all)
hence "sin \<theta> > 0"
using \<open>sin \<phi> * sin \<theta> < 0\<close>
by (smt mult_nonpos_nonpos)
have "?ReByz > 0"
proof-
have "ry * (1 + rz\<^sup>2) * sin \<theta> > 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
using \<open>sin \<theta> > 0\<close>
by (smt mult_pos_neg mult_pos_pos zero_less_power)
moreover
have "rz * (1 + ry\<^sup>2) * sin \<phi> < 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
using \<open>sin \<phi> < 0\<close>
by (smt mult_pos_neg mult_pos_pos zero_less_power)
ultimately
show ?thesis
by simp
qed
have "?Ayz < 0"
using \<open>Re ?Byz / ?Ayz < -1\<close> \<open>?Ayz \<noteq> 0\<close> \<open>Re ?Byz = ?ReByz\<close> \<open>?ReByz > 0\<close>
by (smt divide_less_0_iff)
hence "sin (\<phi> - \<theta>) < 0"
using \<open>ry > 0\<close> \<open>rz > 0\<close>
by (smt mult_nonneg_nonneg)
have "1 / ru + ru \<ge> 1 / rz + rz"
using Pasch_fun_mono[of ru rz] \<open>0 < ru\<close> \<open>ru \<le> rz\<close> \<open>rz < 1\<close>
by simp
hence "sin \<theta> * (1 / ru + ru) \<ge> sin \<theta> * (1 / rz + rz)"
using \<open>sin \<theta> > 0\<close>
by auto
thus ?thesis
using \<open>ru > 0\<close> \<open>rz > 0\<close> \<open>ru \<le> rz\<close> \<open>rz < 1\<close> \<open>?Ayz < 0\<close> \<open>sin (\<phi> - \<theta>) < 0\<close>
using divide_right_mono_neg[of "sin \<theta> * (1 / rz + rz)" "sin \<theta> * (1 / ru + ru)" "2 * sin (\<phi> - \<theta>)"]
by (subst Kz, subst Ku) simp
qed
have "intersects_x_axis_positive ?yu"
using * \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>u' = cor ru * cis \<phi>\<close>
apply simp
apply (transfer fixing: ry rz ru \<theta> \<phi>)
apply (transfer fixing: ry rz ru \<theta> \<phi>)
proof-
have "Re ?Byu / ?Ayu < -1"
using \<open>Re ?Byz / ?Ayz < -1\<close> \<open>?ReByz / ?Ayz \<ge> ?ReByu / ?Ayu\<close>
by (subst (asm) \<open>Re ?Byz = ?ReByz\<close>, subst \<open>Re ?Byu = ?ReByu\<close>) simp
thus "intersects_x_axis_positive_cmat (poincare_line_cvec_cmat (of_complex_cvec (cor ry * cis \<theta>)) (of_complex_cvec (cor ru * cis \<phi>)))"
using yu
by simp
qed
let ?a = "calc_x_axis_intersection ?yu"
have "?a \<in> positive_x_axis" "?a \<in> circline_set ?yu" "?a \<in> unit_disc"
using \<open>intersects_x_axis_positive ?yu\<close>
using intersects_x_axis_positive_iff'[of ?yu] \<open>y \<noteq> u\<close>
by auto
then obtain a' where a': "?a = of_complex a'" "is_real a'" "Re a' > 0" "Re a' < 1"
unfolding positive_x_axis_def circline_set_x_axis
by (auto simp add: cmod_eq_Re)
have "intersects_x_axis ?yz" "intersects_x_axis ?yu"
using \<open>intersects_x_axis_positive ?yz\<close> \<open>intersects_x_axis_positive ?yu\<close>
by auto
show "\<exists>a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between 0\<^sub>h a (of_complex v)"
proof (rule_tac x="?a" in exI, safe)
show "poincare_between u ?a y"
using poincare_between_x_axis_intersection[of y u ?a]
using calc_x_axis_intersection[OF is_poincare_line_poincare_line[OF \<open>y \<noteq> u\<close>] \<open>intersects_x_axis ?yu\<close>]
using calc_x_axis_intersection_in_unit_disc[OF is_poincare_line_poincare_line[OF \<open>y \<noteq> u\<close>] \<open>intersects_x_axis ?yu\<close>]
using in_disc \<open>y \<noteq> u\<close> \<open>y \<notin> circline_set x_axis\<close> \<open>u \<notin> circline_set x_axis\<close>
using * \<open>\<phi> = arg u'\<close> \<open>\<theta> = arg y'\<close> \<open>\<phi> * \<theta> < 0\<close>
by (subst poincare_between_rev, auto simp add: mult.commute)
next
show "poincare_between 0\<^sub>h ?a (of_complex v)"
proof-
have "-?ReByz / ?Ayz \<le> -?ReByu / ?Ayu"
using \<open>?ReByz / ?Ayz \<ge> ?ReByu / ?Ayu\<close>
by linarith
have "outward ?yz ?yu"
using * \<open>y' = cor ry * cis \<theta>\<close> \<open>z' = cor rz * cis \<phi>\<close> \<open>u' = cor ru * cis \<phi>\<close>
apply simp
apply (transfer fixing: ry rz ru \<theta> \<phi>)
apply (transfer fixing: ry rz ru \<theta> \<phi>)
apply (subst yz yu)+
unfolding outward_cmat_def
apply (simp only: Let_def prod.case)
apply (subst yz yu)+
using \<open>-?ReByz / ?Ayz \<le> -?ReByu / ?Ayu\<close>
by simp
hence "Re a' \<le> Re v"
using \<open>?v = of_complex v\<close>
using \<open>?a = of_complex a'\<close>
using \<open>intersects_x_axis_positive ?yz\<close> \<open>intersects_x_axis_positive ?yu\<close>
using outward[OF is_poincare_line_poincare_line[OF \<open>y \<noteq> z\<close>] is_poincare_line_poincare_line[OF \<open>y \<noteq> u\<close>]]
by simp
thus ?thesis
using \<open>?v = of_complex v\<close>
using poincare_between_x_axis_0uv[of "Re a'" "Re v"] a' v
by simp
qed
next
show "?a \<in> unit_disc"
by fact
qed
qed
next
show "x \<in> unit_disc" "v \<in> unit_disc" "x \<noteq> v"
using assms
by auto
next
fix M x v
let ?Mx = "moebius_pt M x" and ?Mv = "moebius_pt M v"
assume 1: "unit_disc_fix M" "x \<in> unit_disc" "v \<in> unit_disc" "x \<noteq> v"
assume 2: "?P ?Mx ?Mv"
show "?P x v"
proof safe
fix y z u
let ?My = "moebius_pt M y" and ?Mz = "moebius_pt M z" and ?Mu = "moebius_pt M u"
assume "distinct [x, y, z, u, v]" "\<not> poincare_collinear {x, y, z}" "y \<in> unit_disc" "z \<in> unit_disc" "u \<in> unit_disc"
"poincare_between x u z" "poincare_between y v z"
hence "\<exists> Ma. Ma \<in> unit_disc \<and> poincare_between ?Mu Ma ?My \<and> poincare_between ?Mx Ma ?Mv"
using 1 2[rule_format, of ?My ?Mz ?Mu]
by simp
then obtain Ma where Ma: "Ma \<in> unit_disc" "poincare_between ?Mu Ma ?My \<and> poincare_between ?Mx Ma ?Mv"
by blast
let ?a = "moebius_pt (-M) Ma"
let ?Ma = "moebius_pt M ?a"
have "?Ma = Ma"
by (metis moebius_pt_invert uminus_moebius_def)
hence "?Ma \<in> unit_disc" "poincare_between ?Mu ?Ma ?My \<and> poincare_between ?Mx ?Ma ?Mv"
using Ma
by auto
thus "\<exists>a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between x a v"
using unit_disc_fix_moebius_inv[OF \<open>unit_disc_fix M\<close>] \<open>unit_disc_fix M\<close> \<open>Ma \<in> unit_disc\<close>
using \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close> \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
by (rule_tac x="?a" in exI, simp del: moebius_pt_comp_inv_right)
qed
qed
thus ?thesis
using assms
by auto
qed
text\<open>Pasch axiom, only degenerative cases.\<close>
lemma Pasch_deg:
assumes "x \<in> unit_disc" and "y \<in> unit_disc" and "z \<in> unit_disc" and "u \<in> unit_disc" and "v \<in> unit_disc"
assumes "\<not> distinct [x, y, z, u, v] \<or> poincare_collinear {x, y, z}"
assumes "poincare_between x u z" and "poincare_between y v z"
shows "\<exists> a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between x a v"
proof(cases "poincare_collinear {x, y, z}")
case True
hence "poincare_between x y z \<or> poincare_between y x z \<or> poincare_between y z x"
using assms(1, 2, 3) poincare_collinear3_between poincare_between_rev by blast
show ?thesis
proof(cases "poincare_between x y z")
case True
have "poincare_between x y v"
using True assms poincare_between_transitivity
by (meson poincare_between_rev)
thus ?thesis
using assms(2)
by (rule_tac x="y" in exI, simp)
next
case False
hence "poincare_between y x z \<or> poincare_between y z x"
using \<open>poincare_between x y z \<or> poincare_between y x z \<or> poincare_between y z x\<close>
by simp
show ?thesis
proof(cases "poincare_between y x z")
case True
hence "poincare_between u x y"
using assms
by (meson poincare_between_rev poincare_between_transitivity)
thus ?thesis
using assms
by (rule_tac x="x" in exI, simp)
next
case False
hence "poincare_between y z x"
using \<open>poincare_between y x z \<or> poincare_between y z x\<close>
by auto
hence "poincare_between x z v"
using assms
by (meson poincare_between_rev poincare_between_transitivity)
hence "poincare_between x u v"
using assms poincare_between_transitivity poincare_between_rev
by (smt poincare_between_sum_distances)
thus ?thesis
using assms
by (rule_tac x="u" in exI, simp)
qed
qed
next
case False
hence "\<not> distinct [x, y, z, u, v]"
using assms(6) by auto
show ?thesis
proof(cases "u=z")
case True
thus ?thesis
using assms
apply(rule_tac x="v" in exI)
by(simp add:poincare_between_rev)
next
case False (* "u \<noteq> z" *)
hence "x \<noteq> z"
using assms poincare_between_sandwich by blast
show ?thesis
proof(cases "v=z")
case True
thus ?thesis
using assms
by (rule_tac x="u" in exI, simp)
next
case False (* v \<noteq> z *)
hence "y \<noteq> z"
using assms poincare_between_sandwich by blast
show ?thesis
proof(cases "u = x")
case True
thus ?thesis
using assms
by (rule_tac x="x" in exI, simp)
next
case False (*u \<noteq> x*)
have "x \<noteq> y"
using assms \<open>\<not> poincare_collinear {x, y, z}\<close>
by fastforce
have "x \<noteq> v"
using assms \<open>\<not> poincare_collinear {x, y, z}\<close>
by (metis insert_commute poincare_between_poincare_collinear)
have "u \<noteq> y"
using assms \<open>\<not> poincare_collinear {x, y, z}\<close>
using poincare_between_poincare_collinear by blast
have "u \<noteq> v"
proof(rule ccontr)
assume "\<not> u \<noteq> v"
hence "poincare_between x v z"
using assms by auto
hence "x \<in> circline_set (poincare_line z v)"
using poincare_between_rev[of x v z]
using poincare_between_poincare_line_uvz[of z v x]
using assms \<open>v \<noteq> z\<close>
by auto
have "y \<in> circline_set (poincare_line z v)"
using assms \<open>\<not> u \<noteq> v\<close>
using poincare_between_rev[of y v z]
using poincare_between_poincare_line_uvz[of z v y]
using assms \<open>v \<noteq> z\<close>
by auto
have "z \<in> circline_set (poincare_line z v)"
using ex_poincare_line_two_points[of z v] \<open>v \<noteq> z\<close>
by auto
have "is_poincare_line (poincare_line z v)"
using \<open>v \<noteq> z\<close>
by auto
hence "poincare_collinear {x, y, z}"
using \<open>x \<in> circline_set (poincare_line z v)\<close>
using \<open>y \<in> circline_set (poincare_line z v)\<close>
using \<open>z \<in> circline_set (poincare_line z v)\<close>
unfolding poincare_collinear_def
by (rule_tac x="poincare_line z v" in exI, simp)
thus False
using \<open>\<not> poincare_collinear {x, y, z}\<close> by simp
qed
have "v = y"
using \<open>u \<noteq> v\<close> \<open>u \<noteq> y\<close> \<open>x \<noteq> v\<close> \<open>x \<noteq> y\<close> \<open>u \<noteq> x\<close> \<open>y \<noteq> z\<close> \<open>v \<noteq> z\<close> \<open>x \<noteq> z\<close> \<open>u \<noteq> z\<close>
using \<open>\<not> distinct [x, y, z, u, v]\<close>
by auto
thus ?thesis
using assms
by (rule_tac x="y" in exI, simp)
qed
qed
qed
qed
text \<open>Axiom of Pasch\<close>
lemma Pasch:
assumes "x \<in> unit_disc" and "y \<in> unit_disc" and "z \<in> unit_disc" and "u \<in> unit_disc" and "v \<in> unit_disc"
assumes "poincare_between x u z" and "poincare_between y v z"
shows "\<exists> a. a \<in> unit_disc \<and> poincare_between u a y \<and> poincare_between x a v"
proof(cases "distinct [x, y, z, u, v] \<and> \<not> poincare_collinear {x, y, z}")
case True
thus ?thesis
using assms Pasch_nondeg by auto
next
case False
thus ?thesis
using assms Pasch_deg by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Segment construction axiom\<close>
(* ------------------------------------------------------------------ *)
lemma segment_construction:
assumes "x \<in> unit_disc" and "y \<in> unit_disc"
assumes "a \<in> unit_disc" and "b \<in> unit_disc"
shows "\<exists> z. z \<in> unit_disc \<and> poincare_between x y z \<and> poincare_distance y z = poincare_distance a b"
proof-
obtain d where d: "d = poincare_distance a b"
by auto
have "d \<ge> 0"
using assms
by (simp add: d poincare_distance_ge0)
have "\<exists> z. z \<in> unit_disc \<and> poincare_between x y z \<and> poincare_distance y z = d" (is "?P x y")
proof (cases "x = y")
case True
have "\<exists> z. z \<in> unit_disc \<and> poincare_distance x z = d"
proof (rule wlog_zero)
show "\<exists> z. z \<in> unit_disc \<and> poincare_distance 0\<^sub>h z = d"
using ex_x_axis_poincare_distance_negative[of d] \<open>d \<ge> 0\<close>
by blast
next
show "x \<in> unit_disc"
by fact
next
fix a u
assume "u \<in> unit_disc" "cmod a < 1"
assume "\<exists>z. z \<in> unit_disc \<and> poincare_distance (moebius_pt (blaschke a) u) z = d"
then obtain z where *: "z \<in> unit_disc" "poincare_distance (moebius_pt (blaschke a) u) z = d"
by auto
obtain z' where z': "z = moebius_pt (blaschke a) z'" "z' \<in> unit_disc"
using \<open>z \<in> unit_disc\<close>
using unit_disc_fix_iff[of "blaschke a"] \<open>cmod a < 1\<close>
using blaschke_unit_disc_fix[of a]
by blast
show "\<exists>z. z \<in> unit_disc \<and> poincare_distance u z = d"
using * z' \<open>u : unit_disc\<close>
using blaschke_unit_disc_fix[of a] \<open>cmod a < 1\<close>
by (rule_tac x=z' in exI, simp)
qed
thus ?thesis
using \<open>x = y\<close>
unfolding poincare_between_def
by auto
next
case False
show ?thesis
proof (rule wlog_positive_x_axis[where P="\<lambda> y x. ?P x y"])
fix x
assume "is_real x" "0 < Re x" "Re x < 1"
then obtain z where z: "is_real z" "Re z \<le> 0" "- 1 < Re z" "of_complex z \<in> unit_disc"
"of_complex z \<in> unit_disc" "of_complex z \<in> circline_set x_axis" "poincare_distance 0\<^sub>h (of_complex z) = d"
using ex_x_axis_poincare_distance_negative[of d] \<open>d \<ge> 0\<close>
by auto
have "poincare_between (of_complex x) 0\<^sub>h (of_complex z)"
proof (cases "z = 0")
case True
thus ?thesis
unfolding poincare_between_def
by auto
next
case False
have "x \<noteq> 0"
using \<open>is_real x\<close> \<open>Re x > 0\<close>
by auto
thus ?thesis
using poincare_between_x_axis_u0v[of x z]
using z \<open>is_real x\<close> \<open>x \<noteq> 0\<close> \<open>Re x > 0\<close> False
using complex_eq_if_Re_eq mult_pos_neg
by fastforce
qed
thus "?P (of_complex x) 0\<^sub>h"
using \<open>poincare_distance 0\<^sub>h (of_complex z) = d\<close> \<open>of_complex z \<in> unit_disc\<close>
by blast
next
show "x \<in> unit_disc" "y \<in> unit_disc"
by fact+
next
show "y \<noteq> x" using \<open>x \<noteq> y\<close> by simp
next
fix M u v
assume "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assume "?P (moebius_pt M v) (moebius_pt M u)"
then obtain z where *: "z \<in> unit_disc" "poincare_between (moebius_pt M v) (moebius_pt M u) z" "poincare_distance (moebius_pt M u) z = d"
by auto
obtain z' where z': "z = moebius_pt M z'" "z' \<in> unit_disc"
using \<open>z \<in> unit_disc\<close>
using unit_disc_fix_iff[of M] \<open>unit_disc_fix M\<close>
by blast
thus "?P v u"
using * \<open>u \<in> unit_disc\<close> \<open>v \<in> unit_disc\<close> \<open>unit_disc_fix M\<close>
by auto
qed
qed
thus ?thesis
using assms d
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Five segment axiom\<close>
(* ------------------------------------------------------------------ *)
lemma five_segment_axiom:
assumes
in_disc: "x \<in> unit_disc" "y \<in> unit_disc" "z \<in> unit_disc" "u \<in> unit_disc" and
in_disc': "x' \<in> unit_disc" "y' \<in> unit_disc" "z' \<in> unit_disc" "u' \<in> unit_disc" and
"x \<noteq> y" and
betw: "poincare_between x y z" "poincare_between x' y' z'" and
xy: "poincare_distance x y = poincare_distance x' y'" and
xu: "poincare_distance x u = poincare_distance x' u'" and
yu: "poincare_distance y u = poincare_distance y' u'" and
yz: "poincare_distance y z = poincare_distance y' z'"
shows
"poincare_distance z u = poincare_distance z' u'"
proof-
from assms obtain M where
M: "unit_disc_fix_f M" "M x = x'" "M u = u'" "M y = y'"
using unit_disc_fix_f_congruent_triangles[of x y u]
by blast
have "M z = z'"
proof (rule unique_poincare_distance_on_ray[where u=x' and v=y' and y="M z" and z=z' and d="poincare_distance x z"])
show "0 \<le> poincare_distance x z"
using poincare_distance_ge0 in_disc
by simp
next
show "x' \<noteq> y'"
using M \<open>x \<noteq> y\<close>
using in_disc in_disc' poincare_distance_eq_0_iff xy
by auto
next
show "poincare_distance x' (M z) = poincare_distance x z"
using M in_disc
unfolding unit_disc_fix_f_def
by auto
next
show "M z \<in> unit_disc"
using M in_disc
unfolding unit_disc_fix_f_def
by auto
next
show "poincare_distance x' z' = poincare_distance x z"
using xy yz betw
using poincare_between_sum_distances[of x y z]
using poincare_between_sum_distances[of x' y' z']
using in_disc in_disc'
by auto
next
show "poincare_between x' y' (M z)"
using M
using in_disc betw
unfolding unit_disc_fix_f_def
by auto
qed fact+
thus ?thesis
using \<open>unit_disc_fix_f M\<close>
using in_disc in_disc'
\<open>M u = u'\<close>
unfolding unit_disc_fix_f_def
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Upper dimension axiom\<close>
(* ------------------------------------------------------------------ *)
lemma upper_dimension_axiom:
assumes in_disc: "x \<in> unit_disc" "y \<in> unit_disc" "z \<in> unit_disc" "u \<in> unit_disc" "v \<in> unit_disc"
assumes "poincare_distance x u = poincare_distance x v"
"poincare_distance y u = poincare_distance y v"
"poincare_distance z u = poincare_distance z v"
"u \<noteq> v"
shows "poincare_between x y z \<or> poincare_between y z x \<or> poincare_between z x y"
proof (cases "x = y \<or> y = z \<or> x = z")
case True
thus ?thesis
using in_disc
by auto
next
case False
hence "x \<noteq> y" "x \<noteq> z" "y \<noteq> z"
by auto
let ?cong = "\<lambda> a b a' b'. poincare_distance a b = poincare_distance a' b'"
have "\<forall> z u v. z \<in> unit_disc \<and> u \<in> unit_disc \<and> v \<in> unit_disc \<and>
?cong x u x v \<and> ?cong y u y v \<and> ?cong z u z v \<and> u \<noteq> v \<longrightarrow>
poincare_collinear {x, y, z}" (is "?P x y")
proof (rule wlog_positive_x_axis[where P="?P"])
fix x
assume x: "is_real x" "0 < Re x" "Re x < 1"
hence "x \<noteq> 0"
by auto
have "0\<^sub>h \<in> circline_set x_axis"
by simp
show "?P 0\<^sub>h (of_complex x)"
proof safe
fix z u v
assume in_disc: "z \<in> unit_disc" "u \<in> unit_disc" "v \<in> unit_disc"
then obtain z' u' v' where zuv: "z = of_complex z'" "u = of_complex u'" "v = of_complex v'"
using inf_or_of_complex[of z] inf_or_of_complex[of u] inf_or_of_complex[of v]
by auto
assume cong: "?cong 0\<^sub>h u 0\<^sub>h v" "?cong (of_complex x) u (of_complex x) v" "?cong z u z v" "u \<noteq> v"
let ?r0 = "poincare_distance 0\<^sub>h u" and
?rx = "poincare_distance (of_complex x) u"
have "?r0 > 0" "?rx > 0"
using in_disc cong
using poincare_distance_eq_0_iff[of "0\<^sub>h" u] poincare_distance_ge0[of "0\<^sub>h" u]
using poincare_distance_eq_0_iff[of "0\<^sub>h" v] poincare_distance_ge0[of "0\<^sub>h" v]
using poincare_distance_eq_0_iff[of "of_complex x" u] poincare_distance_ge0[of "of_complex x" u]
using poincare_distance_eq_0_iff[of "of_complex x" v] poincare_distance_ge0[of "of_complex x" v]
using x
by (auto simp add: cmod_eq_Re)
let ?pc0 = "poincare_circle 0\<^sub>h ?r0" and
?pcx = "poincare_circle (of_complex x) ?rx"
have "u \<in> ?pc0 \<inter> ?pcx" "v \<in> ?pc0 \<inter> ?pcx"
using in_disc cong
by (auto simp add: poincare_circle_def)
hence "u = conjugate v"
using intersect_poincare_circles_x_axis[of 0 x ?r0 ?rx u v]
using x \<open>x \<noteq> 0\<close> \<open>u \<noteq> v\<close> \<open>?r0 > 0\<close> \<open>?rx > 0\<close>
by simp
let ?ru = "poincare_distance u z"
have "?ru > 0"
using poincare_distance_ge0[of u z] in_disc
using cong
using poincare_distance_eq_0_iff[of z u] poincare_distance_eq_0_iff[of z v]
using poincare_distance_eq_0_iff
by force
have "z \<in> poincare_circle u ?ru \<inter> poincare_circle v ?ru"
using cong in_disc
unfolding poincare_circle_def
by (simp add: poincare_distance_sym)
hence "is_real z'"
using intersect_poincare_circles_conjugate_centers[of u v ?ru z] \<open>u = conjugate v\<close> zuv
using in_disc \<open>u \<noteq> v\<close> \<open>?ru > 0\<close>
by simp
thus "poincare_collinear {0\<^sub>h, of_complex x, z}"
using poincare_line_0_real_is_x_axis[of "of_complex x"] x \<open>x \<noteq> 0\<close> zuv \<open>0\<^sub>h \<in> circline_set x_axis\<close>
unfolding poincare_collinear_def
by (rule_tac x=x_axis in exI, auto simp add: circline_set_x_axis)
qed
next
fix M x y
assume 1: "unit_disc_fix M" "x \<in> unit_disc" "y \<in> unit_disc" "x \<noteq> y"
assume 2: "?P (moebius_pt M x) (moebius_pt M y)"
show "?P x y"
proof safe
fix z u v
assume "z \<in> unit_disc" "u \<in> unit_disc" "v \<in> unit_disc"
"?cong x u x v" "?cong y u y v" "?cong z u z v" "u \<noteq> v"
hence "poincare_collinear {moebius_pt M x, moebius_pt M y, moebius_pt M z}"
using 1 2[rule_format, of "moebius_pt M z" "moebius_pt M u" "moebius_pt M v"]
by simp
then obtain p where "is_poincare_line p" "{moebius_pt M x, moebius_pt M y, moebius_pt M z} \<subseteq> circline_set p"
unfolding poincare_collinear_def
by auto
thus "poincare_collinear {x, y, z}"
using \<open>unit_disc_fix M\<close>
unfolding poincare_collinear_def
by (rule_tac x="moebius_circline (-M) p" in exI, auto)
qed
qed fact+
thus ?thesis
using assms
using poincare_collinear3_between[of x y z]
using poincare_between_rev
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Lower dimension axiom\<close>
(* ------------------------------------------------------------------ *)
lemma lower_dimension_axiom:
shows "\<exists> a \<in> unit_disc. \<exists> b \<in> unit_disc. \<exists> c \<in> unit_disc.
\<not> poincare_between a b c \<and> \<not> poincare_between b c a \<and> \<not> poincare_between c a b"
proof-
let ?u = "of_complex (1/2)" and ?v = "of_complex (\<i>/2)"
have 1: "0\<^sub>h \<in> unit_disc" and 2: "?u \<in> unit_disc" and 3: "?v \<in> unit_disc"
by simp_all
have *: "\<not> poincare_collinear {0\<^sub>h, ?u, ?v}"
proof (rule ccontr)
assume "\<not> ?thesis"
then obtain p where "is_poincare_line p" "{0\<^sub>h, ?u, ?v} \<subseteq> circline_set p"
unfolding poincare_collinear_def
by auto
moreover
have "of_complex (1 / 2) \<noteq> of_complex (\<i> / 2)"
using of_complex_inj
by fastforce
ultimately
have "0\<^sub>h \<in> circline_set (poincare_line ?u ?v)"
using unique_poincare_line[of ?u ?v p]
by auto
thus False
unfolding circline_set_def
by simp (transfer, transfer, simp add: vec_cnj_def)
qed
show ?thesis
apply (rule_tac x="0\<^sub>h" in bexI, rule_tac x="?u" in bexI, rule_tac x="?v" in bexI)
apply (rule ccontr, auto)
using *
using poincare_between_poincare_collinear[OF 1 2 3]
using poincare_between_poincare_collinear[OF 2 3 1]
using poincare_between_poincare_collinear[OF 3 1 2]
by (metis insert_commute)+
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Negated Euclidean axiom\<close>
(* ------------------------------------------------------------------ *)
lemma negated_euclidean_axiom_aux:
assumes "on_circline H (of_complex (1/2 + \<i>/2))" and "is_poincare_line H"
assumes "intersects_x_axis_positive H"
shows "\<not> intersects_y_axis_positive H"
using assms
proof (transfer, transfer)
fix H
assume hh: "hermitean H \<and> H \<noteq> mat_zero" "is_poincare_line_cmat H"
obtain A B C D where "H = (A, B, C, D)"
by (cases H, auto)
hence *: "is_real A" "H = (A, B, cnj B, A)" "(cmod B)\<^sup>2 > (cmod A)\<^sup>2"
using hermitean_elems[of A B C D] hh
by auto
assume "intersects_x_axis_positive_cmat H"
hence "Re A \<noteq> 0" "Re B / Re A < - 1"
using *
by auto
assume "on_circline_cmat_cvec H (of_complex_cvec (1 / 2 + \<i> / 2))"
hence "6*A + 4*Re B + 4*Im B = 0"
using *
unfolding cor_mult
apply (subst Re_express_cnj[of B])
apply (subst Im_express_cnj[of B])
apply (simp add: vec_cnj_def)
apply (simp add: field_simps)
done
hence "Re (6*A + 4*Re B + 4*Im B) = 0"
by simp
hence "3*Re A + 2*Re B + 2*Im B = 0"
using \<open>is_real A\<close>
by simp
hence "3/2 + Re B/Re A + Im B/Re A = 0"
using \<open>Re A \<noteq> 0\<close>
by (simp add: field_simps)
hence "-Im B/Re A - 3/2 < -1"
using \<open>Re B / Re A < -1\<close>
by simp
hence "Im B/Re A > -1/2"
by (simp add: field_simps)
thus "\<not> intersects_y_axis_positive_cmat H"
using *
by simp
qed
lemma negated_euclidean_axiom:
shows "\<exists> a b c d t.
a \<in> unit_disc \<and> b \<in> unit_disc \<and> c \<in> unit_disc \<and> d \<in> unit_disc \<and> t \<in> unit_disc \<and>
poincare_between a d t \<and> poincare_between b d c \<and> a \<noteq> d \<and>
(\<forall> x y. x \<in> unit_disc \<and> y \<in> unit_disc \<and>
poincare_between a b x \<and> poincare_between x t y \<longrightarrow> \<not> poincare_between a c y)"
proof-
let ?a = "0\<^sub>h"
let ?b = "of_complex (1/2)"
let ?c = "of_complex (\<i>/2)"
let ?dl = "(5 - sqrt 17) / 4"
let ?d = "of_complex (?dl + \<i>*?dl)"
let ?t = "of_complex (1/2 + \<i>/2)"
have "?dl \<noteq> 0"
proof-
have "(sqrt 17)\<^sup>2 \<noteq> 5\<^sup>2"
by simp
hence "sqrt 17 \<noteq> 5"
by force
thus ?thesis
by simp
qed
have "?d \<noteq> ?a"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "?dl + \<i>*?dl = 0"
by simp
hence "Re (?dl + \<i>*?dl) = 0"
by simp
thus False
using \<open>?dl \<noteq> 0\<close>
by simp
qed
have "?dl > 0"
proof-
have "(sqrt 17)\<^sup>2 < 5\<^sup>2"
by (simp add: power2_eq_square)
hence "sqrt 17 < 5"
by (rule power2_less_imp_less, simp)
thus ?thesis
by simp
qed
have "?a \<noteq> ?b"
by (metis divide_eq_0_iff of_complex_zero_iff zero_neq_numeral zero_neq_one)
have "?a \<noteq> ?c"
by (metis complex_i_not_zero divide_eq_0_iff of_complex_zero_iff zero_neq_numeral)
show ?thesis
proof (rule_tac x="?a" in exI, rule_tac x="?b" in exI, rule_tac x="?c" in exI, rule_tac x="?d" in exI, rule_tac x="?t" in exI, safe)
show "?a \<in> unit_disc" "?b \<in> unit_disc" "?c \<in> unit_disc" "?t \<in> unit_disc"
by (auto simp add: cmod_def power2_eq_square)
have cmod_d: "cmod (?dl + \<i>*?dl) = ?dl * sqrt 2"
using \<open>?dl > 0\<close>
unfolding cmod_def
by (simp add: real_sqrt_mult)
show "?d \<in> unit_disc"
proof-
have "?dl < 1 / sqrt 2"
proof-
have "17\<^sup>2 < (5 * sqrt 17)\<^sup>2"
by (simp add: field_simps)
hence "17 < 5 * sqrt 17"
by (rule power2_less_imp_less, simp)
hence "?dl\<^sup>2 < (1 / sqrt 2)\<^sup>2"
by (simp add: power2_eq_square field_simps)
thus "?dl < 1 / sqrt 2"
by (rule power2_less_imp_less, simp)
qed
thus ?thesis
using cmod_d
by (simp add: field_simps)
qed
have cmod_d: "1 - (cmod (to_complex ?d))\<^sup>2 = (-17 + 5*sqrt 17) / 4" (is "_ = ?cmod_d")
apply (simp only: to_complex_of_complex)
apply (subst cmod_d)
apply (simp add: power_mult_distrib)
apply (simp add: power2_eq_square field_simps)
done
have cmod_d_c: "(cmod (to_complex ?d - to_complex ?c))\<^sup>2 = (17 - 4*sqrt 17) / 4" (is "_ = ?cmod_dc")
unfolding cmod_square
by (simp add: field_simps)
have cmod_c: "1 - (cmod (to_complex ?c))\<^sup>2 = 3/4" (is "_ = ?cmod_c")
by (simp add: power2_eq_square)
have xx: "\<And> x::real. x + x = 2*x"
by simp
have "cmod ((to_complex ?b) - (to_complex ?d)) = cmod ((to_complex ?d) - (to_complex ?c))"
by (simp add: cmod_def power2_eq_square field_simps)
moreover
have "cmod (to_complex ?b) = cmod (to_complex ?c)"
by simp
ultimately
have *: "poincare_distance_formula' (to_complex ?b) (to_complex ?d) =
poincare_distance_formula' (to_complex ?d) (to_complex ?c)"
unfolding poincare_distance_formula'_def
by simp
have **: "poincare_distance_formula' (to_complex ?d) (to_complex ?c) = (sqrt 17) / 3"
unfolding poincare_distance_formula'_def
proof (subst cmod_d, subst cmod_c, subst cmod_d_c)
have "(sqrt 17 * 15)\<^sup>2 \<noteq> 51\<^sup>2"
by simp
hence "sqrt 17 * 15 \<noteq> 51"
by force
hence "sqrt 17 * 15 - 51 \<noteq> 0"
by simp
have "(5 * sqrt 17)\<^sup>2 \<noteq> 17\<^sup>2"
by simp
hence "5 * sqrt 17 \<noteq> 17"
by force
hence "?cmod_d * ?cmod_c \<noteq> 0"
by simp
hence "1 + 2 * (?cmod_dc / (?cmod_d * ?cmod_c)) = (?cmod_d * ?cmod_c + 2 * ?cmod_dc) / (?cmod_d * ?cmod_c)"
using add_frac_num[of "?cmod_d * ?cmod_c" "2 * ?cmod_dc" 1]
by (simp add: field_simps)
also have "... = (64 * (85 - sqrt 17 * 17)) / (64 * (sqrt 17 * 15 - 51))"
by (simp add: field_simps)
also have "... = (85 - sqrt 17 * 17) / (sqrt 17 * 15 - 51)"
by (rule mult_divide_mult_cancel_left, simp)
also have "... = sqrt 17 / 3"
by (subst frac_eq_eq, fact, simp, simp add: field_simps)
finally
show "1 + 2 * (?cmod_dc / (?cmod_d * ?cmod_c)) = sqrt 17 / 3"
.
qed
have "sqrt 17 \<ge> 3"
proof-
have "(sqrt 17)\<^sup>2 \<ge> 3\<^sup>2"
by simp
thus ?thesis
by (rule power2_le_imp_le, simp)
qed
thus "poincare_between ?b ?d ?c"
unfolding poincare_between_sum_distances[OF \<open>?b \<in> unit_disc\<close> \<open>?d \<in> unit_disc\<close> \<open>?c \<in> unit_disc\<close>]
unfolding poincare_distance_formula[OF \<open>?b \<in> unit_disc\<close> \<open>?d \<in> unit_disc\<close>]
unfolding poincare_distance_formula[OF \<open>?d \<in> unit_disc\<close> \<open>?c \<in> unit_disc\<close>]
unfolding poincare_distance_formula[OF \<open>?b \<in> unit_disc\<close> \<open>?c \<in> unit_disc\<close>]
unfolding poincare_distance_formula_def
apply (subst *, subst xx, subst **, subst arcosh_double)
apply (simp_all add: cmod_def power2_eq_square)
done
show "poincare_between ?a ?d ?t"
proof (subst poincare_between_0uv[OF \<open>?d \<in> unit_disc\<close> \<open>?t \<in> unit_disc\<close> \<open>?d \<noteq> ?a\<close>])
show "?t \<noteq> 0\<^sub>h"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "1/2 + \<i>/2 = 0"
by simp
hence "Re (1/2 + \<i>/2) = 0"
by simp
thus False
by simp
qed
next
have "19\<^sup>2 \<le> (5 * sqrt 17)\<^sup>2"
by simp
hence "19 \<le> 5 * sqrt 17"
by (rule power2_le_imp_le, simp)
hence "cmod (to_complex ?d) \<le> cmod (to_complex ?t)"
by (simp add: Let_def cmod_def power2_eq_square field_simps)
moreover
have "arg (to_complex ?d) = arg (to_complex ?t)"
proof-
have 1: "to_complex ?d = ((5 - sqrt 17) / 4) * (1 + \<i>)"
by (simp add: field_simps)
have 2: "to_complex ?t = (cor (1/2)) * (1 + \<i>)"
by (simp add: field_simps)
have "(sqrt 17)\<^sup>2 < 5\<^sup>2"
by simp
hence "sqrt 17 < 5"
by (rule power2_less_imp_less, simp)
hence 3: "(5 - sqrt 17) / 4 > 0"
by simp
have 4: "(1::real) / 2 > 0"
by simp
show ?thesis
apply (subst 1, subst 2)
apply (subst arg_mult_real_positive[OF 3])
apply (subst arg_mult_real_positive[OF 4])
by simp
qed
ultimately
show "let d' = to_complex ?d; t' = to_complex ?t in arg d' = arg t' \<and> cmod d' \<le> cmod t'"
by simp
qed
show "?a = ?d \<Longrightarrow> False"
using \<open>?d \<noteq> ?a\<close>
by simp
fix x y
assume "x \<in> unit_disc" "y \<in> unit_disc"
assume abx: "poincare_between ?a ?b x"
hence "x \<in> circline_set x_axis"
using poincare_between_poincare_line_uvz[of ?a ?b x] \<open>x \<in> unit_disc\<close> \<open>?a \<noteq> ?b\<close>
using poincare_line_0_real_is_x_axis[of ?b]
by (auto simp add: circline_set_x_axis)
have "x \<noteq> 0\<^sub>h"
using abx poincare_between_sandwich[of ?a ?b] \<open>?a \<noteq> ?b\<close>
by auto
have "x \<in> positive_x_axis"
using \<open>x \<in> circline_set x_axis\<close> \<open>x \<noteq> 0\<^sub>h\<close> \<open>x \<in> unit_disc\<close>
using abx poincare_between_x_axis_0uv[of "1/2" "Re (to_complex x)"]
unfolding circline_set_x_axis positive_x_axis_def
by (auto simp add: cmod_eq_Re abs_less_iff complex_eq_if_Re_eq)
assume acy: "poincare_between ?a ?c y"
hence "y \<in> circline_set y_axis"
using poincare_between_poincare_line_uvz[of ?a ?c y] \<open>y \<in> unit_disc\<close> \<open>?a \<noteq> ?c\<close>
using poincare_line_0_imag_is_y_axis[of ?c]
by (auto simp add: circline_set_y_axis)
have "y \<noteq> 0\<^sub>h"
using acy poincare_between_sandwich[of ?a ?c] \<open>?a \<noteq> ?c\<close>
by auto
have "y \<in> positive_y_axis"
proof-
have " \<And>x. \<lbrakk>x \<noteq> 0; poincare_between 0\<^sub>h (of_complex (\<i> / 2)) (of_complex x); is_imag x; - 1 < Im x\<rbrakk> \<Longrightarrow> 0 < Im x"
by (smt add.left_neutral complex.expand divide_complex_def complex_eq divide_less_0_1_iff divide_less_eq_1_pos imaginary_unit.simps(1) mult.left_neutral of_real_1 of_real_add of_real_divide of_real_eq_0_iff one_add_one poincare_between_y_axis_0uv zero_complex.simps(1) zero_complex.simps(2) zero_less_divide_1_iff)
thus ?thesis
using \<open>y \<in> circline_set y_axis\<close> \<open>y \<noteq> 0\<^sub>h\<close> \<open>y \<in> unit_disc\<close>
using acy
unfolding circline_set_y_axis positive_y_axis_def
by (auto simp add: cmod_eq_Im abs_less_iff)
qed
have "x \<noteq> y"
using \<open>x \<in> positive_x_axis\<close> \<open>y \<in> positive_y_axis\<close>
unfolding positive_x_axis_def positive_y_axis_def circline_set_x_axis circline_set_y_axis
by auto
assume xty: "poincare_between x ?t y"
let ?xy = "poincare_line x y"
have "?t \<in> circline_set ?xy"
using xty poincare_between_poincare_line_uzv[OF \<open>x \<noteq> y\<close> \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>?t \<in> unit_disc\<close>]
by simp
moreover
have "?xy \<noteq> x_axis"
using poincare_line_circline_set[OF \<open>x \<noteq> y\<close>] \<open>y \<in> positive_y_axis\<close>
by (auto simp add: circline_set_x_axis positive_y_axis_def)
hence "intersects_x_axis_positive ?xy"
using intersects_x_axis_positive_iff[of "?xy"] \<open>x \<noteq> y\<close> \<open>x \<in> unit_disc\<close> \<open>x \<in> positive_x_axis\<close>
by auto
moreover
have "?xy \<noteq> y_axis"
using poincare_line_circline_set[OF \<open>x \<noteq> y\<close>] \<open>x \<in> positive_x_axis\<close>
by (auto simp add: circline_set_y_axis positive_x_axis_def)
hence "intersects_y_axis_positive ?xy"
using intersects_y_axis_positive_iff[of "?xy"] \<open>x \<noteq> y\<close> \<open>y \<in> unit_disc\<close> \<open>y \<in> positive_y_axis\<close>
by auto
ultimately
show False
using negated_euclidean_axiom_aux[of ?xy] \<open>x \<noteq> y\<close>
unfolding circline_set_def
by auto
qed
qed
text \<open>Alternate form of the Euclidean axiom -- this one is much easier to prove\<close>
lemma negated_euclidean_axiom':
shows "\<exists> a b c.
a \<in> unit_disc \<and> b \<in> unit_disc \<and> c \<in> unit_disc \<and> \<not>(poincare_collinear {a, b, c}) \<and>
\<not>(\<exists> x. x \<in> unit_disc \<and>
poincare_distance a x = poincare_distance b x \<and>
poincare_distance a x = poincare_distance c x)"
proof-
let ?a = "of_complex (\<i>/2)"
let ?b = "of_complex (-\<i>/2)"
let ?c = "of_complex (1/5)"
have "(\<i>/2) \<noteq> (-\<i>/2)"
by simp
hence "?a \<noteq> ?b"
by (metis to_complex_of_complex)
have "(\<i>/2) \<noteq> (1/5)"
by simp
hence "?a \<noteq> ?c"
by (metis to_complex_of_complex)
have "(-\<i>/2) \<noteq> (1/5)"
by (metis add.inverse_inverse cmod_divide div_by_1 divide_divide_eq_right inverse_eq_divide minus_divide_left mult.commute norm_ii norm_minus_cancel norm_numeral norm_one numeral_One numeral_eq_iff semiring_norm(88))
hence "?b \<noteq> ?c"
by (metis to_complex_of_complex)
have "?a \<in> unit_disc" "?b \<in> unit_disc" "?c \<in> unit_disc"
by auto
moreover
have "\<not>(poincare_collinear {?a, ?b, ?c})"
unfolding poincare_collinear_def
proof(rule ccontr)
assume " \<not> (\<nexists>p. is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p)"
then obtain p where "is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p"
by auto
let ?ab = "poincare_line ?a ?b"
have "p = ?ab"
using \<open>is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p\<close>
using unique_poincare_line[of ?a ?b] \<open>?a \<noteq> ?b\<close> \<open>?a \<in> unit_disc\<close> \<open>?b \<in> unit_disc\<close>
by auto
have "?c \<notin> circline_set ?ab"
proof(rule ccontr)
assume "\<not> ?c \<notin> circline_set ?ab"
have "poincare_between ?a 0\<^sub>h ?b"
unfolding poincare_between_def
using cross_ratio_0inf by auto
hence "0\<^sub>h \<in> circline_set ?ab"
using \<open>?a \<noteq> ?b\<close> \<open>?a \<in> unit_disc\<close> \<open>?b \<in> unit_disc\<close>
using poincare_between_poincare_line_uzv zero_in_unit_disc
by blast
hence "?ab = poincare_line 0\<^sub>h ?a"
using unique_poincare_line[of ?a ?b] \<open>?a \<noteq> ?b\<close> \<open>?a \<in> unit_disc\<close> \<open>?b \<in> unit_disc\<close>
using \<open>is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p\<close>
using \<open>p = ?ab\<close> poincare_line_circline_set(1) unique_poincare_line
by (metis add.inverse_neutral divide_minus_left of_complex_zero_iff zero_in_unit_disc)
hence "(\<i>/2) * cnj(1/5) = cnj(\<i>/2) * (1/5)"
using poincare_collinear_zero_iff[of "(\<i>/2)" "(1/5)"]
using \<open>?a \<noteq> ?c\<close> \<open>\<not> ?c \<notin> circline_set ?ab\<close> \<open>?a \<in> unit_disc\<close> \<open>?c \<in> unit_disc\<close> \<open>p = ?ab\<close>
using \<open>0\<^sub>h \<in> circline_set ?ab\<close> \<open>is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p\<close>
using poincare_collinear_def by auto
thus False
by simp
qed
thus False
using \<open>p = ?ab\<close> \<open>is_poincare_line p \<and> {?a, ?b, ?c} \<subseteq> circline_set p\<close>
by auto
qed
moreover
have "\<not>(\<exists> x. x \<in> unit_disc \<and>
poincare_distance ?a x = poincare_distance ?b x \<and>
poincare_distance ?a x = poincare_distance ?c x)"
proof(rule ccontr)
assume "\<not> ?thesis"
then obtain x where "x \<in> unit_disc" "poincare_distance ?a x = poincare_distance ?b x"
"poincare_distance ?a x = poincare_distance ?c x"
by blast
let ?x = "to_complex x"
have "poincare_distance_formula' (\<i>/2) ?x = poincare_distance_formula' (-\<i>/2) ?x"
using \<open>poincare_distance ?a x = poincare_distance ?b x\<close>
using \<open>x \<in> unit_disc\<close> \<open>?a \<in> unit_disc\<close> \<open>?b \<in> unit_disc\<close>
by (metis cosh_dist to_complex_of_complex)
hence "(cmod (\<i> / 2 - ?x))\<^sup>2 = (cmod (- \<i> / 2 - ?x))\<^sup>2"
unfolding poincare_distance_formula'_def
apply (simp add:field_simps)
using \<open>x \<in> unit_disc\<close> unit_disc_cmod_square_lt_1 by fastforce
hence "Im ?x = 0"
unfolding cmod_def
by (simp add: power2_eq_iff)
have "1 - (Re ?x)\<^sup>2 \<noteq> 0"
using \<open>x \<in> unit_disc\<close> unit_disc_cmod_square_lt_1
using cmod_power2 by force
hence "24 - 24 * (Re ?x)\<^sup>2 \<noteq> 0"
by simp
have "poincare_distance_formula' (\<i>/2) ?x = poincare_distance_formula' (1/5) ?x"
using \<open>poincare_distance ?a x = poincare_distance ?c x\<close>
using \<open>x \<in> unit_disc\<close> \<open>?a \<in> unit_disc\<close> \<open>?c \<in> unit_disc\<close>
by (metis cosh_dist to_complex_of_complex)
hence "(2 + 8 * (Re ?x)\<^sup>2) /(3 - 3 * (Re ?x)\<^sup>2) = 2 * (1 - Re ?x * 5)\<^sup>2 / (24 - 24 * (Re ?x)\<^sup>2)" (is "?lhs = ?rhs")
unfolding poincare_distance_formula'_def
apply (simp add:field_simps)
unfolding cmod_def
using \<open>Im ?x = 0\<close>
by (simp add:field_simps)
hence *: "?lhs * (24 - 24 * (Re ?x)\<^sup>2) = ?rhs * (24 - 24 * (Re ?x)\<^sup>2) "
using \<open>(24 - 24 * (Re ?x)\<^sup>2) \<noteq> 0\<close>
by simp
have "?lhs * (24 - 24 * (Re ?x)\<^sup>2) = (2 + 8 * (Re ?x)\<^sup>2) * 8"
using \<open>(24 - 24 * (Re ?x)\<^sup>2) \<noteq> 0\<close> \<open>1 - (Re ?x)\<^sup>2 \<noteq> 0\<close>
by (simp add:field_simps)
have "?rhs * (24 - 24 * (Re ?x)\<^sup>2) = 2 * (1 - Re ?x * 5)\<^sup>2"
using \<open>(24 - 24 * (Re ?x)\<^sup>2) \<noteq> 0\<close> \<open>1 - (Re ?x)\<^sup>2 \<noteq> 0\<close>
by (simp add:field_simps)
hence "(2 + 8 * (Re ?x)\<^sup>2) * 8 = 2 * (1 - Re ?x * 5)\<^sup>2"
using * \<open>?lhs * (24 - 24 * (Re ?x)\<^sup>2) = (2 + 8 * (Re ?x)\<^sup>2) * 8\<close>
by simp
hence "7 * (Re ?x)\<^sup>2 + 10 * (Re ?x) + 7 = 0"
by (simp add:field_simps comm_ring_1_class.power2_diff)
thus False
using discriminant_iff[of 7 "Re (to_complex x)" 10 7] discrim_def[of 7 10 7]
by auto
qed
ultimately show ?thesis
apply (rule_tac x="?a" in exI)
apply (rule_tac x="?b" in exI)
apply (rule_tac x="?c" in exI)
by auto
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Continuity axiom\<close>
(* ------------------------------------------------------------------ *)
text \<open>The set $\phi$ is on the left of the set $\psi$\<close>
abbreviation set_order where
"set_order A \<phi> \<psi> \<equiv> \<forall>x\<in> unit_disc. \<forall>y\<in> unit_disc. \<phi> x \<and> \<psi> y \<longrightarrow> poincare_between A x y"
text \<open>The point $B$ is between the sets $\phi$ and $\psi$\<close>
abbreviation point_between_sets where
"point_between_sets \<phi> B \<psi> \<equiv> \<forall>x\<in> unit_disc. \<forall>y\<in> unit_disc. \<phi> x \<and> \<psi> y \<longrightarrow> poincare_between x B y"
lemma continuity:
assumes "\<exists> A \<in> unit_disc. set_order A \<phi> \<psi>"
shows "\<exists> B \<in> unit_disc. point_between_sets \<phi> B \<psi>"
proof (cases "(\<exists> x0 \<in> unit_disc. \<phi> x0) \<and> (\<exists> y0 \<in> unit_disc. \<psi> y0)")
case False
thus ?thesis
using assms by blast
next
case True
then obtain Y0 where "\<psi> Y0" "Y0 \<in> unit_disc"
by auto
obtain A where *: "A \<in> unit_disc" "set_order A \<phi> \<psi>"
using assms
by auto
show ?thesis
proof(cases "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> x = A")
case True
thus ?thesis
using \<open>A \<in> unit_disc\<close>
using poincare_between_nonstrict(1) by blast
next
case False
then obtain X0 where "\<phi> X0" "X0 \<noteq> A" "X0 \<in> unit_disc"
by auto
have "Y0 \<noteq> A"
proof(rule ccontr)
assume "\<not> Y0 \<noteq> A"
hence "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> poincare_between A x A"
using * \<open>\<psi> Y0\<close>
by (cases A) force
hence "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> x = A"
using * poincare_between_sandwich by blast
thus False
using False by auto
qed
show ?thesis
proof (cases "\<exists> B \<in> unit_disc. \<phi> B \<and> \<psi> B")
case True
then obtain B where "B \<in> unit_disc" "\<phi> B" "\<psi> B"
by auto
hence "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> poincare_between A x B"
using * by auto
have "\<forall> y \<in> unit_disc. \<psi> y \<longrightarrow> poincare_between A B y"
using * \<open>B \<in> unit_disc\<close> \<open>\<phi> B\<close>
by auto
show ?thesis
proof(rule+)
show "B \<in> unit_disc"
by fact
next
fix x y
assume "x \<in> unit_disc" "y \<in> unit_disc" "\<phi> x \<and> \<psi> y"
hence "poincare_between A x B" "poincare_between A B y"
using \<open>\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> poincare_between A x B\<close>
using \<open>\<forall> y \<in> unit_disc. \<psi> y \<longrightarrow> poincare_between A B y\<close>
by simp+
thus "poincare_between x B y"
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>B \<in> unit_disc\<close> \<open>A \<in> unit_disc\<close>
using poincare_between_transitivity[of A x B y]
by simp
qed
next
case False
have "poincare_between A X0 Y0"
using \<open>\<phi> X0\<close> \<open>\<psi> Y0\<close> * \<open>Y0 \<in> unit_disc\<close> \<open>X0 \<in> unit_disc\<close>
by auto
have "\<forall> \<phi>. \<forall> \<psi>. set_order A \<phi> \<psi> \<and> \<not> (\<exists> B \<in> unit_disc. \<phi> B \<and> \<psi> B) \<and> \<phi> X0 \<and>
(\<exists> y \<in> unit_disc. \<psi> y) \<and> (\<exists> x \<in> unit_disc. \<phi> x)
\<longrightarrow> (\<exists> B \<in> unit_disc. point_between_sets \<phi> B \<psi>)"
(is "?P A X0")
proof (rule wlog_positive_x_axis[where P="?P"])
show "A \<in> unit_disc"
by fact
next
show "X0 \<in> unit_disc"
by fact
next
show "A \<noteq> X0"
using \<open>X0 \<noteq> A\<close> by simp
next
fix M u v
let ?M = "\<lambda> x. moebius_pt M x"
let ?Mu = "?M u" and ?Mv = "?M v"
assume hip: "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
"?P ?Mu ?Mv"
show "?P u v"
proof safe
fix \<phi> \<psi> x y
assume "set_order u \<phi> \<psi>" "\<not> (\<exists>B\<in>unit_disc. \<phi> B \<and> \<psi> B)" "\<phi> v"
"y \<in> unit_disc" "\<psi> y" "x \<in> unit_disc" "\<phi> x"
let ?M\<phi> = "\<lambda> X'. \<exists> X. \<phi> X \<and> ?M X = X'"
let ?M\<psi> = "\<lambda> X'. \<exists> X. \<psi> X \<and> ?M X = X'"
obtain M\<phi> where "M\<phi> = ?M\<phi>" by simp
obtain M\<psi> where "M\<psi> = ?M\<psi>" by simp
have "M\<phi> ?Mv"
using \<open>\<phi> v\<close> using \<open>M\<phi> = ?M\<phi>\<close>
by blast
moreover
have "\<not> (\<exists> B \<in>unit_disc. M\<phi> B \<and> M\<psi> B)"
using \<open>\<not> (\<exists>B\<in>unit_disc. \<phi> B \<and> \<psi> B)\<close>
using \<open>M\<phi> = ?M\<phi>\<close> \<open>M\<psi> = ?M\<psi>\<close>
by (metis hip(1) moebius_pt_invert unit_disc_fix_discI unit_disc_fix_moebius_inv)
moreover
have "\<exists> y \<in> unit_disc. M\<psi> y"
using \<open>y \<in> unit_disc\<close> \<open>\<psi> y\<close> \<open>M\<psi> = ?M\<psi>\<close> \<open>unit_disc_fix M\<close>
by auto
moreover
have "set_order ?Mu ?M\<phi> ?M\<psi>"
proof ((rule ballI)+, rule impI)
fix Mx My
assume "Mx \<in> unit_disc" "My \<in> unit_disc" "?M\<phi> Mx \<and> ?M\<psi> My"
then obtain x y where "\<phi> x \<and> ?M x = Mx" "\<psi> y \<and> ?M y = My"
by blast
hence "x \<in> unit_disc" "y \<in> unit_disc"
using \<open>Mx \<in> unit_disc\<close> \<open>My \<in> unit_disc\<close> \<open>unit_disc_fix M\<close>
by (metis moebius_pt_comp_inv_left unit_disc_fix_discI unit_disc_fix_moebius_inv)+
hence "poincare_between u x y"
using \<open>set_order u \<phi> \<psi>\<close>
using \<open>Mx \<in> unit_disc\<close> \<open>My \<in> unit_disc\<close> \<open>\<phi> x \<and> ?M x = Mx\<close> \<open>\<psi> y \<and> ?M y = My\<close>
by blast
then show "poincare_between ?Mu Mx My"
using \<open>\<phi> x \<and> ?M x = Mx\<close> \<open>\<psi> y \<and> ?M y = My\<close>
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>u \<in> unit_disc\<close> \<open>unit_disc_fix M\<close>
using unit_disc_fix_moebius_preserve_poincare_between by blast
qed
hence "set_order ?Mu M\<phi> M\<psi>"
using \<open>M\<phi> = ?M\<phi>\<close> \<open>M\<psi> = ?M\<psi>\<close>
by simp
ultimately
have "\<exists> Mb \<in> unit_disc. point_between_sets M\<phi> Mb M\<psi>"
using hip(5)
by blast
then obtain Mb where bbb:
"Mb \<in> unit_disc" "point_between_sets ?M\<phi> Mb ?M\<psi>"
using \<open>M\<phi> = ?M\<phi>\<close> \<open>M\<psi> = ?M\<psi>\<close>
by auto
let ?b = "moebius_pt (moebius_inv M) Mb"
show "\<exists> b \<in> unit_disc. point_between_sets \<phi> b \<psi>"
proof (rule_tac x="?b" in bexI, (rule ballI)+, rule impI)
fix x y
assume "x \<in> unit_disc" "y \<in> unit_disc" "\<phi> x \<and> \<psi> y"
hence "poincare_between u x y"
using \<open>set_order u \<phi> \<psi>\<close>
by blast
let ?Mx = "?M x" and ?My = "?M y"
have "?M\<phi> ?Mx" "?M\<psi> ?My"
using \<open>\<phi> x \<and> \<psi> y\<close>
by blast+
have "?Mx \<in> unit_disc" "?My \<in> unit_disc"
using \<open>x \<in> unit_disc\<close> \<open>unit_disc_fix M\<close> \<open>y \<in> unit_disc\<close>
by auto
hence "poincare_between ?Mx Mb ?My"
using \<open>?M\<phi> ?Mx\<close> \<open>?M\<psi> ?My\<close> \<open>?Mx \<in> unit_disc\<close> \<open>?My \<in> unit_disc\<close> bbb
by auto
then show "poincare_between x ?b y"
using \<open>unit_disc_fix M\<close>
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>Mb \<in> unit_disc\<close> \<open>?Mx \<in> unit_disc\<close> \<open>?My \<in> unit_disc\<close>
using unit_disc_fix_moebius_preserve_poincare_between[of M x ?b y]
by auto
next
show "?b \<in> unit_disc"
using bbb \<open>unit_disc_fix M\<close>
by auto
qed
qed
next
fix X
assume xx: "is_real X" "0 < Re X" "Re X < 1"
let ?X = "of_complex X"
show "?P 0\<^sub>h ?X"
proof ((rule allI)+, rule impI, (erule conjE)+)
fix \<phi> \<psi>
assume "set_order 0\<^sub>h \<phi> \<psi>" "\<not> (\<exists>B\<in>unit_disc. \<phi> B \<and> \<psi> B)" "\<phi> ?X"
"\<exists>y\<in>unit_disc. \<psi> y" "\<exists>x\<in>unit_disc. \<phi> x"
have "?X \<in> unit_disc"
using xx
by (simp add: cmod_eq_Re)
have \<psi>pos: "\<forall> y \<in> unit_disc. \<psi> y \<longrightarrow> (is_real (to_complex y) \<and> Re (to_complex y) > 0)"
proof(rule ballI, rule impI)
fix y
let ?y = "to_complex y"
assume "y \<in> unit_disc" "\<psi> y"
hence "poincare_between 0\<^sub>h ?X y"
using \<open>set_order 0\<^sub>h \<phi> \<psi>\<close>
using \<open>?X \<in> unit_disc\<close> \<open>\<phi> ?X\<close>
by auto
thus "is_real ?y \<and> 0 < Re ?y"
using xx \<open>?X \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
by (metis (mono_tags, hide_lams) arg_0_iff of_complex_zero_iff poincare_between_0uv poincare_between_sandwich to_complex_of_complex unit_disc_to_complex_inj zero_in_unit_disc)
qed
have \<phi>noneg: "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> (is_real (to_complex x) \<and> Re (to_complex x) \<ge> 0)"
proof(rule ballI, rule impI)
fix x
assume "x \<in> unit_disc" "\<phi> x"
obtain y where "y \<in> unit_disc" "\<psi> y"
using \<open>\<exists> y \<in> unit_disc. \<psi> y\<close> by blast
let ?x = "to_complex x" and ?y = "to_complex y"
have "is_real ?y" "Re ?y > 0"
using \<psi>pos \<open>\<psi> y\<close> \<open>y \<in> unit_disc\<close>
by auto
have "poincare_between 0\<^sub>h x y"
using \<open>set_order 0\<^sub>h \<phi> \<psi>\<close>
using \<open>x \<in> unit_disc\<close> \<open>\<phi> x\<close> \<open>y\<in>unit_disc\<close> \<open>\<psi> y\<close>
by auto
thus "is_real ?x \<and> 0 \<le> Re ?x"
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>is_real (to_complex y)\<close> \<open>\<psi> y\<close>
using \<open>set_order 0\<^sub>h \<phi> \<psi>\<close>
using \<open>\<phi> ?X\<close> \<open>?X \<in> unit_disc\<close> \<open>Re ?y > 0\<close>
by (metis arg_0_iff le_less of_complex_zero poincare_between_0uv to_complex_of_complex zero_complex.simps(1) zero_complex.simps(2))
qed
have \<phi>less\<psi>: "\<forall>x\<in>unit_disc. \<forall>y\<in>unit_disc. \<phi> x \<and> \<psi> y \<longrightarrow> Re (to_complex x) < Re (to_complex y)"
proof((rule ballI)+, rule impI)
fix x y
let ?x = "to_complex x" and ?y = "to_complex y"
assume "x \<in> unit_disc" "y \<in> unit_disc" "\<phi> x \<and> \<psi> y"
hence "poincare_between 0\<^sub>h x y"
using \<open>set_order 0\<^sub>h \<phi> \<psi>\<close>
by auto
moreover
have "is_real ?x" "Re ?x \<ge> 0"
using \<phi>noneg
using \<open>x \<in> unit_disc\<close> \<open>\<phi> x \<and> \<psi> y\<close> by auto
moreover
have "is_real ?y" "Re ?y > 0"
using \<psi>pos
using \<open>y \<in> unit_disc\<close> \<open>\<phi> x \<and> \<psi> y\<close> by auto
ultimately
have "Re ?x \<le> Re ?y"
using \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
by (metis Re_complex_of_real arg_0_iff le_less of_complex_zero poincare_between_0uv rcis_cmod_arg rcis_zero_arg to_complex_of_complex)
have "Re ?x \<noteq> Re ?y"
using \<open>\<phi> x \<and> \<psi> y\<close> \<open>is_real ?x\<close> \<open>is_real ?y\<close>
using \<open>\<not> (\<exists>B\<in>unit_disc. \<phi> B \<and> \<psi> B)\<close> \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close>
by (metis complex.expand unit_disc_to_complex_inj)
thus "Re ?x < Re ?y"
using \<open>Re ?x \<le> Re ?y\<close> by auto
qed
have "\<exists> b \<in> unit_disc. \<forall> x \<in> unit_disc. \<forall> y \<in> unit_disc.
is_real (to_complex b) \<and>
(\<phi> x \<and> \<psi> y \<longrightarrow> (Re (to_complex x) \<le> Re (to_complex b) \<and> Re (to_complex b) \<le> Re (to_complex y)))"
proof-
let ?Phi = "{x. (of_complex (cor x)) \<in> unit_disc \<and> \<phi> (of_complex (cor x))}"
have "\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> Re (to_complex x) \<le> Sup ?Phi"
proof(safe)
fix x
let ?x = "to_complex x"
assume "x \<in> unit_disc" "\<phi> x"
hence "is_real ?x" "Re ?x \<ge> 0"
using \<phi>noneg
by auto
hence "cor (Re ?x) = ?x"
using complex_of_real_Re by blast
hence "of_complex (cor (Re ?x)) \<in> unit_disc"
using \<open>x \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
moreover
have "\<phi> (of_complex (cor (Re ?x)))"
using \<open>cor (Re ?x) = ?x\<close> \<open>\<phi> x\<close> \<open>x \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
ultimately
have "Re ?x \<in> ?Phi"
by auto
have "\<exists>M. \<forall>x \<in> ?Phi. x \<le> M"
using \<phi>less\<psi>
using \<open>\<exists> y \<in> unit_disc. \<psi> y\<close>
by (metis (mono_tags, lifting) Re_complex_of_real le_less mem_Collect_eq to_complex_of_complex)
thus "Re ?x \<le> Sup ?Phi"
using cSup_upper[of "Re ?x" ?Phi]
unfolding bdd_above_def
using \<open>Re ?x \<in> ?Phi\<close>
by auto
qed
have "\<forall> y \<in> unit_disc. \<psi> y \<longrightarrow> Sup ?Phi \<le> Re (to_complex y)"
proof (safe)
fix y
let ?y = "to_complex y"
assume "\<psi> y" "y \<in> unit_disc"
show "Sup ?Phi \<le> Re ?y"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "Re ?y < Sup ?Phi"
by auto
have "\<exists> x. \<phi> (of_complex (cor x)) \<and> (of_complex (cor x)) \<in> unit_disc"
proof -
obtain x' where "x' \<in> unit_disc" "\<phi> x'"
using \<open>\<exists> x \<in> unit_disc. \<phi> x\<close> by blast
let ?x' = "to_complex x'"
have "is_real ?x'"
using \<open>x' \<in> unit_disc\<close> \<open>\<phi> x'\<close>
using \<phi>noneg
by auto
hence "cor (Re ?x') = ?x'"
using complex_of_real_Re by blast
hence "x' = of_complex (cor (Re ?x'))"
using \<open>x' \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
show ?thesis
apply (rule_tac x="Re ?x'" in exI)
using \<open>x' \<in> unit_disc\<close>
apply (subst (asm) \<open>x' = of_complex (cor (Re ?x'))\<close>, simp)
using \<open>\<phi> x'\<close>
by (subst (asm) (2) \<open>x' = of_complex (cor (Re ?x'))\<close>, simp)
qed
hence "?Phi \<noteq> {}"
by auto
then obtain x where "\<phi> (of_complex (cor x))" "Re ?y < x"
"(of_complex (cor x)) \<in> unit_disc"
using \<open>Re ?y < Sup ?Phi\<close>
using less_cSupE[of "Re ?y" ?Phi]
by auto
moreover
have "Re ?y < Re (to_complex (of_complex (cor x)))"
using \<open>Re ?y < x\<close>
by simp
ultimately
show False
using \<phi>less\<psi>
using \<open>\<psi> y\<close> \<open>y \<in> unit_disc\<close>
by (metis less_not_sym)
qed
qed
thus ?thesis
using \<open>\<forall> x \<in> unit_disc. \<phi> x \<longrightarrow> Re (to_complex x) \<le> Sup ?Phi\<close>
apply (rule_tac x="(of_complex (cor (Sup ?Phi)))" in bexI, simp)
using \<open>\<exists>y\<in>unit_disc. \<psi> y\<close> \<open>\<phi> ?X\<close> \<open>?X \<in> unit_disc\<close>
using \<open>\<forall>y\<in>unit_disc. \<psi> y \<longrightarrow> is_real (to_complex y) \<and> 0 < Re (to_complex y)\<close>
by (smt complex_of_real_Re inf_notin_unit_disc norm_of_real of_complex_to_complex to_complex_of_complex unit_disc_iff_cmod_lt_1 xx(2))
qed
then obtain B where "B \<in> unit_disc" "is_real (to_complex B)"
"\<forall>x\<in>unit_disc. \<forall>y\<in>unit_disc. \<phi> x \<and> \<psi> y \<longrightarrow> Re (to_complex x) \<le> Re (to_complex B) \<and>
Re (to_complex B) \<le> Re (to_complex y)"
by blast
show "\<exists> b \<in> unit_disc. point_between_sets \<phi> b \<psi>"
proof (rule_tac x="B" in bexI)
show "B \<in> unit_disc"
by fact
next
show "point_between_sets \<phi> B \<psi>"
proof ((rule ballI)+, rule impI)
fix x y
let ?x = "to_complex x" and ?y = "to_complex y" and ?B = "to_complex B"
assume "x \<in> unit_disc" "y \<in> unit_disc" "\<phi> x \<and> \<psi> y"
hence "Re ?x \<le> Re ?B \<and> Re ?B \<le> Re ?y"
using \<open>\<forall>x\<in>unit_disc. \<forall>y\<in>unit_disc. \<phi> x \<and> \<psi> y \<longrightarrow> Re (to_complex x) \<le> Re ?B \<and>
Re (to_complex B) \<le> Re (to_complex y)\<close>
by auto
moreover
have "is_real ?x" "Re ?x \<ge> 0"
using \<phi>noneg
using \<open>x \<in> unit_disc\<close> \<open>\<phi> x \<and> \<psi> y\<close>
by auto
moreover
have "is_real ?y" "Re ?y > 0"
using \<psi>pos
using \<open>y \<in> unit_disc\<close> \<open>\<phi> x \<and> \<psi> y\<close>
by auto
moreover
have "cor (Re ?x) = ?x"
using complex_of_real_Re \<open>is_real ?x\<close> by blast
hence "x = of_complex (cor (Re ?x))"
using \<open>x \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
moreover
have "cor (Re ?y) = ?y"
using complex_of_real_Re \<open>is_real ?y\<close> by blast
hence "y = of_complex (cor (Re ?y))"
using \<open>y \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
moreover
have "cor (Re ?B) = ?B"
using complex_of_real_Re \<open>is_real (to_complex B)\<close> by blast
hence "B = of_complex (cor (Re ?B))"
using \<open>B \<in> unit_disc\<close>
by (metis inf_notin_unit_disc of_complex_to_complex)
ultimately
show "poincare_between x B y"
using \<open>is_real (to_complex B)\<close> \<open>x \<in> unit_disc\<close> \<open>y \<in> unit_disc\<close> \<open>B \<in> unit_disc\<close>
using poincare_between_x_axis_uvw[of "Re (to_complex x)" "Re (to_complex B)" "Re (to_complex y)"]
by (smt Re_complex_of_real arg_0_iff poincare_between_nonstrict(1) rcis_cmod_arg rcis_zero_arg unit_disc_iff_cmod_lt_1)
qed
qed
qed
qed
thus ?thesis
using False \<open>\<phi> X0\<close> \<open>\<psi> Y0\<close> * \<open>Y0 \<in> unit_disc\<close> \<open>X0 \<in> unit_disc\<close>
by auto
qed
qed
qed
(* ------------------------------------------------------------------ *)
subsection\<open>Limiting parallels axiom\<close>
(* ------------------------------------------------------------------ *)
text \<open>Auxiliary definitions\<close>
definition poincare_on_line where
"poincare_on_line p a b \<longleftrightarrow> poincare_collinear {p, a, b}"
definition poincare_on_ray where
"poincare_on_ray p a b \<longleftrightarrow> poincare_between a p b \<or> poincare_between a b p"
definition poincare_in_angle where
"poincare_in_angle p a b c \<longleftrightarrow>
b \<noteq> a \<and> b \<noteq> c \<and> p \<noteq> b \<and> (\<exists> x \<in> unit_disc. poincare_between a x c \<and> x \<noteq> a \<and> x \<noteq> c \<and> poincare_on_ray p b x)"
definition poincare_ray_meets_line where
"poincare_ray_meets_line a b c d \<longleftrightarrow> (\<exists> x \<in> unit_disc. poincare_on_ray x a b \<and> poincare_on_line x c d)"
text \<open>All points on ray are collinear\<close>
lemma poincare_on_ray_poincare_collinear:
assumes "p \<in> unit_disc" and "a \<in> unit_disc" and "b \<in> unit_disc" and "poincare_on_ray p a b"
shows "poincare_collinear {p, a, b}"
using assms poincare_between_poincare_collinear
unfolding poincare_on_ray_def
by (metis insert_commute)
text \<open>H-isometries preserve all defined auxiliary relations\<close>
lemma unit_disc_fix_preserves_poincare_on_line [simp]:
assumes "unit_disc_fix M" and "p \<in> unit_disc" "a \<in> unit_disc" "b \<in> unit_disc"
shows "poincare_on_line (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) \<longleftrightarrow> poincare_on_line p a b"
using assms
unfolding poincare_on_line_def
by auto
lemma unit_disc_fix_preserves_poincare_on_ray [simp]:
assumes "unit_disc_fix M" "p \<in> unit_disc" "a \<in> unit_disc" "b \<in> unit_disc"
shows "poincare_on_ray (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) \<longleftrightarrow> poincare_on_ray p a b"
using assms
unfolding poincare_on_ray_def
by auto
lemma unit_disc_fix_preserves_poincare_in_angle [simp]:
assumes "unit_disc_fix M" "p \<in> unit_disc" "a \<in> unit_disc" "b \<in> unit_disc" "c \<in> unit_disc"
shows "poincare_in_angle (moebius_pt M p) (moebius_pt M a) (moebius_pt M b) (moebius_pt M c) \<longleftrightarrow> poincare_in_angle p a b c" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume "?lhs"
then obtain Mx where *: "Mx \<in> unit_disc"
"poincare_between (moebius_pt M a) Mx (moebius_pt M c)"
"Mx \<noteq> moebius_pt M a" "Mx \<noteq> moebius_pt M c" "poincare_on_ray (moebius_pt M p) (moebius_pt M b) Mx"
"moebius_pt M b \<noteq> moebius_pt M a" "moebius_pt M b \<noteq> moebius_pt M c" "moebius_pt M p \<noteq> moebius_pt M b"
unfolding poincare_in_angle_def
by auto
obtain x where "Mx = moebius_pt M x" "x \<in> unit_disc"
by (metis "*"(1) assms(1) image_iff unit_disc_fix_iff)
thus ?rhs
using * assms
unfolding poincare_in_angle_def
by auto
next
assume ?rhs
then obtain x where *: "x \<in> unit_disc"
"poincare_between a x c"
"x \<noteq> a" "x \<noteq> c" "poincare_on_ray p b x"
"b \<noteq> a" "b \<noteq> c" "p \<noteq> b"
unfolding poincare_in_angle_def
by auto
thus ?lhs
using assms
unfolding poincare_in_angle_def
by auto (rule_tac x="moebius_pt M x" in bexI, auto)
qed
lemma unit_disc_fix_preserves_poincare_ray_meets_line [simp]:
assumes "unit_disc_fix M" "a \<in> unit_disc" "b \<in> unit_disc" "c \<in> unit_disc" "d \<in> unit_disc"
shows "poincare_ray_meets_line (moebius_pt M a) (moebius_pt M b) (moebius_pt M c) (moebius_pt M d) \<longleftrightarrow> poincare_ray_meets_line a b c d" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
then obtain Mx where *: "Mx \<in> unit_disc" "poincare_on_ray Mx (moebius_pt M a) (moebius_pt M b)"
"poincare_on_line Mx (moebius_pt M c) (moebius_pt M d)"
unfolding poincare_ray_meets_line_def
by auto
obtain x where "Mx = moebius_pt M x" "x \<in> unit_disc"
by (metis "*"(1) assms(1) image_iff unit_disc_fix_iff)
thus ?rhs
using assms *
unfolding poincare_ray_meets_line_def poincare_on_line_def
by auto
next
assume ?rhs
then obtain x where *: "x \<in> unit_disc" "poincare_on_ray x a b"
"poincare_on_line x c d"
unfolding poincare_ray_meets_line_def
by auto
thus ?lhs
using assms *
unfolding poincare_ray_meets_line_def poincare_on_line_def
by auto (rule_tac x="moebius_pt M x" in bexI, auto)
qed
text \<open>H-lines that intersect on the absolute do not meet (they do not share a common h-point)\<close>
lemma tangent_not_meet:
assumes "x1 \<in> unit_disc" and "x2 \<in> unit_disc" and "x1 \<noteq> x2" and "\<not> poincare_collinear {0\<^sub>h, x1, x2}"
assumes "i \<in> ideal_points (poincare_line x1 x2)" "a \<in> unit_disc" "a \<noteq> 0\<^sub>h" "poincare_collinear {0\<^sub>h, a, i}"
shows "\<not> poincare_ray_meets_line 0\<^sub>h a x1 x2"
proof (rule ccontr)
assume "\<not> ?thesis"
then obtain x where "x \<in> unit_disc" "poincare_on_ray x 0\<^sub>h a" "poincare_collinear {x, x1, x2}"
unfolding poincare_ray_meets_line_def poincare_on_line_def
by auto
have "poincare_collinear {0\<^sub>h, a, x}"
- using `poincare_on_ray x 0\<^sub>h a` `x \<in> unit_disc` `a \<in> unit_disc`
+ using \<open>poincare_on_ray x 0\<^sub>h a\<close> \<open>x \<in> unit_disc\<close> \<open>a \<in> unit_disc\<close>
by (meson poincare_between_poincare_collinear poincare_between_rev poincare_on_ray_def poincare_on_ray_poincare_collinear zero_in_unit_disc)
have "x \<noteq> 0\<^sub>h"
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}` `poincare_collinear {x, x1, x2}`
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close> \<open>poincare_collinear {x, x1, x2}\<close>
unfolding poincare_collinear_def
by (auto simp add: assms(2) assms(3) poincare_between_rev)
let ?l1 = "poincare_line 0\<^sub>h a"
let ?l2 = "poincare_line x1 x2"
have "i \<in> circline_set unit_circle"
- using `i \<in> ideal_points (poincare_line x1 x2)`
+ using \<open>i \<in> ideal_points (poincare_line x1 x2)\<close>
using assms(3) ideal_points_on_unit_circle is_poincare_line_poincare_line by blast
have "i \<in> circline_set ?l1"
- using `poincare_collinear {0\<^sub>h, a, i}`
+ using \<open>poincare_collinear {0\<^sub>h, a, i}\<close>
unfolding poincare_collinear_def
using \<open>a \<in> unit_disc\<close> \<open>a \<noteq> 0\<^sub>h\<close>
by (metis insert_subset unique_poincare_line zero_in_unit_disc)
moreover
have "x \<in> circline_set ?l1"
- using `a \<in> unit_disc` `a \<noteq> 0\<^sub>h` `poincare_collinear {0\<^sub>h, a, x}` `x \<in> unit_disc`
+ using \<open>a \<in> unit_disc\<close> \<open>a \<noteq> 0\<^sub>h\<close> \<open>poincare_collinear {0\<^sub>h, a, x}\<close> \<open>x \<in> unit_disc\<close>
by (metis poincare_collinear3_between poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_line_sym zero_in_unit_disc)
moreover
have "inversion x \<in> circline_set ?l1"
- using `poincare_collinear {0\<^sub>h, a, x}`
- using poincare_line_inversion_full[of "0\<^sub>h" a x] `a \<in> unit_disc` `a \<noteq> 0\<^sub>h` `x \<in> unit_disc`
+ using \<open>poincare_collinear {0\<^sub>h, a, x}\<close>
+ using poincare_line_inversion_full[of "0\<^sub>h" a x] \<open>a \<in> unit_disc\<close> \<open>a \<noteq> 0\<^sub>h\<close> \<open>x \<in> unit_disc\<close>
by (metis poincare_collinear3_between is_poincare_line_inverse_point is_poincare_line_poincare_line poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_line_sym zero_in_unit_disc)
moreover
have "x \<in> circline_set ?l2"
- using `poincare_collinear {x, x1, x2}` `x1 \<noteq> x2` `x1 \<in> unit_disc` `x2 \<in> unit_disc` `x \<in> unit_disc`
+ using \<open>poincare_collinear {x, x1, x2}\<close> \<open>x1 \<noteq> x2\<close> \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>x \<in> unit_disc\<close>
by (metis insert_commute inversion_noteq_unit_disc poincare_between_poincare_line_uvz poincare_between_poincare_line_uzv poincare_collinear3_iff poincare_line_sym_general)
moreover
hence "inversion x \<in> circline_set ?l2"
- using `x1 \<noteq> x2` `x1 \<in> unit_disc` `x2 \<in> unit_disc` `x \<in> unit_disc`
+ using \<open>x1 \<noteq> x2\<close> \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>x \<in> unit_disc\<close>
using poincare_line_inversion_full[of x1 x2 x]
unfolding circline_set_def
by auto
moreover
have "i \<in> circline_set ?l2"
- using `x1 \<noteq> x2` `x1 \<in> unit_disc` `x2 \<in> unit_disc`
- using `i \<in> ideal_points ?l2`
+ using \<open>x1 \<noteq> x2\<close> \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close>
+ using \<open>i \<in> ideal_points ?l2\<close>
by (simp add: ideal_points_on_circline)
moreover
have "x \<noteq> inversion x"
- using `x \<in> unit_disc`
+ using \<open>x \<in> unit_disc\<close>
using inversion_noteq_unit_disc by fastforce
moreover
have "x \<noteq> i"
- using `x \<in> unit_disc`
+ using \<open>x \<in> unit_disc\<close>
using \<open>i \<in> circline_set unit_circle\<close> circline_set_def inversion_noteq_unit_disc
by fastforce+
moreover
have "inversion x \<noteq> i"
using \<open>i \<in> circline_set unit_circle\<close> \<open>x \<noteq> i\<close> circline_set_def inversion_unit_circle
by fastforce
ultimately
have "?l1 = ?l2"
using unique_circline_set[of x "inversion x" i]
by blast
hence "0\<^sub>h \<in> circline_set ?l2"
by (metis \<open>a \<noteq> 0\<^sub>h\<close> poincare_line_circline_set(1))
thus False
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}`
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
unfolding poincare_collinear_def
- using \<open>poincare_collinear {x, x1, x2}\<close> \<open>x1 \<noteq> x2\<close> `x1 \<in> unit_disc` `x2 \<in> unit_disc` poincare_collinear_def unique_poincare_line
+ using \<open>poincare_collinear {x, x1, x2}\<close> \<open>x1 \<noteq> x2\<close> \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> poincare_collinear_def unique_poincare_line
by auto
qed
lemma limiting_parallels:
assumes "a \<in> unit_disc" and "x1 \<in> unit_disc" and "x2 \<in> unit_disc" and "\<not> poincare_on_line a x1 x2"
shows "\<exists>a1\<in>unit_disc. \<exists>a2\<in>unit_disc.
\<not> poincare_on_line a a1 a2 \<and>
\<not> poincare_ray_meets_line a a1 x1 x2 \<and> \<not> poincare_ray_meets_line a a2 x1 x2 \<and>
(\<forall>a'\<in>unit_disc. poincare_in_angle a' a1 a a2 \<longrightarrow> poincare_ray_meets_line a a' x1 x2)" (is "?P a x1 x2")
proof-
have "\<not> poincare_collinear {a, x1, x2}"
- using `\<not> poincare_on_line a x1 x2`
+ using \<open>\<not> poincare_on_line a x1 x2\<close>
unfolding poincare_on_line_def
by simp
have "\<forall> x1 x2. x1 \<in> unit_disc \<and> x2 \<in> unit_disc \<and> \<not> poincare_collinear {a, x1, x2} \<longrightarrow> ?P a x1 x2" (is "?Q a")
- proof (rule wlog_zero[OF `a \<in> unit_disc`])
+ proof (rule wlog_zero[OF \<open>a \<in> unit_disc\<close>])
fix a u
assume *: "u \<in> unit_disc" "cmod a < 1"
hence uf: "unit_disc_fix (blaschke a)"
by simp
assume **: "?Q (moebius_pt (blaschke a) u)"
show "?Q u"
proof safe
fix x1 x2
let ?M = "moebius_pt (blaschke a)"
assume xx: "x1 \<in> unit_disc" "x2 \<in> unit_disc" "\<not> poincare_collinear {u, x1, x2}"
hence MM: "?M x1 \<in> unit_disc \<and> ?M x2\<in> unit_disc \<and> \<not> poincare_collinear {?M u, ?M x1, ?M x2}"
using *
by auto
show "?P u x1 x2" (is "\<exists>a1\<in>unit_disc. \<exists>a2\<in>unit_disc. ?P' a1 a2 u x1 x2")
proof-
obtain Ma1 Ma2 where MM: "Ma1 \<in> unit_disc" "Ma2 \<in> unit_disc" "?P' Ma1 Ma2 (?M u) (?M x1) (?M x2)"
using **[rule_format, OF MM]
by blast
hence MM': "\<forall>a'\<in>unit_disc. poincare_in_angle a' Ma1 (?M u) Ma2 \<longrightarrow> poincare_ray_meets_line (?M u) a' (?M x1) (?M x2)"
by auto
obtain a1 a2 where a: "a1 \<in> unit_disc" "a2 \<in> unit_disc" "?M a1 = Ma1" "?M a2 = Ma2"
using uf
by (metis \<open>Ma1 \<in> unit_disc\<close> \<open>Ma2 \<in> unit_disc\<close> image_iff unit_disc_fix_iff)
have "\<forall>a'\<in>unit_disc. poincare_in_angle a' a1 u a2 \<longrightarrow> poincare_ray_meets_line u a' x1 x2"
proof safe
fix a'
assume "a' \<in> unit_disc" "poincare_in_angle a' a1 u a2"
thus "poincare_ray_meets_line u a' x1 x2"
using MM(1-2) MM'[rule_format, of "?M a'"] * uf a xx
by (meson unit_disc_fix_discI unit_disc_fix_preserves_poincare_in_angle unit_disc_fix_preserves_poincare_ray_meets_line)
qed
hence "?P' a1 a2 u x1 x2"
using MM * uf xx a
by auto
thus ?thesis
- using `a1 \<in> unit_disc` `a2 \<in> unit_disc`
+ using \<open>a1 \<in> unit_disc\<close> \<open>a2 \<in> unit_disc\<close>
by blast
qed
qed
next
show "?Q 0\<^sub>h"
proof safe
fix x1 x2
assume "x1 \<in> unit_disc" "x2 \<in> unit_disc"
assume "\<not> poincare_collinear {0\<^sub>h, x1, x2}"
show "?P 0\<^sub>h x1 x2"
proof-
let ?lx = "poincare_line x1 x2"
have "x1 \<noteq> x2"
- using `x1 \<in> unit_disc` `x2 \<in> unit_disc``\<not> poincare_collinear {0\<^sub>h, x1, x2}`
+ using \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close>\<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
using poincare_collinear3_between
by auto
have lx: "is_poincare_line ?lx"
- using is_poincare_line_poincare_line[OF `x1 \<noteq> x2`]
+ using is_poincare_line_poincare_line[OF \<open>x1 \<noteq> x2\<close>]
by simp
obtain i1 i2 where "ideal_points ?lx = {i1, i2}"
by (meson \<open>x1 \<noteq> x2\<close> is_poincare_line_poincare_line obtain_ideal_points)
let ?li = "poincare_line i1 i2"
let ?i1 = "to_complex i1"
let ?i2 = "to_complex i2"
have "i1 \<in> unit_circle_set" "i2 \<in> unit_circle_set"
using lx \<open>ideal_points ?lx = {i1, i2}\<close>
unfolding unit_circle_set_def
by (metis ideal_points_on_unit_circle insertI1, metis ideal_points_on_unit_circle insertI1 insertI2)
have "i1 \<noteq> i2"
using \<open>ideal_points ?lx = {i1, i2}\<close> \<open>x1 \<in> unit_disc\<close> \<open>x1 \<noteq> x2\<close> \<open>x2 \<in> unit_disc\<close> ideal_points_different(1)
by blast
let ?a1 = "of_complex (?i1 / 2)"
let ?a2 = "of_complex (?i2 / 2)"
let ?la = "poincare_line ?a1 ?a2"
have "?a1 \<in> unit_disc" "?a2 \<in> unit_disc"
- using `i1 \<in> unit_circle_set` `i2 \<in> unit_circle_set`
+ using \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close>
unfolding unit_circle_set_def unit_disc_def disc_def circline_set_def
by auto (transfer, transfer, case_tac i1, case_tac i2, simp add: vec_cnj_def)+
have "?a1 \<noteq> 0\<^sub>h" "?a2 \<noteq> 0\<^sub>h"
- using `i1 \<in> unit_circle_set` `i2 \<in> unit_circle_set`
+ using \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close>
unfolding unit_circle_set_def
by auto
have "?a1 \<noteq> ?a2"
- using `i1 \<noteq> i2`
+ using \<open>i1 \<noteq> i2\<close>
by (metis \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close> circline_set_def divide_cancel_right inversion_infty inversion_unit_circle mem_Collect_eq of_complex_to_complex of_complex_zero to_complex_of_complex unit_circle_set_def zero_neq_numeral)
have "poincare_collinear {0\<^sub>h, ?a1, i1}"
unfolding poincare_collinear_def
- using `?a1 \<noteq> 0\<^sub>h`[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a1]
+ using \<open>?a1 \<noteq> 0\<^sub>h\<close>[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a1]
unfolding circline_set_def
apply (rule_tac x="poincare_line 0\<^sub>h ?a1" in exI, auto)
apply (transfer, transfer, auto simp add: vec_cnj_def)
done
have "poincare_collinear {0\<^sub>h, ?a2, i2}"
unfolding poincare_collinear_def
- using `?a2 \<noteq> 0\<^sub>h`[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a2]
+ using \<open>?a2 \<noteq> 0\<^sub>h\<close>[symmetric] is_poincare_line_poincare_line[of "0\<^sub>h" ?a2]
unfolding circline_set_def
apply (rule_tac x="poincare_line 0\<^sub>h ?a2" in exI, auto)
apply (transfer, transfer, auto simp add: vec_cnj_def)
done
have "\<not> poincare_ray_meets_line 0\<^sub>h ?a1 x1 x2"
using tangent_not_meet[of x1 x2 i1 ?a1]
- using `x1 \<in> unit_disc` `x2 \<in> unit_disc` `?a1 \<in> unit_disc` `x1 \<noteq> x2` `\<not> poincare_collinear {0\<^sub>h, x1, x2}`
- using `ideal_points ?lx = {i1, i2}` `?a1 \<noteq> 0\<^sub>h` `poincare_collinear {0\<^sub>h, ?a1, i1}`
+ using \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>?a1 \<in> unit_disc\<close> \<open>x1 \<noteq> x2\<close> \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
+ using \<open>ideal_points ?lx = {i1, i2}\<close> \<open>?a1 \<noteq> 0\<^sub>h\<close> \<open>poincare_collinear {0\<^sub>h, ?a1, i1}\<close>
by simp
moreover
have "\<not> poincare_ray_meets_line 0\<^sub>h ?a2 x1 x2"
using tangent_not_meet[of x1 x2 i2 ?a2]
- using `x1 \<in> unit_disc` `x2 \<in> unit_disc` `?a2 \<in> unit_disc` `x1 \<noteq> x2` `\<not> poincare_collinear {0\<^sub>h, x1, x2}`
- using `ideal_points ?lx = {i1, i2}` `?a2 \<noteq> 0\<^sub>h` `poincare_collinear {0\<^sub>h, ?a2, i2}`
+ using \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> \<open>x1 \<noteq> x2\<close> \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
+ using \<open>ideal_points ?lx = {i1, i2}\<close> \<open>?a2 \<noteq> 0\<^sub>h\<close> \<open>poincare_collinear {0\<^sub>h, ?a2, i2}\<close>
by simp
moreover
have "\<forall>a' \<in> unit_disc. poincare_in_angle a' ?a1 0\<^sub>h ?a2 \<longrightarrow> poincare_ray_meets_line 0\<^sub>h a' x1 x2"
unfolding poincare_in_angle_def
proof safe
fix a' a
assume *: "a' \<in> unit_disc" "a \<in> unit_disc" "poincare_on_ray a' 0\<^sub>h a" "a' \<noteq> 0\<^sub>h"
"poincare_between ?a1 a ?a2" "a \<noteq> ?a1" "a \<noteq> ?a2"
show "poincare_ray_meets_line 0\<^sub>h a' x1 x2"
proof-
have "\<forall> a' a1 a2 x1 x2 i1 i2.
a' \<in> unit_disc \<and> x1 \<in> unit_disc \<and> x2 \<in> unit_disc \<and> x1 \<noteq> x2 \<and>
\<not> poincare_collinear {0\<^sub>h, x1, x2} \<and> ideal_points (poincare_line x1 x2) = {i1, i2} \<and>
a1 = of_complex (to_complex i1 / 2) \<and> a2 = of_complex (to_complex i2 / 2) \<and>
i1 \<noteq> i2 \<and> a1 \<noteq> a2 \<and> poincare_collinear {0\<^sub>h, a1, i1} \<and> poincare_collinear {0\<^sub>h, a2, i2} \<and>
a1 \<in> unit_disc \<and> a2 \<in> unit_disc \<and> i1 \<in> unit_circle_set \<and> i2 \<in> unit_circle_set \<and>
poincare_on_ray a' 0\<^sub>h a \<and> a' \<noteq> 0\<^sub>h \<and> poincare_between a1 a a2 \<and> a \<noteq> a1 \<and> a \<noteq> a2 \<longrightarrow>
poincare_ray_meets_line 0\<^sub>h a' x1 x2" (is "\<forall> a' a1 a2 x1 x2 i1 i2. ?R 0\<^sub>h a' a1 a2 x1 x2 i1 i2 a")
- proof (rule wlog_rotation_to_positive_x_axis[OF `a \<in> unit_disc`])
+ proof (rule wlog_rotation_to_positive_x_axis[OF \<open>a \<in> unit_disc\<close>])
let ?R' = "\<lambda> a zero. \<forall> a' a1 a2 x1 x2 i1 i2. ?R zero a' a1 a2 x1 x2 i1 i2 a"
fix xa
assume xa: "is_real xa" "0 < Re xa" "Re xa < 1"
let ?a = "of_complex xa"
show "?R' ?a 0\<^sub>h"
proof safe
fix a' a1 a2 x1 x2 i1 i2
let ?i1 = "to_complex i1" and ?i2 = "to_complex i2"
let ?a1 = "of_complex (?i1 / 2)" and ?a2 = "of_complex (?i2 / 2)"
let ?la = "poincare_line ?a1 ?a2" and ?lx = "poincare_line x1 x2" and ?li = "poincare_line i1 i2"
assume "a' \<in> unit_disc" "x1 \<in> unit_disc" "x2 \<in> unit_disc" "x1 \<noteq> x2"
assume "\<not> poincare_collinear {0\<^sub>h, x1, x2}" "ideal_points ?lx = {i1, i2}"
assume "poincare_on_ray a' 0\<^sub>h ?a" "a' \<noteq> 0\<^sub>h"
assume "poincare_between ?a1 ?a ?a2" "?a \<noteq> ?a1" "?a \<noteq> ?a2"
assume "i1 \<noteq> i2" "?a1 \<noteq> ?a2" "poincare_collinear {0\<^sub>h, ?a1, i1}" "poincare_collinear {0\<^sub>h, ?a2, i2}"
assume "?a1 \<in> unit_disc" "?a2 \<in> unit_disc"
assume "i1 \<in> unit_circle_set" "i2 \<in> unit_circle_set"
show "poincare_ray_meets_line 0\<^sub>h a' x1 x2"
proof-
have "?lx = ?li"
using \<open>ideal_points ?lx = {i1, i2}\<close> \<open>x1 \<noteq> x2\<close> ideal_points_line_unique
by auto
have lx: "is_poincare_line ?lx"
- using is_poincare_line_poincare_line[OF `x1 \<noteq> x2`]
+ using is_poincare_line_poincare_line[OF \<open>x1 \<noteq> x2\<close>]
by simp
have "x1 \<in> circline_set ?lx" "x2 \<in> circline_set ?lx"
using lx \<open>x1 \<noteq> x2\<close>
by auto
have "?lx \<noteq> x_axis"
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}` `x1 \<in> circline_set ?lx` `x2 \<in> circline_set ?lx` lx
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close> \<open>x1 \<in> circline_set ?lx\<close> \<open>x2 \<in> circline_set ?lx\<close> lx
unfolding poincare_collinear_def
by auto
have "0\<^sub>h \<notin> circline_set ?lx"
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}` lx `x1 \<in> circline_set ?lx` `x2 \<in> circline_set ?lx`
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close> lx \<open>x1 \<in> circline_set ?lx\<close> \<open>x2 \<in> circline_set ?lx\<close>
unfolding poincare_collinear_def
by auto
have "xa \<noteq> 0" "?a \<noteq> 0\<^sub>h"
using xa
by auto
hence "0\<^sub>h \<noteq> ?a"
by metis
have "?a \<in> positive_x_axis"
using xa
unfolding positive_x_axis_def
by simp
have "?a \<in> unit_disc"
using xa
by (auto simp add: cmod_eq_Re)
have "?a \<in> circline_set ?la"
- using `poincare_between ?a1 ?a ?a2`
+ using \<open>poincare_between ?a1 ?a ?a2\<close>
using \<open>?a1 \<noteq> ?a2\<close> \<open>?a \<in> unit_disc\<close> \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> poincare_between_poincare_line_uzv
by blast
have "?a1 \<in> circline_set ?la" "?a2 \<in> circline_set ?la"
by (auto simp add: \<open>?a1 \<noteq> ?a2\<close>)
have la: "is_poincare_line ?la"
- using is_poincare_line_poincare_line[OF `?a1 \<noteq> ?a2`]
+ using is_poincare_line_poincare_line[OF \<open>?a1 \<noteq> ?a2\<close>]
by simp
have inv: "inversion i1 = i1" "inversion i2 = i2"
- using `i1 \<in> unit_circle_set` `i2 \<in> unit_circle_set`
+ using \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close>
by (auto simp add: circline_set_def unit_circle_set_def)
have "i1 \<noteq> \<infinity>\<^sub>h" "i2 \<noteq> \<infinity>\<^sub>h"
using inv
by auto
have "?a1 \<notin> circline_set x_axis \<and> ?a2 \<notin> circline_set x_axis"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "?a1 \<in> circline_set x_axis \<or> ?a2 \<in> circline_set x_axis"
by auto
hence "?la = x_axis"
proof
assume "?a1 \<in> circline_set x_axis"
hence "{?a, ?a1} \<subseteq> circline_set ?la \<inter> circline_set x_axis"
- using `?a \<in> circline_set ?la` `?a1 \<in> circline_set ?la``?a \<in> positive_x_axis`
+ using \<open>?a \<in> circline_set ?la\<close> \<open>?a1 \<in> circline_set ?la\<close>\<open>?a \<in> positive_x_axis\<close>
using circline_set_x_axis_I xa(1)
by blast
thus "?la = x_axis"
using unique_is_poincare_line[of ?a ?a1 ?la x_axis]
- using `?a1 \<in> unit_disc` `?a \<in> unit_disc` la `?a \<noteq> ?a1`
+ using \<open>?a1 \<in> unit_disc\<close> \<open>?a \<in> unit_disc\<close> la \<open>?a \<noteq> ?a1\<close>
by auto
next
assume "?a2 \<in> circline_set x_axis"
hence "{?a, ?a2} \<subseteq> circline_set ?la \<inter> circline_set x_axis"
- using `?a \<in> circline_set ?la` `?a2 \<in> circline_set ?la` `?a \<in> positive_x_axis`
+ using \<open>?a \<in> circline_set ?la\<close> \<open>?a2 \<in> circline_set ?la\<close> \<open>?a \<in> positive_x_axis\<close>
using circline_set_x_axis_I xa(1)
by blast
thus "?la = x_axis"
using unique_is_poincare_line[of ?a ?a2 ?la x_axis]
- using `?a2 \<in> unit_disc` `?a \<in> unit_disc` la `?a \<noteq> ?a2`
+ using \<open>?a2 \<in> unit_disc\<close> \<open>?a \<in> unit_disc\<close> la \<open>?a \<noteq> ?a2\<close>
by auto
qed
hence "i1 \<in> circline_set x_axis \<and> i2 \<in> circline_set x_axis"
- using `?a1 \<in> circline_set ?la` `?a2 \<in> circline_set ?la`
+ using \<open>?a1 \<in> circline_set ?la\<close> \<open>?a2 \<in> circline_set ?la\<close>
by (metis \<open>i1 \<noteq> \<infinity>\<^sub>h\<close> \<open>i2 \<noteq> \<infinity>\<^sub>h\<close> \<open>of_complex (to_complex i1 / 2) \<in> unit_disc\<close> \<open>of_complex (to_complex i2 / 2) \<in> unit_disc\<close> \<open>poincare_collinear {0\<^sub>h, of_complex (to_complex i1 / 2), i1}\<close> \<open>poincare_collinear {0\<^sub>h, of_complex (to_complex i2 / 2), i2}\<close> divide_eq_0_iff inf_not_of_complex inv(1) inv(2) inversion_noteq_unit_disc of_complex_to_complex of_complex_zero_iff poincare_collinear3_poincare_lines_equal_general poincare_line_0_real_is_x_axis poincare_line_circline_set(2) zero_in_unit_disc zero_neq_numeral)
thus False
- using `?lx \<noteq> x_axis` unique_is_poincare_line_general[of i1 i2 ?li x_axis] `i1 \<noteq> i2` inv `?lx = ?li`
+ using \<open>?lx \<noteq> x_axis\<close> unique_is_poincare_line_general[of i1 i2 ?li x_axis] \<open>i1 \<noteq> i2\<close> inv \<open>?lx = ?li\<close>
by auto
qed
hence "?la \<noteq> x_axis"
using \<open>?a1 \<noteq> ?a2\<close> poincare_line_circline_set(1)
by fastforce
have "intersects_x_axis_positive ?la"
- using intersects_x_axis_positive_iff[of ?la] `?la \<noteq> x_axis` `?a \<in> circline_set ?la` la
- using `?a \<in> unit_disc` `?a \<in> positive_x_axis`
+ using intersects_x_axis_positive_iff[of ?la] \<open>?la \<noteq> x_axis\<close> \<open>?a \<in> circline_set ?la\<close> la
+ using \<open>?a \<in> unit_disc\<close> \<open>?a \<in> positive_x_axis\<close>
by auto
have "intersects_x_axis ?lx"
proof-
have "arg (to_complex ?a1) * arg (to_complex ?a2) < 0"
- using `poincare_between ?a1 ?a ?a2` `?a1 \<in> unit_disc` `?a2 \<in> unit_disc`
+ using \<open>poincare_between ?a1 ?a ?a2\<close> \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close>
using poincare_between_x_axis_intersection[of ?a1 ?a2 "of_complex xa"]
- using `?a1 \<noteq> ?a2` `?a \<in> unit_disc` `?a1 \<notin> circline_set x_axis \<and> ?a2 \<notin> circline_set x_axis` `?a \<in> positive_x_axis`
- using `?a \<in> circline_set ?la`
+ using \<open>?a1 \<noteq> ?a2\<close> \<open>?a \<in> unit_disc\<close> \<open>?a1 \<notin> circline_set x_axis \<and> ?a2 \<notin> circline_set x_axis\<close> \<open>?a \<in> positive_x_axis\<close>
+ using \<open>?a \<in> circline_set ?la\<close>
unfolding positive_x_axis_def
by simp
moreover
have "\<And> x y x' y' :: real. \<lbrakk>sgn x' = sgn x; sgn y' = sgn y\<rbrakk> \<Longrightarrow> x*y < 0 \<longleftrightarrow> x'*y' < 0"
by (metis sgn_less sgn_mult)
ultimately
have "Im (to_complex ?a1) * Im (to_complex ?a2) < 0"
using arg_Im_sgn[of "to_complex ?a1"] arg_Im_sgn[of "to_complex ?a2"]
- using `?a1 \<in> unit_disc` `?a2 \<in> unit_disc` `?a1 \<notin> circline_set x_axis \<and> ?a2 \<notin> circline_set x_axis`
+ using \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> \<open>?a1 \<notin> circline_set x_axis \<and> ?a2 \<notin> circline_set x_axis\<close>
using inf_or_of_complex[of ?a1] inf_or_of_complex[of ?a2] circline_set_x_axis
by (metis circline_set_x_axis_I to_complex_of_complex)
thus ?thesis
using ideal_points_intersects_x_axis[of ?lx i1 i2]
- using `ideal_points ?lx = {i1, i2}` lx `?lx \<noteq> x_axis`
+ using \<open>ideal_points ?lx = {i1, i2}\<close> lx \<open>?lx \<noteq> x_axis\<close>
by simp
qed
have "intersects_x_axis_positive ?lx"
proof-
have "cmod ?i1 = 1" "cmod ?i2 = 1"
using \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close>
unfolding unit_circle_set_def
by auto
let ?a1' = "?i1 / 2" and ?a2' = "?i2 / 2"
let ?Aa1 = "\<i> * (?a1' * cnj ?a2' - ?a2' * cnj ?a1')" and
?Ba1 = "\<i> * (?a2' * cor ((cmod ?a1')\<^sup>2 + 1) - ?a1' * cor ((cmod ?a2')\<^sup>2 + 1))"
have "?Aa1 \<noteq> 0 \<or> ?Ba1 \<noteq> 0"
- using `cmod (to_complex i1) = 1` `cmod (to_complex i2) = 1` `?a1 \<noteq> ?a2`
+ using \<open>cmod (to_complex i1) = 1\<close> \<open>cmod (to_complex i2) = 1\<close> \<open>?a1 \<noteq> ?a2\<close>
by (auto simp add: power_divide complex_mult_cnj_cmod)
have "is_real ?Aa1"
by simp
have "?a1 \<noteq> inversion ?a2"
using \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> inversion_noteq_unit_disc by fastforce
hence "Re ?Ba1 / Re ?Aa1 < -1"
- using `intersects_x_axis_positive ?la` `?a1 \<noteq> ?a2`
- using intersects_x_axis_positive_mk_circline[of ?Aa1 ?Ba1] `?Aa1 \<noteq> 0 \<or> ?Ba1 \<noteq> 0` `is_real ?Aa1`
+ using \<open>intersects_x_axis_positive ?la\<close> \<open>?a1 \<noteq> ?a2\<close>
+ using intersects_x_axis_positive_mk_circline[of ?Aa1 ?Ba1] \<open>?Aa1 \<noteq> 0 \<or> ?Ba1 \<noteq> 0\<close> \<open>is_real ?Aa1\<close>
using poincare_line_non_homogenous[of ?a1 ?a2]
by (simp add: Let_def)
moreover
let ?i1' = "to_complex i1" and ?i2' = "to_complex i2"
let ?Ai1 = "\<i> * (?i1' * cnj ?i2' - ?i2' * cnj ?i1')" and
?Bi1 = "\<i> * (?i2' * cor ((cmod ?i1')\<^sup>2 + 1) - ?i1' * cor ((cmod ?i2')\<^sup>2 + 1))"
have "?Ai1 \<noteq> 0 \<or> ?Bi1 \<noteq> 0"
- using `cmod (to_complex i1) = 1` `cmod (to_complex i2) = 1` `?a1 \<noteq> ?a2`
+ using \<open>cmod (to_complex i1) = 1\<close> \<open>cmod (to_complex i2) = 1\<close> \<open>?a1 \<noteq> ?a2\<close>
by (auto simp add: power_divide complex_mult_cnj_cmod)
have "is_real ?Ai1"
by simp
have "sgn (Re ?Bi1 / Re ?Ai1) = sgn (Re ?Ba1 / Re ?Aa1)"
proof-
have "Re ?Bi1 / Re ?Ai1 = (Im ?i1 * 2 - Im ?i2 * 2) /
(Im ?i2 * (Re ?i1 * 2) - Im ?i1 * (Re ?i2 * 2))"
- using `cmod ?i1 = 1` `cmod ?i2 = 1`
+ using \<open>cmod ?i1 = 1\<close> \<open>cmod ?i2 = 1\<close>
by (auto simp add: complex_mult_cnj_cmod field_simps)
also have "... = (Im ?i1 - Im ?i2) /
(Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2))" (is "... = ?expr")
apply (subst left_diff_distrib[symmetric])
apply (subst semiring_normalization_rules(18))+
apply (subst left_diff_distrib[symmetric])
by (metis mult.commute mult_divide_mult_cancel_left_if zero_neq_numeral)
finally have 1: "Re ?Bi1 / Re ?Ai1 = (Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2))"
.
have "Re ?Ba1 / Re ?Aa1 = (Im ?i1 * 20 - Im ?i2 * 20) /
(Im ?i2 * (Re ?i1 * 16) - Im ?i1 * (Re ?i2 * 16))"
- using `cmod (to_complex i1) = 1` `cmod (to_complex i2) = 1`
+ using \<open>cmod (to_complex i1) = 1\<close> \<open>cmod (to_complex i2) = 1\<close>
by (auto simp add: complex_mult_cnj_cmod field_simps)
also have "... = (20 / 16) * ((Im ?i1 - Im ?i2) /
(Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2)))"
apply (subst left_diff_distrib[symmetric])+
apply (subst semiring_normalization_rules(18))+
apply (subst left_diff_distrib[symmetric])+
by (metis (no_types, hide_lams) field_class.field_divide_inverse mult.commute times_divide_times_eq)
finally have 2: "Re ?Ba1 / Re ?Aa1 = (5 / 4) * ((Im ?i1 - Im ?i2) / (Im ?i2 * (Re ?i1) - Im ?i1 * (Re ?i2)))"
by simp
have "?expr \<noteq> 0"
- using `Re ?Ba1 / Re ?Aa1 < -1`
+ using \<open>Re ?Ba1 / Re ?Aa1 < -1\<close>
apply (subst (asm) 2)
by linarith
thus ?thesis
apply (subst 1, subst 2)
apply (simp only: sgn_mult)
by simp
qed
moreover
have "i1 \<noteq> inversion i2"
by (simp add: \<open>i1 \<noteq> i2\<close> inv(2))
have "(Re ?Bi1 / Re ?Ai1)\<^sup>2 > 1"
proof-
have "?Ai1 = 0 \<or> (Re ?Bi1)\<^sup>2 > (Re ?Ai1)\<^sup>2"
- using `intersects_x_axis ?lx`
- using `i1 \<noteq> i2` `i1 \<noteq> \<infinity>\<^sub>h` `i2 \<noteq> \<infinity>\<^sub>h` `i1 \<noteq> inversion i2`
- using intersects_x_axis_mk_circline[of ?Ai1 ?Bi1] `?Ai1 \<noteq> 0 \<or> ?Bi1 \<noteq> 0` `is_real ?Ai1`
- using poincare_line_non_homogenous[of i1 i2] `?lx = ?li`
+ using \<open>intersects_x_axis ?lx\<close>
+ using \<open>i1 \<noteq> i2\<close> \<open>i1 \<noteq> \<infinity>\<^sub>h\<close> \<open>i2 \<noteq> \<infinity>\<^sub>h\<close> \<open>i1 \<noteq> inversion i2\<close>
+ using intersects_x_axis_mk_circline[of ?Ai1 ?Bi1] \<open>?Ai1 \<noteq> 0 \<or> ?Bi1 \<noteq> 0\<close> \<open>is_real ?Ai1\<close>
+ using poincare_line_non_homogenous[of i1 i2] \<open>?lx = ?li\<close>
by metis
moreover
have "?Ai1 \<noteq> 0"
proof (rule ccontr)
assume "\<not> ?thesis"
hence "0\<^sub>h \<in> circline_set ?li"
unfolding circline_set_def
apply simp
apply (transfer, transfer, case_tac i1, case_tac i2)
by (auto simp add: vec_cnj_def field_simps)
thus False
- using `0\<^sub>h \<notin> circline_set ?lx` `?lx = ?li`
+ using \<open>0\<^sub>h \<notin> circline_set ?lx\<close> \<open>?lx = ?li\<close>
by simp
qed
ultimately
have "(Re ?Bi1)\<^sup>2 > (Re ?Ai1)\<^sup>2"
by auto
moreover
have "Re ?Ai1 \<noteq> 0"
- using `is_real ?Ai1` `?Ai1 \<noteq> 0`
+ using \<open>is_real ?Ai1\<close> \<open>?Ai1 \<noteq> 0\<close>
by (simp add: complex_eq_iff)
ultimately
show ?thesis
by (simp add: power_divide)
qed
moreover
{
fix x1 x2 :: real
assume "sgn x1 = sgn x2" "x1 < -1" "x2\<^sup>2 > 1"
hence "x2 < -1"
by (smt one_power2 real_sqrt_abs real_sqrt_less_iff sgn_neg sgn_pos)
}
ultimately
have "Re ?Bi1 / Re ?Ai1 < -1"
by metis
thus ?thesis
- using `i1 \<noteq> i2` `i1 \<noteq> \<infinity>\<^sub>h` `i2 \<noteq> \<infinity>\<^sub>h` `i1 \<noteq> inversion i2`
- using intersects_x_axis_positive_mk_circline[of ?Ai1 ?Bi1] `?Ai1 \<noteq> 0 \<or> ?Bi1 \<noteq> 0` `is_real ?Ai1`
- using poincare_line_non_homogenous[of i1 i2] `?lx = ?li`
+ using \<open>i1 \<noteq> i2\<close> \<open>i1 \<noteq> \<infinity>\<^sub>h\<close> \<open>i2 \<noteq> \<infinity>\<^sub>h\<close> \<open>i1 \<noteq> inversion i2\<close>
+ using intersects_x_axis_positive_mk_circline[of ?Ai1 ?Bi1] \<open>?Ai1 \<noteq> 0 \<or> ?Bi1 \<noteq> 0\<close> \<open>is_real ?Ai1\<close>
+ using poincare_line_non_homogenous[of i1 i2] \<open>?lx = ?li\<close>
by (simp add: Let_def)
qed
then obtain x where x: "x \<in> unit_disc" "x \<in> circline_set ?lx \<inter> positive_x_axis"
- using intersects_x_axis_positive_iff[OF lx `?lx \<noteq> x_axis`]
+ using intersects_x_axis_positive_iff[OF lx \<open>?lx \<noteq> x_axis\<close>]
by auto
have "poincare_on_ray x 0\<^sub>h a' \<and> poincare_collinear {x1, x2, x}"
proof
show "poincare_collinear {x1, x2, x}"
- using x lx `x1 \<in> circline_set ?lx` `x2 \<in> circline_set ?lx`
+ using x lx \<open>x1 \<in> circline_set ?lx\<close> \<open>x2 \<in> circline_set ?lx\<close>
unfolding poincare_collinear_def
by auto
next
show "poincare_on_ray x 0\<^sub>h a'"
unfolding poincare_on_ray_def
proof-
have "a' \<in> circline_set x_axis"
- using `poincare_on_ray a' 0\<^sub>h ?a` xa `0\<^sub>h \<noteq> ?a` `xa \<noteq> 0` `a' \<in> unit_disc`
+ using \<open>poincare_on_ray a' 0\<^sub>h ?a\<close> xa \<open>0\<^sub>h \<noteq> ?a\<close> \<open>xa \<noteq> 0\<close> \<open>a' \<in> unit_disc\<close>
unfolding poincare_on_ray_def
using poincare_line_0_real_is_x_axis[of "of_complex xa"]
using poincare_between_poincare_line_uvz[of "0\<^sub>h" "of_complex xa" a']
using poincare_between_poincare_line_uzv[of "0\<^sub>h" "of_complex xa" a']
by (auto simp add: cmod_eq_Re)
then obtain xa' where xa': "a' = of_complex xa'" "is_real xa'"
- using `a' \<in> unit_disc`
+ using \<open>a' \<in> unit_disc\<close>
using circline_set_def on_circline_x_axis
by auto
hence "-1 < Re xa'" "Re xa' < 1" "xa' \<noteq> 0"
- using `a' \<in> unit_disc` `a' \<noteq> 0\<^sub>h`
+ using \<open>a' \<in> unit_disc\<close> \<open>a' \<noteq> 0\<^sub>h\<close>
by (auto simp add: cmod_eq_Re)
hence "Re xa' > 0" "Re xa' < 1" "is_real xa'"
- using `poincare_on_ray a' 0\<^sub>h (of_complex xa)`
+ using \<open>poincare_on_ray a' 0\<^sub>h (of_complex xa)\<close>
using poincare_between_x_axis_0uv[of "Re xa'" "Re xa"]
using poincare_between_x_axis_0uv[of "Re xa" "Re xa'"]
using circline_set_positive_x_axis_I[of "Re xa'"]
using xa xa' complex_of_real_Re
unfolding poincare_on_ray_def
by (smt of_real_0, linarith, blast)
moreover
obtain xx where "is_real xx" "Re xx > 0" "Re xx < 1" "x = of_complex xx"
using x
unfolding positive_x_axis_def
using circline_set_def cmod_eq_Re on_circline_x_axis
by auto
ultimately
show "poincare_between 0\<^sub>h x a' \<or> poincare_between 0\<^sub>h a' x"
- using `a' = of_complex xa'`
+ using \<open>a' = of_complex xa'\<close>
by (smt \<open>a' \<in> unit_disc\<close> arg_0_iff poincare_between_0uv poincare_between_def to_complex_of_complex x(1))
qed
qed
thus ?thesis
- using `x \<in> unit_disc`
+ using \<open>x \<in> unit_disc\<close>
unfolding poincare_ray_meets_line_def poincare_on_line_def
by (metis insert_commute)
qed
qed
next
show "a \<noteq> 0\<^sub>h"
proof (rule ccontr)
assume "\<not> ?thesis"
then obtain k where "k<0" "to_complex ?a1 = cor k * to_complex ?a2"
- using poincare_between_u0v[OF `?a1 \<in> unit_disc` `?a2 \<in> unit_disc` `?a1 \<noteq> 0\<^sub>h` `?a2 \<noteq> 0\<^sub>h`]
- using `poincare_between ?a1 a ?a2`
+ using poincare_between_u0v[OF \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> \<open>?a1 \<noteq> 0\<^sub>h\<close> \<open>?a2 \<noteq> 0\<^sub>h\<close>]
+ using \<open>poincare_between ?a1 a ?a2\<close>
by auto
hence "to_complex i1 = cor k * to_complex i2" "k < 0"
by auto
hence "0\<^sub>h \<in> circline_set (poincare_line x1 x2)"
- using ideal_points_proportional[of "poincare_line x1 x2" i1 i2 k] `ideal_points (poincare_line x1 x2) = {i1, i2}`
- using is_poincare_line_poincare_line[OF `x1 \<noteq> x2`]
+ using ideal_points_proportional[of "poincare_line x1 x2" i1 i2 k] \<open>ideal_points (poincare_line x1 x2) = {i1, i2}\<close>
+ using is_poincare_line_poincare_line[OF \<open>x1 \<noteq> x2\<close>]
by simp
thus False
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}`
- using is_poincare_line_poincare_line[OF `x1 \<noteq> x2`]
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
+ using is_poincare_line_poincare_line[OF \<open>x1 \<noteq> x2\<close>]
unfolding poincare_collinear_def
by (meson \<open>x1 \<noteq> x2\<close> empty_subsetI insert_subset poincare_line_circline_set(1) poincare_line_circline_set(2))
qed
next
fix \<phi> u
let ?R' = "\<lambda> a zero. \<forall> a' a1 a2 x1 x2 i1 i2. ?R zero a' a1 a2 x1 x2 i1 i2 a"
let ?M = "moebius_pt (moebius_rotation \<phi>)"
assume *: "u \<in> unit_disc" "u \<noteq> 0\<^sub>h" and **: "?R' (?M u) 0\<^sub>h"
have uf: "unit_disc_fix (moebius_rotation \<phi>)"
by simp
have "?M 0\<^sub>h = 0\<^sub>h"
by auto
hence **: "?R' (?M u) (?M 0\<^sub>h)"
using **
by simp
show "?R' u 0\<^sub>h"
proof (rule allI)+
fix a' a1 a2 x1 x2 i1 i2
have i1: "i1 \<in> unit_circle_set \<longrightarrow> moebius_pt (moebius_rotation \<phi>) (of_complex (to_complex i1 / 2)) = of_complex (to_complex (moebius_pt (moebius_rotation \<phi>) i1) / 2)"
using unit_circle_set_def by force
have i2: "i2 \<in> unit_circle_set \<longrightarrow> moebius_pt (moebius_rotation \<phi>) (of_complex (to_complex i2 / 2)) = of_complex (to_complex (moebius_pt (moebius_rotation \<phi>) i2) / 2)"
using unit_circle_set_def by force
show "?R 0\<^sub>h a' a1 a2 x1 x2 i1 i2 u"
using **[rule_format, of "?M a'" "?M x1" "?M x2" "?M i1" "?M i2" "?M a1" "?M a2"] uf *
apply (auto simp del: moebius_pt_moebius_rotation_zero moebius_pt_moebius_rotation)
using i1 i2
by simp
qed
qed
thus ?thesis
- using `a' \<in> unit_disc` `x1 \<in> unit_disc` `x2 \<in> unit_disc` `x1 \<noteq> x2`
- using `\<not> poincare_collinear {0\<^sub>h, x1, x2}` `ideal_points ?lx = {i1, i2}` `i1 \<noteq> i2`
- using `?a1 \<noteq> ?a2` `poincare_collinear {0\<^sub>h, ?a1, i1}` `poincare_collinear {0\<^sub>h, ?a2, i2}`
- using `?a1 \<in> unit_disc` `?a2 \<in> unit_disc` `i1 \<in> unit_circle_set` `i2 \<in> unit_circle_set`
- using `poincare_on_ray a' 0\<^sub>h a` `a' \<noteq> 0\<^sub>h` `poincare_between ?a1 a ?a2` `a \<noteq> ?a1` `a \<noteq> ?a2`
+ using \<open>a' \<in> unit_disc\<close> \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>x1 \<noteq> x2\<close>
+ using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close> \<open>ideal_points ?lx = {i1, i2}\<close> \<open>i1 \<noteq> i2\<close>
+ using \<open>?a1 \<noteq> ?a2\<close> \<open>poincare_collinear {0\<^sub>h, ?a1, i1}\<close> \<open>poincare_collinear {0\<^sub>h, ?a2, i2}\<close>
+ using \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close> \<open>i1 \<in> unit_circle_set\<close> \<open>i2 \<in> unit_circle_set\<close>
+ using \<open>poincare_on_ray a' 0\<^sub>h a\<close> \<open>a' \<noteq> 0\<^sub>h\<close> \<open>poincare_between ?a1 a ?a2\<close> \<open>a \<noteq> ?a1\<close> \<open>a \<noteq> ?a2\<close>
by blast
qed
qed
moreover
have "\<not> poincare_on_line 0\<^sub>h ?a1 ?a2"
proof
assume *: "poincare_on_line 0\<^sub>h ?a1 ?a2"
hence "poincare_collinear {0\<^sub>h, ?a1, ?a2}"
unfolding poincare_on_line_def
by simp
hence "poincare_line 0\<^sub>h ?a1 = poincare_line 0\<^sub>h ?a2"
using poincare_collinear3_poincare_lines_equal_general[of "0\<^sub>h" ?a1 ?a2]
using \<open>?a1 \<in> unit_disc\<close> \<open>?a1 \<noteq> 0\<^sub>h\<close> \<open>?a2 \<in> unit_disc\<close> \<open>?a2 \<noteq> 0\<^sub>h\<close>
by (metis inversion_noteq_unit_disc zero_in_unit_disc)
have "i1 \<in> circline_set (poincare_line 0\<^sub>h ?a1)"
- using `poincare_collinear {0\<^sub>h, ?a1, i1}`
+ using \<open>poincare_collinear {0\<^sub>h, ?a1, i1}\<close>
using poincare_collinear3_poincare_line_general[of i1 "0\<^sub>h" ?a1]
- using \<open>?a1 \<in> unit_disc\<close> `?a1 \<noteq> 0\<^sub>h`
+ using \<open>?a1 \<in> unit_disc\<close> \<open>?a1 \<noteq> 0\<^sub>h\<close>
by (metis insert_commute inversion_noteq_unit_disc zero_in_unit_disc)
moreover
have "i2 \<in> circline_set (poincare_line 0\<^sub>h ?a1)"
- using `poincare_collinear {0\<^sub>h, ?a2, i2}`
+ using \<open>poincare_collinear {0\<^sub>h, ?a2, i2}\<close>
using poincare_collinear3_poincare_line_general[of i2 "0\<^sub>h" ?a2]
- using \<open>?a2 \<in> unit_disc\<close> `?a2 \<noteq> 0\<^sub>h` \<open>poincare_line 0\<^sub>h ?a1 = poincare_line 0\<^sub>h ?a2\<close>
+ using \<open>?a2 \<in> unit_disc\<close> \<open>?a2 \<noteq> 0\<^sub>h\<close> \<open>poincare_line 0\<^sub>h ?a1 = poincare_line 0\<^sub>h ?a2\<close>
by (metis insert_commute inversion_noteq_unit_disc zero_in_unit_disc)
ultimately
have "poincare_collinear {0\<^sub>h, i1, i2}"
using \<open>?a1 \<in> unit_disc\<close> \<open>?a1 \<noteq> 0\<^sub>h\<close> \<open>poincare_collinear {0\<^sub>h, ?a1, i1}\<close>
by (smt insert_subset poincare_collinear_def unique_poincare_line zero_in_unit_disc)
hence "0\<^sub>h \<in> circline_set (poincare_line i1 i2)"
using poincare_collinear3_poincare_line_general[of "0\<^sub>h" i1 i2]
using \<open>i1 \<noteq> i2\<close> \<open>i2 \<in> unit_circle_set\<close> unit_circle_set_def
by force
moreover
have "?lx = ?li"
using \<open>ideal_points ?lx = {i1, i2}\<close> \<open>x1 \<noteq> x2\<close> ideal_points_line_unique
by auto
ultimately
show False
using \<open>\<not> poincare_collinear {0\<^sub>h, x1, x2}\<close>
using \<open>x1 \<noteq> x2\<close> poincare_line_poincare_collinear3_general
by auto
qed
ultimately
show ?thesis
- using `?a1 \<in> unit_disc` `?a2 \<in> unit_disc`
+ using \<open>?a1 \<in> unit_disc\<close> \<open>?a2 \<in> unit_disc\<close>
by blast
qed
qed
qed
thus ?thesis
- using `x1 \<in> unit_disc` `x2 \<in> unit_disc` `\<not> poincare_collinear {a, x1, x2}`
+ using \<open>x1 \<in> unit_disc\<close> \<open>x2 \<in> unit_disc\<close> \<open>\<not> poincare_collinear {a, x1, x2}\<close>
by blast
qed
subsection\<open>Interpretation of locales\<close>
global_interpretation PoincareTarskiAbsolute: TarskiAbsolute where cong = p_congruent and betw = p_between
defines p_on_line = PoincareTarskiAbsolute.on_line and
p_on_ray = PoincareTarskiAbsolute.on_ray and
p_in_angle = PoincareTarskiAbsolute.in_angle and
p_ray_meets_line = PoincareTarskiAbsolute.ray_meets_line
proof-
show "TarskiAbsolute p_congruent p_between"
proof
text\<open> 1. Reflexivity of congruence \<close>
fix x y
show "p_congruent x y y x"
unfolding p_congruent_def
by transfer (simp add: poincare_distance_sym)
next
text\<open> 2. Transitivity of congruence \<close>
fix x y z u v w
show "p_congruent x y z u \<and> p_congruent x y v w \<longrightarrow> p_congruent z u v w"
by (transfer, simp)
next
text\<open> 3. Identity of congruence \<close>
fix x y z
show "p_congruent x y z z \<longrightarrow> x = y"
unfolding p_congruent_def
by transfer (simp add: poincare_distance_eq_0_iff)
next
text\<open> 4. Segment construction \<close>
fix x y a b
show "\<exists> z. p_between x y z \<and> p_congruent y z a b"
using segment_construction
unfolding p_congruent_def
by transfer (simp, blast)
next
text\<open> 5. Five segment \<close>
fix x y z x' y' z' u u'
show "x \<noteq> y \<and> p_between x y z \<and> p_between x' y' z' \<and>
p_congruent x y x' y' \<and> p_congruent y z y' z' \<and>
p_congruent x u x' u' \<and> p_congruent y u y' u' \<longrightarrow>
p_congruent z u z' u'"
unfolding p_congruent_def
apply transfer
using five_segment_axiom
by meson
next
text\<open> 6. Identity of betweeness \<close>
fix x y
show "p_between x y x \<longrightarrow> x = y"
by transfer (simp add: poincare_between_sum_distances poincare_distance_eq_0_iff poincare_distance_sym)
next
text\<open> 7. Pasch \<close>
fix x y z u v
show "p_between x u z \<and> p_between y v z \<longrightarrow> (\<exists> a. p_between u a y \<and> p_between x a v)"
apply transfer
using Pasch
by blast
next
text\<open> 8. Lower dimension \<close>
show "\<exists> a. \<exists> b. \<exists> c. \<not> p_between a b c \<and> \<not> p_between b c a \<and> \<not> p_between c a b"
apply (transfer)
using lower_dimension_axiom
by simp
next
text\<open> 9. Upper dimension \<close>
fix x y z u v
show "p_congruent x u x v \<and> p_congruent y u y v \<and> p_congruent z u z v \<and> u \<noteq> v \<longrightarrow>
p_between x y z \<or> p_between y z x \<or> p_between z x y"
unfolding p_congruent_def
by (transfer, simp add: upper_dimension_axiom)
qed
qed
interpretation PoincareTarskiHyperbolic: TarskiHyperbolic
where cong = p_congruent and betw = p_between
proof
text\<open> 10. Euclid negation \<close>
show "\<exists> a b c d t. p_between a d t \<and> p_between b d c \<and> a \<noteq> d \<and>
(\<forall> x y. p_between a b x \<and> p_between a c y \<longrightarrow> \<not> p_between x t y)"
using negated_euclidean_axiom
by transfer (auto, blast)
next
fix a x1 x2
assume "\<not> TarskiAbsolute.on_line p_between a x1 x2"
hence "\<not> p_on_line a x1 x2"
using TarskiAbsolute.on_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
using PoincareTarskiAbsolute.on_line_def
by simp
text\<open> 11. Limiting parallels \<close>
thus "\<exists>a1 a2.
\<not> TarskiAbsolute.on_line p_between a a1 a2 \<and>
\<not> TarskiAbsolute.ray_meets_line p_between a a1 x1 x2 \<and>
\<not> TarskiAbsolute.ray_meets_line p_between a a2 x1 x2 \<and>
(\<forall>a'. TarskiAbsolute.in_angle p_between a' a1 a a2 \<longrightarrow> TarskiAbsolute.ray_meets_line p_between a a' x1 x2)"
unfolding TarskiAbsolute.in_angle_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
unfolding TarskiAbsolute.on_ray_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
unfolding TarskiAbsolute.ray_meets_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
unfolding TarskiAbsolute.on_ray_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
unfolding TarskiAbsolute.on_line_def[OF PoincareTarskiAbsolute.TarskiAbsolute_axioms]
unfolding PoincareTarskiAbsolute.on_line_def
apply transfer
proof-
fix a x1 x2
assume *: "a \<in> unit_disc" "x1 \<in> unit_disc" "x2 \<in> unit_disc"
"\<not> (poincare_between a x1 x2 \<or> poincare_between x1 a x2 \<or> poincare_between x1 x2 a)"
hence "\<not> poincare_on_line a x1 x2"
using poincare_collinear3_iff[of a x1 x2]
using poincare_between_rev poincare_on_line_def by blast
hence "\<exists>a1\<in>unit_disc.
\<exists>a2\<in>unit_disc.
\<not> poincare_on_line a a1 a2 \<and>
\<not> poincare_ray_meets_line a a1 x1 x2 \<and>
\<not> poincare_ray_meets_line a a2 x1 x2 \<and>
(\<forall>a'\<in>unit_disc.
poincare_in_angle a' a1 a a2 \<longrightarrow>
poincare_ray_meets_line a a' x1 x2)"
using limiting_parallels[of a x1 x2] *
by blast
then obtain a1 a2 where **: "a1\<in>unit_disc" "a2\<in>unit_disc" "\<not> poincare_on_line a a1 a2"
"\<not> poincare_ray_meets_line a a2 x1 x2"
"\<not> poincare_ray_meets_line a a1 x1 x2"
"\<forall>a'\<in>unit_disc.
poincare_in_angle a' a1 a a2 \<longrightarrow>
poincare_ray_meets_line a a' x1 x2"
by blast
have "\<not> (\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a1 \<or>
poincare_between a a1 x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x))"
- using `\<not> poincare_ray_meets_line a a1 x1 x2`
+ using \<open>\<not> poincare_ray_meets_line a a1 x1 x2\<close>
unfolding poincare_on_line_def poincare_ray_meets_line_def poincare_on_ray_def
using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3)
by auto
moreover
have "\<not> (\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a2 \<or>
poincare_between a a2 x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x))"
- using `\<not> poincare_ray_meets_line a a2 x1 x2`
+ using \<open>\<not> poincare_ray_meets_line a a2 x1 x2\<close>
unfolding poincare_on_line_def poincare_ray_meets_line_def poincare_on_ray_def
using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3)
by auto
moreover
have "\<not> (poincare_between a a1 a2 \<or> poincare_between a1 a a2 \<or> poincare_between a1 a2 a)"
- using `\<not> poincare_on_line a a1 a2` poincare_collinear3_iff[of a a1 a2]
+ using \<open>\<not> poincare_on_line a a1 a2\<close> poincare_collinear3_iff[of a a1 a2]
using *(1) **(1-2)
unfolding poincare_on_line_def
by simp
moreover
have "(\<forall>a'\<in>{z. z \<in> unit_disc}.
a \<noteq> a1 \<and>
a \<noteq> a2 \<and>
a' \<noteq> a \<and>
(\<exists>x\<in>{z. z \<in> unit_disc}.
poincare_between a1 x a2 \<and>
x \<noteq> a1 \<and>
x \<noteq> a2 \<and>
(poincare_between a a' x \<or>
poincare_between a x a')) \<longrightarrow>
(\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a' \<or>
poincare_between a a' x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x)))"
using **(6)
unfolding poincare_on_line_def poincare_in_angle_def poincare_ray_meets_line_def poincare_on_ray_def
using poincare_collinear3_iff[of _ x1 x2] poincare_between_rev *(2, 3)
by auto
ultimately
show "\<exists>a1\<in>{z. z \<in> unit_disc}.
\<exists>a2\<in>{z. z \<in> unit_disc}.
\<not> (poincare_between a a1 a2 \<or> poincare_between a1 a a2 \<or> poincare_between a1 a2 a) \<and>
\<not> (\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a1 \<or>
poincare_between a a1 x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x)) \<and>
\<not> (\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a2 \<or>
poincare_between a a2 x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x)) \<and>
(\<forall>a'\<in>{z. z \<in> unit_disc}.
a \<noteq> a1 \<and>
a \<noteq> a2 \<and>
a' \<noteq> a \<and>
(\<exists>x\<in>{z. z \<in> unit_disc}.
poincare_between a1 x a2 \<and>
x \<noteq> a1 \<and>
x \<noteq> a2 \<and>
(poincare_between a a' x \<or>
poincare_between a x a')) \<longrightarrow>
(\<exists>x\<in>{z. z \<in> unit_disc}.
(poincare_between a x a' \<or>
poincare_between a a' x) \<and>
(poincare_between x x1 x2 \<or>
poincare_between x1 x x2 \<or>
poincare_between x1 x2 x)))"
using **(1, 2)
by auto
qed
qed
interpretation PoincareElementaryTarskiHyperbolic: ElementaryTarskiHyperbolic p_congruent p_between
proof
text\<open> 12. Continuity \<close>
fix \<phi> \<psi>
assume "\<exists> a. \<forall> x. \<forall> y. \<phi> x \<and> \<psi> y \<longrightarrow> p_between a x y"
thus "\<exists> b. \<forall> x. \<forall> y. \<phi> x \<and> \<psi> y \<longrightarrow> p_between x b y"
apply transfer
using continuity
by auto
qed
end
diff --git a/thys/QHLProver/Complex_Matrix.thy b/thys/QHLProver/Complex_Matrix.thy
--- a/thys/QHLProver/Complex_Matrix.thy
+++ b/thys/QHLProver/Complex_Matrix.thy
@@ -1,2346 +1,2346 @@
section \<open>Complex matrices\<close>
theory Complex_Matrix
imports
"Jordan_Normal_Form.Matrix"
"Jordan_Normal_Form.Conjugate"
"Jordan_Normal_Form.Jordan_Normal_Form_Existence"
begin
subsection \<open>Trace of a matrix\<close>
definition trace :: "'a::ring mat \<Rightarrow> 'a" where
"trace A = (\<Sum> i \<in> {0 ..< dim_row A}. A $$ (i,i))"
lemma trace_zero [simp]:
"trace (0\<^sub>m n n) = 0"
by (simp add: trace_def)
lemma trace_id [simp]:
"trace (1\<^sub>m n) = n"
by (simp add: trace_def)
lemma trace_comm:
fixes A B :: "'a::comm_ring mat"
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n"
shows "trace (A * B) = trace (B * A)"
proof (simp add: trace_def)
have "(\<Sum>i = 0..<n. (A * B) $$ (i, i)) = (\<Sum>i = 0..<n. \<Sum>j = 0..<n. A $$ (i,j) * B $$ (j,i))"
apply (rule sum.cong) using assms by (auto simp add: scalar_prod_def)
also have "\<dots> = (\<Sum>j = 0..<n. \<Sum>i = 0..<n. A $$ (i,j) * B $$ (j,i))"
by (rule sum.swap)
also have "\<dots> = (\<Sum>j = 0..<n. col A j \<bullet> row B j)"
by (metis (no_types, lifting) A B atLeastLessThan_iff carrier_matD index_col index_row scalar_prod_def sum.cong)
also have "\<dots> = (\<Sum>j = 0..<n. row B j \<bullet> col A j)"
apply (rule sum.cong) apply auto
apply (subst comm_scalar_prod[where n=n]) apply auto
using assms by auto
also have "\<dots> = (\<Sum>j = 0..<n. (B * A) $$ (j, j))"
apply (rule sum.cong) using assms by auto
finally show "(\<Sum>i = 0..<dim_row A. (A * B) $$ (i, i)) = (\<Sum>i = 0..<dim_row B. (B * A) $$ (i, i))"
using A B by auto
qed
lemma trace_add_linear:
fixes A B :: "'a::comm_ring mat"
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n"
shows "trace (A + B) = trace A + trace B" (is "?lhs = ?rhs")
proof -
have "?lhs = (\<Sum>i=0..<n. A$$(i, i) + B$$(i, i))" unfolding trace_def using A B by auto
also have "\<dots> = (\<Sum>i=0..<n. A$$(i, i)) + (\<Sum>i=0..<n. B$$(i, i))" by (auto simp add: sum.distrib)
finally have l: "?lhs = (\<Sum>i=0..<n. A$$(i, i)) + (\<Sum>i=0..<n. B$$(i, i))".
have r: "?rhs = (\<Sum>i=0..<n. A$$(i, i)) + (\<Sum>i=0..<n. B$$(i, i))" unfolding trace_def using A B by auto
from l r show ?thesis by auto
qed
lemma trace_minus_linear:
fixes A B :: "'a::comm_ring mat"
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n"
shows "trace (A - B) = trace A - trace B" (is "?lhs = ?rhs")
proof -
have "?lhs = (\<Sum>i=0..<n. A$$(i, i) - B$$(i, i))" unfolding trace_def using A B by auto
also have "\<dots> = (\<Sum>i=0..<n. A$$(i, i)) - (\<Sum>i=0..<n. B$$(i, i))" by (auto simp add: sum_subtractf)
finally have l: "?lhs = (\<Sum>i=0..<n. A$$(i, i)) - (\<Sum>i=0..<n. B$$(i, i))".
have r: "?rhs = (\<Sum>i=0..<n. A$$(i, i)) - (\<Sum>i=0..<n. B$$(i, i))" unfolding trace_def using A B by auto
from l r show ?thesis by auto
qed
lemma trace_smult:
assumes "A \<in> carrier_mat n n"
shows "trace (c \<cdot>\<^sub>m A) = c * trace A"
proof -
have "trace (c \<cdot>\<^sub>m A) = (\<Sum>i = 0..<dim_row A. c * A $$ (i, i))" unfolding trace_def using assms by auto
also have "\<dots> = c * (\<Sum>i = 0..<dim_row A. A $$ (i, i))"
by (simp add: sum_distrib_left)
also have "\<dots> = c * trace A" unfolding trace_def by auto
ultimately show ?thesis by auto
qed
subsection \<open>Conjugate of a vector\<close>
lemma conjugate_scalar_prod:
fixes v w :: "'a::conjugatable_ring vec"
assumes "dim_vec v = dim_vec w"
shows "conjugate (v \<bullet> w) = conjugate v \<bullet> conjugate w"
using assms by (simp add: scalar_prod_def sum_conjugate conjugate_dist_mul)
subsection \<open>Inner product\<close>
abbreviation inner_prod :: "'a vec \<Rightarrow> 'a vec \<Rightarrow> 'a :: conjugatable_ring"
where "inner_prod v w \<equiv> w \<bullet>c v"
lemma conjugate_scalar_prod_Im [simp]:
"Im (v \<bullet>c v) = 0"
by (simp add: scalar_prod_def conjugate_vec_def sum.neutral)
lemma conjugate_scalar_prod_Re [simp]:
"Re (v \<bullet>c v) \<ge> 0"
by (simp add: scalar_prod_def conjugate_vec_def sum_nonneg)
lemma self_cscalar_prod_geq_0:
fixes v :: "'a::conjugatable_ordered_field vec"
shows "v \<bullet>c v \<ge> 0"
by (auto simp add: scalar_prod_def, rule sum_nonneg, rule conjugate_square_positive)
lemma inner_prod_distrib_left:
fixes u v w :: "('a::conjugatable_field) vec"
assumes dimu: "u \<in> carrier_vec n" and dimv:"v \<in> carrier_vec n" and dimw: "w \<in> carrier_vec n"
shows "inner_prod (v + w) u = inner_prod v u + inner_prod w u" (is "?lhs = ?rhs")
proof -
have dimcv: "conjugate v \<in> carrier_vec n" and dimcw: "conjugate w \<in> carrier_vec n" using assms by auto
have dimvw: "conjugate (v + w) \<in> carrier_vec n" using assms by auto
have "u \<bullet> (conjugate (v + w)) = u \<bullet> conjugate v + u \<bullet> conjugate w"
using dimv dimw dimu dimcv dimcw
by (metis conjugate_add_vec scalar_prod_add_distrib)
then show ?thesis by auto
qed
lemma inner_prod_distrib_right:
fixes u v w :: "('a::conjugatable_field) vec"
assumes dimu: "u \<in> carrier_vec n" and dimv:"v \<in> carrier_vec n" and dimw: "w \<in> carrier_vec n"
shows "inner_prod u (v + w) = inner_prod u v + inner_prod u w" (is "?lhs = ?rhs")
proof -
have dimvw: "v + w \<in> carrier_vec n" using assms by auto
have dimcu: "conjugate u \<in> carrier_vec n" using assms by auto
have "(v + w) \<bullet> (conjugate u) = v \<bullet> conjugate u + w \<bullet> conjugate u"
apply (simp add: comm_scalar_prod[OF dimvw dimcu])
apply (simp add: scalar_prod_add_distrib[OF dimcu dimv dimw])
apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n])
done
then show ?thesis by auto
qed
lemma inner_prod_minus_distrib_right:
fixes u v w :: "('a::conjugatable_field) vec"
assumes dimu: "u \<in> carrier_vec n" and dimv:"v \<in> carrier_vec n" and dimw: "w \<in> carrier_vec n"
shows "inner_prod u (v - w) = inner_prod u v - inner_prod u w" (is "?lhs = ?rhs")
proof -
have dimvw: "v - w \<in> carrier_vec n" using assms by auto
have dimcu: "conjugate u \<in> carrier_vec n" using assms by auto
have "(v - w) \<bullet> (conjugate u) = v \<bullet> conjugate u - w \<bullet> conjugate u"
apply (simp add: comm_scalar_prod[OF dimvw dimcu])
apply (simp add: scalar_prod_minus_distrib[OF dimcu dimv dimw])
apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n])
done
then show ?thesis by auto
qed
lemma inner_prod_smult_right:
fixes u v :: "complex vec"
assumes dimu: "u \<in> carrier_vec n" and dimv:"v \<in> carrier_vec n"
shows "inner_prod (a \<cdot>\<^sub>v u) v = conjugate a * inner_prod u v" (is "?lhs = ?rhs")
using assms apply (simp add: scalar_prod_def conjugate_dist_mul)
apply (subst sum_distrib_left) by (rule sum.cong, auto)
lemma inner_prod_smult_left:
fixes u v :: "complex vec"
assumes dimu: "u \<in> carrier_vec n" and dimv: "v \<in> carrier_vec n"
shows "inner_prod u (a \<cdot>\<^sub>v v) = a * inner_prod u v" (is "?lhs = ?rhs")
using assms apply (simp add: scalar_prod_def)
apply (subst sum_distrib_left) by (rule sum.cong, auto)
lemma inner_prod_smult_left_right:
fixes u v :: "complex vec"
assumes dimu: "u \<in> carrier_vec n" and dimv: "v \<in> carrier_vec n"
shows "inner_prod (a \<cdot>\<^sub>v u) (b \<cdot>\<^sub>v v) = conjugate a * b * inner_prod u v" (is "?lhs = ?rhs")
using assms apply (simp add: scalar_prod_def)
apply (subst sum_distrib_left) by (rule sum.cong, auto)
lemma inner_prod_swap:
fixes x y :: "complex vec"
assumes "y \<in> carrier_vec n" and "x \<in> carrier_vec n"
shows "inner_prod y x = conjugate (inner_prod x y)"
apply (simp add: scalar_prod_def)
apply (rule sum.cong) using assms by auto
text \<open>Cauchy-Schwarz theorem for complex vectors. This is analogous to aux\_Cauchy
and Cauchy\_Schwarz\_ineq in Generalizations2.thy in QR\_Decomposition. Consider
merging and moving to Isabelle library.\<close>
lemma aux_Cauchy:
fixes x y :: "complex vec"
assumes "x \<in> carrier_vec n" and "y \<in> carrier_vec n"
shows "0 \<le> inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y))"
proof -
have "(inner_prod (x+ a \<cdot>\<^sub>v y) (x+a \<cdot>\<^sub>v y)) = (inner_prod (x+a \<cdot>\<^sub>v y) x) + (inner_prod (x+a \<cdot>\<^sub>v y) (a \<cdot>\<^sub>v y))"
apply (subst inner_prod_distrib_right) using assms by auto
also have "\<dots> = inner_prod x x + (a) * (inner_prod x y) + cnj a * ((cnj (inner_prod x y)) + (a) * (inner_prod y y))"
apply (subst (1 2) inner_prod_distrib_left[of _ n]) apply (auto simp add: assms)
apply (subst (1 2) inner_prod_smult_right[of _ n]) apply (auto simp add: assms)
apply (subst inner_prod_smult_left[of _ n]) apply (auto simp add: assms)
apply (subst inner_prod_swap[of y n x]) apply (auto simp add: assms)
unfolding distrib_left
by auto
finally show ?thesis by (metis self_cscalar_prod_geq_0)
qed
lemma Cauchy_Schwarz_complex_vec:
fixes x y :: "complex vec"
assumes "x \<in> carrier_vec n" and "y \<in> carrier_vec n"
shows "inner_prod x y * inner_prod y x \<le> inner_prod x x * inner_prod y y"
proof -
define cnj_a where "cnj_a = - (inner_prod x y)/ cnj (inner_prod y y)"
define a where "a = cnj (cnj_a)"
have cnj_rw: "(cnj a) = cnj_a"
unfolding a_def by (simp)
have rw_0: "cnj (inner_prod x y) + a * (inner_prod y y) = 0"
unfolding a_def cnj_a_def using assms(1) assms(2) conjugate_square_eq_0_vec by fastforce
have "0 \<le> (inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y)))"
using aux_Cauchy assms by auto
also have "\<dots> = (inner_prod x x + a * (inner_prod x y))" unfolding rw_0 by auto
also have "\<dots> = (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y))"
unfolding a_def cnj_a_def by simp
finally have " 0 \<le> (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) " .
hence "0 \<le> (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) * (inner_prod y y)" by auto
also have "\<dots> = ((inner_prod x x)*(inner_prod y y) - (inner_prod x y) * cnj (inner_prod x y))"
by (smt add.inverse_neutral add_diff_cancel diff_0 diff_divide_eq_iff divide_cancel_right mult_eq_0_iff nonzero_mult_div_cancel_right rw_0)
finally have "(inner_prod x y) * cnj (inner_prod x y) \<le> (inner_prod x x)*(inner_prod y y)" by auto
then show ?thesis
apply (subst inner_prod_swap[of y n x]) by (auto simp add: assms)
qed
subsection \<open>Hermitian adjoint of a matrix\<close>
abbreviation adjoint where "adjoint \<equiv> mat_adjoint"
lemma adjoint_dim_row [simp]:
"dim_row (adjoint A) = dim_col A" by (simp add: mat_adjoint_def)
lemma adjoint_dim_col [simp]:
"dim_col (adjoint A) = dim_row A" by (simp add: mat_adjoint_def)
lemma adjoint_dim:
"A \<in> carrier_mat n n \<Longrightarrow> adjoint A \<in> carrier_mat n n"
using adjoint_dim_col adjoint_dim_row by blast
lemma adjoint_def:
"adjoint A = mat (dim_col A) (dim_row A) (\<lambda>(i,j). conjugate (A $$ (j,i)))"
unfolding mat_adjoint_def mat_of_rows_def by auto
lemma adjoint_eval:
assumes "i < dim_col A" "j < dim_row A"
shows "(adjoint A) $$ (i,j) = conjugate (A $$ (j,i))"
using assms by (simp add: adjoint_def)
lemma adjoint_row:
assumes "i < dim_col A"
shows "row (adjoint A) i = conjugate (col A i)"
apply (rule eq_vecI)
using assms by (auto simp add: adjoint_eval)
lemma adjoint_col:
assumes "i < dim_row A"
shows "col (adjoint A) i = conjugate (row A i)"
apply (rule eq_vecI)
using assms by (auto simp add: adjoint_eval)
text \<open>The identity <v, A w> = <A* v, w>\<close>
lemma adjoint_def_alter:
fixes v w :: "'a::conjugatable_field vec"
and A :: "'a::conjugatable_field mat"
assumes dims: "v \<in> carrier_vec n" "w \<in> carrier_vec m" "A \<in> carrier_mat n m"
shows "inner_prod v (A *\<^sub>v w) = inner_prod (adjoint A *\<^sub>v v) w" (is "?lhs = ?rhs")
proof -
from dims have "?lhs = (\<Sum>i=0..<dim_vec v. (\<Sum>j=0..<dim_vec w.
conjugate (v$i) * A$$(i, j) * w$j))"
apply (simp add: scalar_prod_def sum_distrib_right )
apply (rule sum.cong, simp)
apply (rule sum.cong, auto)
done
moreover from assms have "?rhs = (\<Sum>i=0..<dim_vec v. (\<Sum>j=0..<dim_vec w.
conjugate (v$i) * A$$(i, j) * w$j))"
apply (simp add: scalar_prod_def adjoint_eval
sum_conjugate conjugate_dist_mul sum_distrib_left)
apply (subst sum.swap[where ?A = "{0..<n}"])
apply (rule sum.cong, simp)
apply (rule sum.cong, auto)
done
ultimately show ?thesis by simp
qed
lemma adjoint_one:
shows "adjoint (1\<^sub>m n) = (1\<^sub>m n::complex mat)"
apply (rule eq_matI)
by (auto simp add: adjoint_eval)
lemma adjoint_scale:
fixes A :: "'a::conjugatable_field mat"
shows "adjoint (a \<cdot>\<^sub>m A) = (conjugate a) \<cdot>\<^sub>m adjoint A"
apply (rule eq_matI) using conjugatable_ring_class.conjugate_dist_mul
by (auto simp add: adjoint_eval)
lemma adjoint_add:
fixes A B :: "'a::conjugatable_field mat"
assumes "A \<in> carrier_mat n m" "B \<in> carrier_mat n m"
shows "adjoint (A + B) = adjoint A + adjoint B"
apply (rule eq_matI)
using assms conjugatable_ring_class.conjugate_dist_add
by( auto simp add: adjoint_eval)
lemma adjoint_minus:
fixes A B :: "'a::conjugatable_field mat"
assumes "A \<in> carrier_mat n m" "B \<in> carrier_mat n m"
shows "adjoint (A - B) = adjoint A - adjoint B"
apply (rule eq_matI)
using assms apply(auto simp add: adjoint_eval)
by (metis add_uminus_conv_diff conjugate_dist_add conjugate_neg)
lemma adjoint_mult:
fixes A B :: "'a::conjugatable_field mat"
assumes "A \<in> carrier_mat n m" "B \<in> carrier_mat m l"
shows "adjoint (A * B) = adjoint B * adjoint A"
proof (rule eq_matI, auto simp add: adjoint_eval adjoint_row adjoint_col)
fix i j
assume "i < dim_col B" "j < dim_row A"
show "conjugate (row A j \<bullet> col B i) = conjugate (col B i) \<bullet> conjugate (row A j)"
using assms apply (simp add: conjugate_scalar_prod)
apply (subst comm_scalar_prod[where n="dim_row B"])
by (auto simp add: carrier_vecI)
qed
lemma adjoint_adjoint:
fixes A :: "'a::conjugatable_field mat"
shows "adjoint (adjoint A) = A"
by (rule eq_matI, auto simp add: adjoint_eval)
lemma trace_adjoint_positive:
fixes A :: "complex mat"
shows "trace (A * adjoint A) \<ge> 0"
apply (auto simp add: trace_def adjoint_col)
apply (rule sum_nonneg) by auto
subsection \<open>Algebraic manipulations on matrices\<close>
lemma right_add_zero_mat[simp]:
"(A :: 'a :: monoid_add mat) \<in> carrier_mat nr nc \<Longrightarrow> A + 0\<^sub>m nr nc = A"
by (intro eq_matI, auto)
lemma add_carrier_mat':
"A \<in> carrier_mat nr nc \<Longrightarrow> B \<in> carrier_mat nr nc \<Longrightarrow> A + B \<in> carrier_mat nr nc"
by simp
lemma minus_carrier_mat':
"A \<in> carrier_mat nr nc \<Longrightarrow> B \<in> carrier_mat nr nc \<Longrightarrow> A - B \<in> carrier_mat nr nc"
by auto
lemma swap_plus_mat:
fixes A B C :: "'a::semiring_1 mat"
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "C \<in> carrier_mat n n"
shows "A + B + C = A + C + B"
by (metis assms assoc_add_mat comm_add_mat)
lemma uminus_mat:
fixes A :: "complex mat"
assumes "A \<in> carrier_mat n n"
shows "-A = (-1) \<cdot>\<^sub>m A"
by auto
ML_file "mat_alg.ML"
-method_setup mat_assoc = {* mat_assoc_method *}
+method_setup mat_assoc = \<open>mat_assoc_method\<close>
"Normalization of expressions on matrices"
lemma mat_assoc_test:
fixes A B C D :: "complex mat"
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "C \<in> carrier_mat n n" "D \<in> carrier_mat n n"
shows
"(A * B) * (C * D) = A * B * C * D"
"adjoint (A * adjoint B) * C = B * (adjoint A * C)"
"A * 1\<^sub>m n * 1\<^sub>m n * B * 1\<^sub>m n = A * B"
"(A - B) + (B - C) = A + (-B) + B + (-C)"
"A + (B - C) = A + B - C"
"A - (B + C + D) = A - B - C - D"
"(A + B) * (B + C) = A * B + B * B + A * C + B * C"
"A - B = A + (-1) \<cdot>\<^sub>m B"
"A * (B - C) * D = A * B * D - A * C * D"
"trace (A * B * C) = trace (B * C * A)"
"trace (A * B * C * D) = trace (C * D * A * B)"
"trace (A + B * C) = trace A + trace (C * B)"
"A + B = B + A"
"A + B + C = C + B + A"
"A + B + (C + D) = A + C + (B + D)"
using assms by (mat_assoc n)+
subsection \<open>Hermitian matrices\<close>
text \<open>A Hermitian matrix is a matrix that is equal to its Hermitian adjoint.\<close>
definition hermitian :: "'a::conjugatable_field mat \<Rightarrow> bool" where
"hermitian A \<longleftrightarrow> (adjoint A = A)"
lemma hermitian_one:
shows "hermitian ((1\<^sub>m n)::('a::conjugatable_field mat))"
unfolding hermitian_def
proof-
have "conjugate (1::'a) = 1"
apply (subst mult_1_right[symmetric, of "conjugate 1"])
apply (subst conjugate_id[symmetric, of "conjugate 1 * 1"])
apply (subst conjugate_dist_mul)
apply auto
done
then show "adjoint ((1\<^sub>m n)::('a::conjugatable_field mat)) = (1\<^sub>m n)"
by (auto simp add: adjoint_eval)
qed
subsection \<open>Inverse matrices\<close>
lemma inverts_mat_symm:
fixes A B :: "'a::field mat"
assumes dim: "A \<in> carrier_mat n n" "B \<in> carrier_mat n n"
and AB: "inverts_mat A B"
shows "inverts_mat B A"
proof -
have "A * B = 1\<^sub>m n" using dim AB unfolding inverts_mat_def by auto
with dim have "B * A = 1\<^sub>m n" by (rule mat_mult_left_right_inverse)
then show "inverts_mat B A" using dim inverts_mat_def by auto
qed
lemma inverts_mat_unique:
fixes A B C :: "'a::field mat"
assumes dim: "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "C \<in> carrier_mat n n"
and AB: "inverts_mat A B" and AC: "inverts_mat A C"
shows "B = C"
proof -
have AB1: "A * B = 1\<^sub>m n" using AB dim unfolding inverts_mat_def by auto
have "A * C = 1\<^sub>m n" using AC dim unfolding inverts_mat_def by auto
then have CA1: "C * A = 1\<^sub>m n" using mat_mult_left_right_inverse[of A n C] dim by auto
then have "C = C * 1\<^sub>m n" using dim by auto
also have "\<dots> = C * (A * B)" using AB1 by auto
also have "\<dots> = (C * A) * B" using dim by auto
also have "\<dots> = 1\<^sub>m n * B" using CA1 by auto
also have "\<dots> = B" using dim by auto
finally show "B = C" ..
qed
subsection \<open>Unitary matrices\<close>
text \<open>A unitary matrix is a matrix whose Hermitian adjoint is also its inverse.\<close>
definition unitary :: "'a::conjugatable_field mat \<Rightarrow> bool" where
"unitary A \<longleftrightarrow> A \<in> carrier_mat (dim_row A) (dim_row A) \<and> inverts_mat A (adjoint A)"
lemma unitaryD2:
assumes "A \<in> carrier_mat n n"
shows "unitary A \<Longrightarrow> inverts_mat (adjoint A) A"
using assms adjoint_dim inverts_mat_symm unitary_def by blast
lemma unitary_simps [simp]:
"A \<in> carrier_mat n n \<Longrightarrow> unitary A \<Longrightarrow> adjoint A * A = 1\<^sub>m n"
"A \<in> carrier_mat n n \<Longrightarrow> unitary A \<Longrightarrow> A * adjoint A = 1\<^sub>m n"
apply (metis adjoint_dim_row carrier_matD(2) inverts_mat_def unitaryD2)
by (simp add: inverts_mat_def unitary_def)
lemma unitary_adjoint [simp]:
assumes "A \<in> carrier_mat n n" "unitary A"
shows "unitary (adjoint A)"
unfolding unitary_def
using adjoint_dim[OF assms(1)] assms by (auto simp add: unitaryD2[OF assms] adjoint_adjoint)
lemma unitary_one:
shows "unitary ((1\<^sub>m n)::('a::conjugatable_field mat))"
unfolding unitary_def
proof -
define I where I_def[simp]: "I \<equiv> ((1\<^sub>m n)::('a::conjugatable_field mat))"
have dim: "I \<in> carrier_mat n n" by auto
have "hermitian I" using hermitian_one by auto
hence "adjoint I = I" using hermitian_def by auto
with dim show "I \<in> carrier_mat (dim_row I) (dim_row I) \<and> inverts_mat I (adjoint I)"
unfolding inverts_mat_def using dim by auto
qed
lemma unitary_zero:
fixes A :: "'a::conjugatable_field mat"
assumes "A \<in> carrier_mat 0 0"
shows "unitary A"
unfolding unitary_def inverts_mat_def Let_def using assms by auto
lemma unitary_elim:
assumes dims: "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "P \<in> carrier_mat n n"
and uP: "unitary P" and eq: "P * A * adjoint P = P * B * adjoint P"
shows "A = B"
proof -
have dimaP: "adjoint P \<in> carrier_mat n n" using dims by auto
have iv: "inverts_mat P (adjoint P)" using uP unitary_def by auto
then have "P * (adjoint P) = 1\<^sub>m n" using inverts_mat_def dims by auto
then have aPP: "adjoint P * P = 1\<^sub>m n" using mat_mult_left_right_inverse[OF dims(3) dimaP] by auto
have "adjoint P * (P * A * adjoint P) * P = (adjoint P * P) * A * (adjoint P * P)"
using dims dimaP by (mat_assoc n)
also have "\<dots> = 1\<^sub>m n * A * 1\<^sub>m n" using aPP by auto
also have "\<dots> = A" using dims by auto
finally have eqA: "A = adjoint P * (P * A * adjoint P) * P" ..
have "adjoint P * (P * B * adjoint P) * P = (adjoint P * P) * B * (adjoint P * P)"
using dims dimaP by (mat_assoc n)
also have "\<dots> = 1\<^sub>m n * B * 1\<^sub>m n" using aPP by auto
also have "\<dots> = B" using dims by auto
finally have eqB: "B = adjoint P * (P * B * adjoint P) * P" ..
then show ?thesis using eqA eqB eq by auto
qed
lemma unitary_is_corthogonal:
fixes U :: "'a::conjugatable_field mat"
assumes dim: "U \<in> carrier_mat n n"
and U: "unitary U"
shows "corthogonal_mat U"
unfolding corthogonal_mat_def Let_def
proof (rule conjI)
have dima: "adjoint U \<in> carrier_mat n n" using dim by auto
have aUU: "mat_adjoint U * U = (1\<^sub>m n)"
apply (insert U[unfolded unitary_def] dim dima, drule conjunct2)
apply (drule inverts_mat_symm[of "U", OF dim dima], unfold inverts_mat_def, auto)
done
then show "diagonal_mat (mat_adjoint U * U)"
by (simp add: diagonal_mat_def)
show "\<forall>i<dim_col U. (mat_adjoint U * U) $$ (i, i) \<noteq> 0" using dim by (simp add: aUU)
qed
lemma unitary_times_unitary:
fixes P Q :: "'a:: conjugatable_field mat"
assumes dim: "P \<in> carrier_mat n n" "Q \<in> carrier_mat n n"
and uP: "unitary P" and uQ: "unitary Q"
shows "unitary (P * Q)"
proof -
have dim_pq: "P * Q \<in> carrier_mat n n" using dim by auto
have "(P * Q) * adjoint (P * Q) = P * (Q * adjoint Q) * adjoint P" using dim by (mat_assoc n)
also have "\<dots> = P * (1\<^sub>m n) * adjoint P" using uQ dim by auto
also have "\<dots> = P * adjoint P" using dim by (mat_assoc n)
also have "\<dots> = 1\<^sub>m n" using uP dim by simp
finally have "(P * Q) * adjoint (P * Q) = 1\<^sub>m n" by auto
hence "inverts_mat (P * Q) (adjoint (P * Q))"
using inverts_mat_def dim_pq by auto
thus "unitary (P*Q)" using unitary_def dim_pq by auto
qed
lemma unitary_operator_keep_trace:
fixes U A :: "complex mat"
assumes dU: "U \<in> carrier_mat n n" and dA: "A \<in> carrier_mat n n" and u: "unitary U"
shows "trace A = trace (adjoint U * A * U)"
proof -
have u': "U * adjoint U = 1\<^sub>m n" using u unfolding unitary_def inverts_mat_def using dU by auto
have "trace (adjoint U * A * U) = trace (U * adjoint U * A)" using dU dA by (mat_assoc n)
also have "\<dots> = trace A" using u' dA by auto
finally show ?thesis by auto
qed
subsection \<open>Normalization of vectors\<close>
definition vec_norm :: "complex vec \<Rightarrow> complex" where
"vec_norm v \<equiv> csqrt (v \<bullet>c v)"
lemma vec_norm_geq_0:
fixes v :: "complex vec"
shows "vec_norm v \<ge> 0"
unfolding vec_norm_def by (insert self_cscalar_prod_geq_0[of v], simp)
lemma vec_norm_zero:
fixes v :: "complex vec"
assumes dim: "v \<in> carrier_vec n"
shows "vec_norm v = 0 \<longleftrightarrow> v = 0\<^sub>v n"
unfolding vec_norm_def
by (subst conjugate_square_eq_0_vec[OF dim, symmetric], rule csqrt_eq_0)
lemma vec_norm_ge_0:
fixes v :: "complex vec"
assumes dim_v: "v \<in> carrier_vec n" and neq0: "v \<noteq> 0\<^sub>v n"
shows "vec_norm v > 0"
proof -
have geq: "vec_norm v \<ge> 0" using vec_norm_geq_0 by auto
have neq: "vec_norm v \<noteq> 0"
apply (insert dim_v neq0)
apply (drule vec_norm_zero, auto)
done
show ?thesis using neq geq by (rule dual_order.not_eq_order_implies_strict)
qed
definition vec_normalize :: "complex vec \<Rightarrow> complex vec" where
"vec_normalize v = (if (v = 0\<^sub>v (dim_vec v)) then v else 1 / (vec_norm v) \<cdot>\<^sub>v v)"
lemma normalized_vec_dim[simp]:
assumes "(v::complex vec) \<in> carrier_vec n"
shows "vec_normalize v \<in> carrier_vec n"
unfolding vec_normalize_def using assms by auto
lemma vec_eq_norm_smult_normalized:
shows "v = vec_norm v \<cdot>\<^sub>v vec_normalize v"
proof (cases "v = 0\<^sub>v (dim_vec v)")
define n where "n = dim_vec v"
then have dimv: "v \<in> carrier_vec n" by auto
then have dimnv: "vec_normalize v \<in> carrier_vec n" by auto
{
case True
then have v0: "v = 0\<^sub>v n" using n_def by auto
then have n0: "vec_norm v = 0" using vec_norm_def by auto
have "vec_norm v \<cdot>\<^sub>v vec_normalize v = 0\<^sub>v n"
unfolding smult_vec_def by (auto simp add: n0 carrier_vecD[OF dimnv])
then show ?thesis using v0 by auto
next
case False
then have v: "v \<noteq> 0\<^sub>v n" using n_def by auto
then have ge0: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto
have "vec_normalize v = (1 / vec_norm v) \<cdot>\<^sub>v v" using False vec_normalize_def by auto
then have "vec_norm v \<cdot>\<^sub>v vec_normalize v = (vec_norm v * (1 / vec_norm v)) \<cdot>\<^sub>v v"
using smult_smult_assoc by auto
also have "\<dots> = v" using ge0 by auto
finally have "v = vec_norm v \<cdot>\<^sub>v vec_normalize v"..
then show "v = vec_norm v \<cdot>\<^sub>v vec_normalize v" using v by auto
}
qed
lemma normalized_cscalar_prod:
fixes v w :: "complex vec"
assumes dim_v: "v \<in> carrier_vec n" and dim_w: "w \<in> carrier_vec n"
shows "v \<bullet>c w = (vec_norm v * vec_norm w) * (vec_normalize v \<bullet>c vec_normalize w)"
unfolding vec_normalize_def apply (split if_split, split if_split)
proof (intro conjI impI)
note dim0 = dim_v dim_w
have dim: "dim_vec v = n" "dim_vec w = n" using dim0 by auto
{
assume "w = 0\<^sub>v n" "v = 0\<^sub>v n"
then have lhs: "v \<bullet>c w = 0" by auto
then moreover have rhs: "vec_norm v * vec_norm w * (v \<bullet>c w) = 0" by auto
ultimately have "v \<bullet>c w = vec_norm v * vec_norm w * (v \<bullet>c w)" by auto
}
with dim show "w = 0\<^sub>v (dim_vec w) \<Longrightarrow> v = 0\<^sub>v (dim_vec v) \<Longrightarrow> v \<bullet>c w = vec_norm v * vec_norm w * (v \<bullet>c w)" by auto
{
assume asm: "w = 0\<^sub>v n" "v \<noteq> 0\<^sub>v n"
then have w0: "conjugate w = 0\<^sub>v n" by auto
with dim0 have "(1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c w = 0" by auto
then moreover have rhs: "vec_norm v * vec_norm w * ((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c w) = 0" by auto
moreover have "v \<bullet>c w = 0" using w0 dim0 by auto
ultimately have "v \<bullet>c w = vec_norm v * vec_norm w * ((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c w)" by auto
}
with dim show "w = 0\<^sub>v (dim_vec w) \<Longrightarrow> v \<noteq> 0\<^sub>v (dim_vec v) \<Longrightarrow> v \<bullet>c w = vec_norm v * vec_norm w * ((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c w)" by auto
{
assume asm: "w \<noteq> 0\<^sub>v n" "v = 0\<^sub>v n"
with dim0 have "v \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w) = 0" by auto
then moreover have rhs: "vec_norm v * vec_norm w * (v \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w)) = 0" by auto
moreover have "v \<bullet>c w = 0" using asm dim0 by auto
ultimately have "v \<bullet>c w = vec_norm v * vec_norm w * (v \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w))" by auto
}
with dim show "w \<noteq> 0\<^sub>v (dim_vec w) \<Longrightarrow> v = 0\<^sub>v (dim_vec v) \<Longrightarrow> v \<bullet>c w = vec_norm v * vec_norm w * (v \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w))" by auto
{
assume asmw: "w \<noteq> 0\<^sub>v n" and asmv: "v \<noteq> 0\<^sub>v n"
have "vec_norm w > 0" by (insert asmw dim0, rule vec_norm_ge_0, auto)
then have cw: "conjugate (1 / vec_norm w) = 1 / vec_norm w" by (simp add: complex_eq_iff complex_is_Real_iff)
from dim0 have
"((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w)) = 1 / vec_norm v * (v \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w))" by auto
also have "\<dots> = 1 / vec_norm v * (v \<bullet> (conjugate (1 / vec_norm w) \<cdot>\<^sub>v conjugate w))"
by (subst conjugate_smult_vec, auto)
also have "\<dots> = 1 / vec_norm v * conjugate (1 / vec_norm w) * (v \<bullet> conjugate w)" using dim by auto
also have "\<dots> = 1 / vec_norm v * (1 / vec_norm w) * (v \<bullet>c w)" using vec_norm_ge_0 cw by auto
finally have eq1: "(1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w) = 1 / vec_norm v * (1 / vec_norm w) * (v \<bullet>c w)" .
then have "vec_norm v * vec_norm w * ((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w)) = (v \<bullet>c w)"
by (subst eq1, insert vec_norm_ge_0[of v n, OF dim_v asmv] vec_norm_ge_0[of w n, OF dim_w asmw], auto)
}
with dim show " w \<noteq> 0\<^sub>v (dim_vec w) \<Longrightarrow> v \<noteq> 0\<^sub>v (dim_vec v) \<Longrightarrow> v \<bullet>c w = vec_norm v * vec_norm w * ((1 / vec_norm v \<cdot>\<^sub>v v) \<bullet>c (1 / vec_norm w \<cdot>\<^sub>v w))" by auto
qed
lemma normalized_vec_norm :
fixes v :: "complex vec"
assumes dim_v: "v \<in> carrier_vec n"
and neq0: "v \<noteq> 0\<^sub>v n"
shows "vec_normalize v \<bullet>c vec_normalize v = 1"
unfolding vec_normalize_def
proof (simp, rule conjI)
show "v = 0\<^sub>v (dim_vec v) \<longrightarrow> v \<bullet>c v = 1" using neq0 dim_v by auto
have dim_a: "(vec_normalize v) \<in> carrier_vec n" "conjugate (vec_normalize v) \<in> carrier_vec n" using dim_v vec_normalize_def by auto
note dim = dim_v dim_a
have nvge0: "vec_norm v > 0" using vec_norm_ge_0 neq0 dim_v by auto
then have vvvv: "v \<bullet>c v = (vec_norm v) * (vec_norm v)" unfolding vec_norm_def by (metis power2_csqrt power2_eq_square)
from nvge0 have "conjugate (vec_norm v) = vec_norm v" by (simp add: complex_eq_iff complex_is_Real_iff)
then have "v \<bullet>c (1 / vec_norm v \<cdot>\<^sub>v v) = 1 / vec_norm v * (v \<bullet>c v)"
by (subst conjugate_smult_vec, auto)
also have "\<dots> = 1 / vec_norm v * vec_norm v * vec_norm v" using vvvv by auto
also have "\<dots> = vec_norm v" by auto
finally have "v \<bullet>c (1 / vec_norm v \<cdot>\<^sub>v v) = vec_norm v".
then show "v \<noteq> 0\<^sub>v (dim_vec v) \<longrightarrow> vec_norm v \<noteq> 0 \<and> v \<bullet>c (1 / vec_norm v \<cdot>\<^sub>v v) = vec_norm v"
using neq0 nvge0 by auto
qed
lemma normalize_zero:
assumes "v \<in> carrier_vec n"
shows "vec_normalize v = 0\<^sub>v n \<longleftrightarrow> v = 0\<^sub>v n"
proof
show "v = 0\<^sub>v n \<Longrightarrow> vec_normalize v = 0\<^sub>v n" unfolding vec_normalize_def by auto
next
have "v \<noteq> 0\<^sub>v n \<Longrightarrow> vec_normalize v \<noteq> 0\<^sub>v n" unfolding vec_normalize_def
proof (simp, rule impI)
assume asm: "v \<noteq> 0\<^sub>v n"
then have "vec_norm v > 0" using vec_norm_ge_0 assms by auto
then have nvge0: "1 / vec_norm v > 0" by (simp add: complex_is_Real_iff)
have "\<exists>k < n. v $ k \<noteq> 0" using asm assms by auto
then obtain k where kn: "k < n" and vkneq0: "v $ k \<noteq> 0" by auto
then have "(1 / vec_norm v \<cdot>\<^sub>v v) $ k = (1 / vec_norm v) * (v $ k)"
using assms carrier_vecD index_smult_vec(1) by blast
with nvge0 vkneq0 have "(1 / vec_norm v \<cdot>\<^sub>v v) $ k \<noteq> 0" by auto
then show "1 / vec_norm v \<cdot>\<^sub>v v \<noteq> 0\<^sub>v n" using assms kn by fastforce
qed
then show "vec_normalize v = 0\<^sub>v n \<Longrightarrow> v = 0\<^sub>v n" by auto
qed
lemma normalize_normalize[simp]:
"vec_normalize (vec_normalize v) = vec_normalize v"
proof (rule disjE[of "v = 0\<^sub>v (dim_vec v)" "v \<noteq> 0\<^sub>v (dim_vec v)"], auto)
let ?n = "dim_vec v"
{
assume "v = 0\<^sub>v ?n"
then have "vec_normalize v = v" unfolding vec_normalize_def by auto
then show "vec_normalize (vec_normalize v) = vec_normalize v" by auto
}
assume neq0: "v \<noteq> 0\<^sub>v ?n"
have dim: "v \<in> carrier_vec ?n" by auto
have "vec_norm (vec_normalize v) = 1" unfolding vec_norm_def
using normalized_vec_norm[OF dim neq0] by auto
then show "vec_normalize (vec_normalize v) = vec_normalize v"
by (subst (1) vec_normalize_def, simp)
qed
subsection \<open>Spectral decomposition of normal complex matrices\<close>
lemma normalize_keep_corthogonal:
fixes vs :: "complex vec list"
assumes cor: "corthogonal vs" and dims: "set vs \<subseteq> carrier_vec n"
shows "corthogonal (map vec_normalize vs)"
unfolding corthogonal_def
proof (rule allI, rule impI, rule allI, rule impI, goal_cases)
case c: (1 i j)
let ?m = "length vs"
have len: "length (map vec_normalize vs) = ?m" by auto
have dim: "\<And>k. k < ?m \<Longrightarrow> (vs ! k) \<in> carrier_vec n" using dims by auto
have map: "\<And>k. k < ?m \<Longrightarrow> map vec_normalize vs ! k = vec_normalize (vs ! k)" by auto
have eq1: "\<And>j k. j < ?m \<Longrightarrow> k < ?m \<Longrightarrow> ((vs ! j) \<bullet>c (vs ! k) = 0) = (j \<noteq> k)" using assms unfolding corthogonal_def by auto
then have "\<And>k. k < ?m \<Longrightarrow> (vs ! k) \<bullet>c (vs ! k) \<noteq> 0 " by auto
then have "\<And>k. k < ?m \<Longrightarrow> (vs ! k) \<noteq> (0\<^sub>v n)" using dim
by (auto simp add: conjugate_square_eq_0_vec[of _ n, OF dim])
then have vnneq0: "\<And>k. k < ?m \<Longrightarrow> vec_norm (vs ! k) \<noteq> 0" using vec_norm_zero[OF dim] by auto
then have i0: "vec_norm (vs ! i) \<noteq> 0" and j0: "vec_norm (vs ! j) \<noteq> 0" using c by auto
have "(vs ! i) \<bullet>c (vs ! j) = vec_norm (vs ! i) * vec_norm (vs ! j) * (vec_normalize (vs ! i) \<bullet>c vec_normalize (vs ! j))"
by (subst normalized_cscalar_prod[of "vs ! i" n "vs ! j"], auto, insert dim c, auto)
with i0 j0 have "(vec_normalize (vs ! i) \<bullet>c vec_normalize (vs ! j) = 0) = ((vs ! i) \<bullet>c (vs ! j) = 0)" by auto
with eq1 c have "(vec_normalize (vs ! i) \<bullet>c vec_normalize (vs ! j) = 0) = (i \<noteq> j)" by auto
with map c show "(map vec_normalize vs ! i \<bullet>c map vec_normalize vs ! j = 0) = (i \<noteq> j)" by auto
qed
lemma normalized_corthogonal_mat_is_unitary:
assumes W: "set ws \<subseteq> carrier_vec n"
and orth: "corthogonal ws"
and len: "length ws = n"
shows "unitary (mat_of_cols n (map vec_normalize ws))" (is "unitary ?W")
proof -
define vs where "vs = map vec_normalize ws"
define W where "W = mat_of_cols n vs"
have W': "set vs \<subseteq> carrier_vec n" using assms vs_def by auto
then have W'': "\<And>k. k < length vs \<Longrightarrow> vs ! k \<in> carrier_vec n" by auto
have orth': "corthogonal vs" using assms normalize_keep_corthogonal vs_def by auto
have len'[simp]: "length vs = n" using assms vs_def by auto
have dimW: "W \<in> carrier_mat n n" using W_def len by auto
have "adjoint W \<in> carrier_mat n n" using dimW by auto
then have dimaW: "mat_adjoint W \<in> carrier_mat n n" by auto
{
fix i j assume i: "i < n" and j: "j < n"
have dimws: "(ws ! i) \<in> carrier_vec n" "(ws ! j) \<in> carrier_vec n" using W len i j by auto
have "(ws ! i) \<bullet>c (ws ! i) \<noteq> 0" "(ws ! j) \<bullet>c (ws ! j) \<noteq> 0" using orth corthogonal_def[of ws] len i j by auto
then have neq0: "(ws ! i) \<noteq> 0\<^sub>v n" "(ws ! j) \<noteq> 0\<^sub>v n"
by (auto simp add: conjugate_square_eq_0_vec[of "ws ! i" n])
then have "vec_norm (ws ! i) > 0" "vec_norm (ws ! j) > 0" using vec_norm_ge_0 dimws by auto
then have ge0: "vec_norm (ws ! i) * vec_norm (ws ! j) > 0" by auto
have ws': "vs ! i = vec_normalize (ws ! i)"
"vs ! j = vec_normalize (ws ! j)"
using len i j vs_def by auto
have ii1: "(vs ! i) \<bullet>c (vs ! i) = 1"
apply (simp add: ws')
apply (rule normalized_vec_norm[of "ws ! i"], rule dimws, rule neq0)
done
have ij0: "i \<noteq> j \<Longrightarrow> (ws ! i) \<bullet>c (ws ! j) = 0" using i j
by (insert orth, auto simp add: corthogonal_def[of ws] len)
have "i \<noteq> j \<Longrightarrow> (ws ! i) \<bullet>c (ws ! j) = (vec_norm (ws ! i) * vec_norm (ws ! j)) * ((vs ! i) \<bullet>c (vs ! j))"
apply (auto simp add: ws')
apply (rule normalized_cscalar_prod)
apply (rule dimws, rule dimws)
done
with ij0 have ij0': "i \<noteq> j \<Longrightarrow> (vs ! i) \<bullet>c (vs ! j) = 0" using ge0 by auto
have cWk: "\<And>k. k < n \<Longrightarrow> col W k = vs ! k" unfolding W_def
apply (subst col_mat_of_cols)
apply (auto simp add: W'')
done
have "(mat_adjoint W * W) $$ (j, i) = row (mat_adjoint W) j \<bullet> col W i"
by (insert dimW i j dimaW, auto)
also have "\<dots> = conjugate (col W j) \<bullet> col W i"
by (insert dimW i j dimaW, auto simp add: mat_adjoint_def)
also have "\<dots> = col W i \<bullet> conjugate (col W j)" using comm_scalar_prod[of "col W i" n] dimW by auto
also have "\<dots> = (vs ! i) \<bullet>c (vs ! j)" using W_def col_mat_of_cols i j len cWk by auto
finally have "(mat_adjoint W * W) $$ (j, i) = (vs ! i) \<bullet>c (vs ! j)".
then have "(mat_adjoint W * W) $$ (j, i) = (if (j = i) then 1 else 0)"
by (auto simp add: ii1 ij0')
}
note maWW = this
then have "mat_adjoint W * W = 1\<^sub>m n" unfolding one_mat_def using dimW dimaW
by (auto simp add: maWW adjoint_def)
then have iv0: "adjoint W * W = 1\<^sub>m n" by auto
have dimaW: "adjoint W \<in> carrier_mat n n" using dimaW by auto
then have iv1: "W * adjoint W = 1\<^sub>m n" using mat_mult_left_right_inverse dimW iv0 by auto
then show "unitary W" unfolding unitary_def inverts_mat_def using dimW dimaW iv0 iv1 by auto
qed
lemma normalize_keep_eigenvector:
assumes ev: "eigenvector A v e"
and dim: "A \<in> carrier_mat n n" "v \<in> carrier_vec n"
shows "eigenvector A (vec_normalize v) e"
unfolding eigenvector_def
proof
show "vec_normalize v \<in> carrier_vec (dim_row A)" using dim by auto
have eg: "A *\<^sub>v v = e \<cdot>\<^sub>v v" using ev dim eigenvector_def by auto
have vneq0: "v \<noteq> 0\<^sub>v n" using ev dim unfolding eigenvector_def by auto
then have s0: "vec_normalize v \<noteq> 0\<^sub>v n"
by (insert dim, subst normalize_zero[of v], auto)
from vneq0 have vvge0: "vec_norm v > 0" using vec_norm_ge_0 dim by auto
have s1: "A *\<^sub>v vec_normalize v = e \<cdot>\<^sub>v vec_normalize v" unfolding vec_normalize_def
using vneq0 dim
apply (auto, simp add: mult_mat_vec)
apply (subst eg, auto)
done
with s0 dim show "vec_normalize v \<noteq> 0\<^sub>v (dim_row A) \<and> A *\<^sub>v vec_normalize v = e \<cdot>\<^sub>v vec_normalize v" by auto
qed
lemma four_block_mat_adjoint:
fixes A B C D :: "'a::conjugatable_field mat"
assumes dim: "A \<in> carrier_mat nr1 nc1" "B \<in> carrier_mat nr1 nc2"
"C \<in> carrier_mat nr2 nc1" "D \<in> carrier_mat nr2 nc2"
shows "adjoint (four_block_mat A B C D)
= four_block_mat (adjoint A) (adjoint C) (adjoint B) (adjoint D)"
by (rule eq_matI, insert dim, auto simp add: adjoint_eval)
fun unitary_schur_decomposition :: "complex mat \<Rightarrow> complex list \<Rightarrow> complex mat \<times> complex mat \<times> complex mat" where
"unitary_schur_decomposition A [] = (A, 1\<^sub>m (dim_row A), 1\<^sub>m (dim_row A))"
| "unitary_schur_decomposition A (e # es) = (let
n = dim_row A;
n1 = n - 1;
v' = find_eigenvector A e;
v = vec_normalize v';
ws0 = gram_schmidt n (basis_completion v);
ws = map vec_normalize ws0;
W = mat_of_cols n ws;
W' = corthogonal_inv W;
A' = W' * A * W;
(A1,A2,A0,A3) = split_block A' 1 1;
(B,P,Q) = unitary_schur_decomposition A3 es;
z_row = (0\<^sub>m 1 n1);
z_col = (0\<^sub>m n1 1);
one_1 = 1\<^sub>m 1
in (four_block_mat A1 (A2 * P) A0 B,
W * four_block_mat one_1 z_row z_col P,
four_block_mat one_1 z_row z_col Q * W'))"
theorem unitary_schur_decomposition:
assumes A: "(A::complex mat) \<in> carrier_mat n n"
and c: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
and B: "unitary_schur_decomposition A es = (B,P,Q)"
shows "similar_mat_wit A B P Q \<and> upper_triangular B \<and> diag_mat B = es \<and> unitary P \<and> (Q = adjoint P)"
using assms
proof (induct es arbitrary: n A B P Q)
case Nil
with degree_monic_char_poly[of A n]
show ?case by (auto intro: similar_mat_wit_refl simp: diag_mat_def unitary_zero)
next
case (Cons e es n A C P Q)
let ?n1 = "n - 1"
from Cons have A: "A \<in> carrier_mat n n" and dim: "dim_row A = n" by auto
let ?cp = "char_poly A"
from Cons(3)
have cp: "?cp = [: -e, 1 :] * (\<Prod>e \<leftarrow> es. [:- e, 1:])" by auto
have mon: "monic (\<Prod>e\<leftarrow> es. [:- e, 1:])" by (rule monic_prod_list, auto)
have deg: "degree ?cp = Suc (degree (\<Prod>e\<leftarrow> es. [:- e, 1:]))" unfolding cp
by (subst degree_mult_eq, insert mon, auto)
with degree_monic_char_poly[OF A] have n: "n \<noteq> 0" by auto
define v' where "v' = find_eigenvector A e"
define v where "v = vec_normalize v'"
define b where "b = basis_completion v"
define ws0 where "ws0 = gram_schmidt n b"
define ws where "ws = map vec_normalize ws0"
define W where "W = mat_of_cols n ws"
define W' where "W' = corthogonal_inv W"
define A' where "A' = W' * A * W"
obtain A1 A2 A0 A3 where splitA': "split_block A' 1 1 = (A1,A2,A0,A3)"
by (cases "split_block A' 1 1", auto)
obtain B P' Q' where schur: "unitary_schur_decomposition A3 es = (B,P',Q')"
by (cases "unitary_schur_decomposition A3 es", auto)
let ?P' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) P'"
let ?Q' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) Q'"
have C: "C = four_block_mat A1 (A2 * P') A0 B" and P: "P = W * ?P'" and Q: "Q = ?Q' * W'"
using Cons(4) unfolding unitary_schur_decomposition.simps
Let_def list.sel dim
v'_def[symmetric] v_def[symmetric] b_def[symmetric] ws0_def[symmetric] ws_def[symmetric] W'_def[symmetric] W_def[symmetric]
A'_def[symmetric] split splitA' schur by auto
have e: "eigenvalue A e"
unfolding eigenvalue_root_char_poly[OF A] cp by simp
from find_eigenvector[OF A e] have ev': "eigenvector A v' e" unfolding v'_def .
then have "v' \<in> carrier_vec n" unfolding eigenvector_def using A by auto
with ev' have ev: "eigenvector A v e" unfolding v_def using A dim normalize_keep_eigenvector by auto
from this[unfolded eigenvector_def]
have v[simp]: "v \<in> carrier_vec n" and v0: "v \<noteq> 0\<^sub>v n" using A by auto
interpret cof_vec_space n "TYPE(complex)" .
from basis_completion[OF v v0, folded b_def]
have span_b: "span (set b) = carrier_vec n" and dist_b: "distinct b"
and indep: "\<not> lin_dep (set b)" and b: "set b \<subseteq> carrier_vec n" and hdb: "hd b = v"
and len_b: "length b = n" by auto
from hdb len_b n obtain vs where bv: "b = v # vs" by (cases b, auto)
from gram_schmidt_result[OF b dist_b indep refl, folded ws0_def]
have ws0: "set ws0 \<subseteq> carrier_vec n" "corthogonal ws0" "length ws0 = n"
by (auto simp: len_b)
then have ws: "set ws \<subseteq> carrier_vec n" "corthogonal ws" "length ws = n" unfolding ws_def
using normalize_keep_corthogonal by auto
- have ws0ne: "ws0 \<noteq> []" using `length ws0 = n` n by auto
+ have ws0ne: "ws0 \<noteq> []" using \<open>length ws0 = n\<close> n by auto
from gram_schmidt_hd[OF v, of vs, folded bv] have hdws0: "hd ws0 = (vec_normalize v')" unfolding ws0_def v_def .
have "hd ws = vec_normalize (hd ws0)" unfolding ws_def using hd_map[OF ws0ne] by auto
then have hdws: "hd ws = v" unfolding v_def using normalize_normalize[of v'] hdws0 by auto
have orth_W: "corthogonal_mat W" using orthogonal_mat_of_cols ws unfolding W_def.
have W: "W \<in> carrier_mat n n"
using ws unfolding W_def using mat_of_cols_carrier(1)[of n ws] by auto
have W': "W' \<in> carrier_mat n n" unfolding W'_def corthogonal_inv_def using W
by (auto simp: mat_of_rows_def)
from corthogonal_inv_result[OF orth_W]
have W'W: "inverts_mat W' W" unfolding W'_def .
hence WW': "inverts_mat W W'" using mat_mult_left_right_inverse[OF W' W] W' W
unfolding inverts_mat_def by auto
have A': "A' \<in> carrier_mat n n" using W W' A unfolding A'_def by auto
have A'A_wit: "similar_mat_wit A' A W' W"
by (rule similar_mat_witI[of _ _ n], insert W W' A A' W'W WW', auto simp: A'_def
inverts_mat_def)
hence A'A: "similar_mat A' A" unfolding similar_mat_def by blast
from similar_mat_wit_sym[OF A'A_wit] have simAA': "similar_mat_wit A A' W W'" by auto
have eigen[simp]: "A *\<^sub>v v = e \<cdot>\<^sub>v v" and v0: "v \<noteq> 0\<^sub>v n"
using v_def v'_def find_eigenvector[OF A e] A normalize_keep_eigenvector
unfolding eigenvector_def by auto
let ?f = "(\<lambda> i. if i = 0 then e else 0)"
have col0: "col A' 0 = vec n ?f"
unfolding A'_def W'_def W_def
using corthogonal_col_ev_0[OF A v v0 eigen n hdws ws].
from A' n have "dim_row A' = 1 + ?n1" "dim_col A' = 1 + ?n1" by auto
from split_block[OF splitA' this] have A2: "A2 \<in> carrier_mat 1 ?n1"
and A3: "A3 \<in> carrier_mat ?n1 ?n1"
and A'block: "A' = four_block_mat A1 A2 A0 A3" by auto
have A1id: "A1 = mat 1 1 (\<lambda> _. e)"
using splitA'[unfolded split_block_def Let_def] arg_cong[OF col0, of "\<lambda> v. v $ 0"] A' n
by (auto simp: col_def)
have A1: "A1 \<in> carrier_mat 1 1" unfolding A1id by auto
{
fix i
assume "i < ?n1"
with arg_cong[OF col0, of "\<lambda> v. v $ Suc i"] A'
have "A' $$ (Suc i, 0) = 0" by auto
} note A'0 = this
have A0id: "A0 = 0\<^sub>m ?n1 1"
using splitA'[unfolded split_block_def Let_def] A'0 A' by auto
have A0: "A0 \<in> carrier_mat ?n1 1" unfolding A0id by auto
from cp char_poly_similar[OF A'A]
have cp: "char_poly A' = [: -e,1 :] * (\<Prod> e \<leftarrow> es. [:- e, 1:])" by simp
also have "char_poly A' = char_poly A1 * char_poly A3"
unfolding A'block A0id
by (rule char_poly_four_block_zeros_col[OF A1 A2 A3])
also have "char_poly A1 = [: -e,1 :]"
by (simp add: A1id char_poly_defs det_def signof_def sign_def)
finally have cp: "char_poly A3 = (\<Prod> e \<leftarrow> es. [:- e, 1:])"
by (metis mult_cancel_left pCons_eq_0_iff zero_neq_one)
from Cons(1)[OF A3 cp schur]
have simIH: "similar_mat_wit A3 B P' Q'" and ut: "upper_triangular B" and diag: "diag_mat B = es"
and uP': "unitary P'" and Q'P': "Q' = adjoint P'"
by auto
from similar_mat_witD2[OF A3 simIH]
have B: "B \<in> carrier_mat ?n1 ?n1" and P': "P' \<in> carrier_mat ?n1 ?n1" and Q': "Q' \<in> carrier_mat ?n1 ?n1"
and PQ': "P' * Q' = 1\<^sub>m ?n1" by auto
have A0_eq: "A0 = P' * A0 * 1\<^sub>m 1" unfolding A0id using P' by auto
have simA'C: "similar_mat_wit A' C ?P' ?Q'" unfolding A'block C
by (rule similar_mat_wit_four_block[OF similar_mat_wit_refl[OF A1] simIH _ A0_eq A1 A3 A0],
insert PQ' A2 P' Q', auto)
have ut1: "upper_triangular A1" unfolding A1id by auto
have ut: "upper_triangular C" unfolding C A0id
by (intro upper_triangular_four_block[OF _ B ut1 ut], auto simp: A1id)
from A1id have diagA1: "diag_mat A1 = [e]" unfolding diag_mat_def by auto
from diag_four_block_mat[OF A1 B] have diag: "diag_mat C = e # es" unfolding diag diagA1 C by simp
have aW: "adjoint W \<in> carrier_mat n n" using W by auto
have aW': "adjoint W' \<in> carrier_mat n n" using W' by auto
have "unitary W" using W_def ws_def ws0 normalized_corthogonal_mat_is_unitary by auto
then have ivWaW: "inverts_mat W (adjoint W)" using unitary_def W aW by auto
with WW' have W'aW: "W' = (adjoint W)" using inverts_mat_unique W W' aW by auto
then have "adjoint W' = W" using adjoint_adjoint by auto
with ivWaW have "inverts_mat W' (adjoint W')" using inverts_mat_symm W aW W'aW by auto
then have "unitary W'" using unitary_def W' by auto
have newP': "P' \<in> carrier_mat (n - Suc 0) (n - Suc 0)" using P' by auto
have rl: "\<And> x1 x2 x3 x4 y1 y2 y3 y4 f. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow> f x1 x2 x3 x4 = f y1 y2 y3 y4" by simp
have Q'aP': "?Q' = adjoint ?P'"
apply (subst four_block_mat_adjoint, auto simp add: newP')
apply (rule rl[where f2 = four_block_mat])
apply (auto simp add: eq_matI adjoint_eval Q'P')
done
have "adjoint P = adjoint ?P' * adjoint W" using W newP' n
apply (simp add: P)
apply (subst adjoint_mult[of W, symmetric])
apply (auto simp add: W P' carrier_matD[of W n n])
done
also have "\<dots> = ?Q' * W'" using Q'aP' W'aW by auto
also have "\<dots> = Q" using Q by auto
finally have QaP: "Q = adjoint P" ..
from similar_mat_wit_trans[OF simAA' simA'C, folded P Q] have smw: "similar_mat_wit A C P Q" by blast
then have dimP: "P \<in> carrier_mat n n" and dimQ: "Q \<in> carrier_mat n n" unfolding similar_mat_wit_def using A by auto
from smw have "P * Q = 1\<^sub>m n" unfolding similar_mat_wit_def using A by auto
then have "inverts_mat P Q" using inverts_mat_def dimP by auto
then have uP: "unitary P" using QaP unitary_def dimP by auto
from ut similar_mat_wit_trans[OF simAA' simA'C, folded P Q] diag uP QaP
show ?case by blast
qed
lemma complex_mat_char_poly_factorizable:
fixes A :: "complex mat"
assumes "A \<in> carrier_mat n n"
shows "\<exists>as. char_poly A = (\<Prod> a \<leftarrow> as. [:- a, 1:]) \<and> length as = n"
proof -
let ?ca = "char_poly A"
have ex0: "\<exists>bs. Polynomial.smult (lead_coeff ?ca) (\<Prod>b\<leftarrow>bs. [:- b, 1:]) = ?ca \<and>
length bs = degree ?ca"
by (simp add: fundamental_theorem_algebra_factorized)
then obtain bs where " Polynomial.smult (lead_coeff ?ca) (\<Prod>b\<leftarrow>bs. [:- b, 1:]) = ?ca \<and>
length bs = degree ?ca" by auto
moreover have "lead_coeff ?ca = (1::complex)"
using assms degree_monic_char_poly by blast
ultimately have ex1: "?ca = (\<Prod>b\<leftarrow>bs. [:- b, 1:]) \<and> length bs = degree ?ca" by auto
moreover have "degree ?ca = n"
by (simp add: assms degree_monic_char_poly)
ultimately show ?thesis by auto
qed
lemma complex_mat_has_unitary_schur_decomposition:
fixes A :: "complex mat"
assumes "A \<in> carrier_mat n n"
shows "\<exists>B P es. similar_mat_wit A B P (adjoint P) \<and> unitary P
\<and> char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:]) \<and> diag_mat B = es"
proof -
have "\<exists>es. char_poly A = (\<Prod> e \<leftarrow> es. [:- e, 1:]) \<and> length es = n"
using assms by (simp add: complex_mat_char_poly_factorizable)
then obtain es where es: "char_poly A = (\<Prod> e \<leftarrow> es. [:- e, 1:]) \<and> length es = n" by auto
obtain B P Q where B: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto)
have "similar_mat_wit A B P Q \<and> upper_triangular B \<and> unitary P \<and> (Q = adjoint P) \<and>
char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:]) \<and> diag_mat B = es" using assms es B
by (auto simp add: unitary_schur_decomposition)
then show ?thesis by auto
qed
lemma normal_upper_triangular_matrix_is_diagonal:
fixes A :: "'a::conjugatable_ordered_field mat"
assumes "A \<in> carrier_mat n n"
and tri: "upper_triangular A"
and norm: "A * adjoint A = adjoint A * A"
shows "diagonal_mat A"
proof (rule disjE[of "n = 0" "n > 0"], blast)
have dim: "dim_row A = n" "dim_col A = n" using assms by auto
from norm have eq0: "\<And>i j. (A * adjoint A)$$(i,j) = (adjoint A * A)$$(i,j)" by auto
have nat_induct_strong:
"\<And>P. (P::nat\<Rightarrow>bool) 0 \<Longrightarrow> (\<And>i. i < n \<Longrightarrow> (\<And>k. k < i \<Longrightarrow> P k) \<Longrightarrow> P i) \<Longrightarrow> (\<And>i. i < n \<Longrightarrow> P i)"
by (metis dual_order.strict_trans infinite_descent0 linorder_neqE_nat)
show "n = 0 \<Longrightarrow> ?thesis" using dim unfolding diagonal_mat_def by auto
show "n > 0 \<Longrightarrow> ?thesis" unfolding diagonal_mat_def dim
apply (rule allI, rule impI)
apply (rule nat_induct_strong)
proof (rule allI, rule impI, rule impI)
assume asm: "n > 0"
from tri upper_triangularD[of A 0 j] dim have z0: "\<And>j. 0< j \<Longrightarrow> j < n \<Longrightarrow> A$$(j, 0) = 0"
by auto
then have ada00: "(adjoint A * A)$$(0,0) = conjugate (A$$(0,0)) * A$$(0,0)"
using asm dim by (auto simp add: scalar_prod_def adjoint_eval sum.atLeast_Suc_lessThan)
have aad00: "(A * adjoint A)$$(0,0) = (\<Sum>k=0..<n. A$$(0, k) * conjugate (A$$(0, k)))"
using asm dim by (auto simp add: scalar_prod_def adjoint_eval)
moreover have
"\<dots> = A$$(0,0) * conjugate (A$$(0,0))
+ (\<Sum>k=1..<n. A$$(0, k) * conjugate (A$$(0, k)))"
using dim asm by (subst sum.atLeast_Suc_lessThan[of 0 n "\<lambda>k. A$$(0, k) * conjugate (A$$(0, k))"], auto)
ultimately have f1tneq0: "(\<Sum>k=(Suc 0)..<n. A$$(0, k) * conjugate (A$$(0, k))) = 0"
using eq0 ada00 by (simp)
have geq0: "\<And>k. k < n \<Longrightarrow> A$$(0, k) * conjugate (A$$(0, k)) \<ge> 0"
using conjugate_square_positive by auto
have "\<And>k. 1 \<le> k \<Longrightarrow> k < n \<Longrightarrow> A$$(0, k) * conjugate (A$$(0, k)) = 0"
by (rule sum_nonneg_0[of "{1..<n}"], auto, rule geq0, auto, rule f1tneq0)
with dim asm show
case0: "\<And>j. 0 < n \<Longrightarrow> j < n \<Longrightarrow> 0 \<noteq> j \<Longrightarrow> A $$ (0, j) = 0"
by auto
{
fix i
assume asm: "n > 0" "i < n" "i > 0"
and ih: "\<And>k. k < i \<Longrightarrow> \<forall>j<n. k \<noteq> j \<longrightarrow> A $$ (k, j) = 0"
then have "\<And>j. j<n \<Longrightarrow> i \<noteq> j \<Longrightarrow> A $$ (i, j) = 0"
proof -
have inter_part: "\<And>b m e. (b::nat) < e \<Longrightarrow> b < m \<Longrightarrow> m < e \<Longrightarrow> {b..<m} \<union> {m..<e} = {b..<e}" by auto
then have
"\<And>b m e f. (b::nat) < e \<Longrightarrow> b < m \<Longrightarrow> m < e
\<Longrightarrow> (\<Sum>k=b..<e. f k) = (\<Sum>k\<in>{b..<m}\<union>{m..<e}. f k)"
using sum.union_disjoint by auto
then have sum_part:
"\<And>b m e f. (b::nat) < e \<Longrightarrow> b < m \<Longrightarrow> m < e
\<Longrightarrow> (\<Sum>k=b..<e. f k) = (\<Sum>k=b..<m. f k) + (\<Sum>k=m..<e. f k)"
by (auto simp add: sum.union_disjoint)
from tri upper_triangularD[of A j i] asm dim have
zsi0: "\<And>j. j < i \<Longrightarrow> A$$(i, j) = 0" by auto
from tri upper_triangularD[of A j i] asm dim have
zsi1: "\<And>k. i < k \<Longrightarrow> k < n \<Longrightarrow> A$$(k, i) = 0" by auto
have
"(A * adjoint A)$$(i, i)
= (\<Sum>k=0..<n. conjugate (A$$(i, k)) * A$$(i, k))" using asm dim
apply (auto simp add: scalar_prod_def adjoint_eval)
apply (rule sum.cong, auto)
done
also have
"\<dots> = (\<Sum>k=0..<i. conjugate (A$$(i, k)) * A$$(i, k))
+ (\<Sum>k=i..<n. conjugate (A$$(i, k)) * A$$(i, k))"
using asm
by (auto simp add: sum_part[of 0 n i])
also have
"\<dots> = (\<Sum>k=i..<n. conjugate (A$$(i, k)) * A$$(i, k))"
using zsi0
by auto
also have
"\<dots> = conjugate (A$$(i, i)) * A$$(i, i)
+ (\<Sum>k=(Suc i)..<n. conjugate (A$$(i, k)) * A$$(i, k))"
using asm
by (auto simp add: sum.atLeast_Suc_lessThan)
finally have
adaii: "(A * adjoint A)$$(i, i)
= conjugate (A$$(i, i)) * A$$(i, i)
+ (\<Sum>k=(Suc i)..<n. conjugate (A$$(i, k)) * A$$(i, k))" .
have
"(adjoint A * A)$$(i, i) = (\<Sum>k=0..<n. conjugate (A$$(k, i)) * A$$(k, i))"
using asm dim by (auto simp add: scalar_prod_def adjoint_eval)
also have
"\<dots> = (\<Sum>k=0..<i. conjugate (A$$(k, i)) * A$$(k, i))
+ (\<Sum>k=i..<n. conjugate (A$$(k, i)) * A$$(k, i))"
using asm by (auto simp add: sum_part[of 0 n i])
also have
"\<dots> = (\<Sum>k=i..<n. conjugate (A$$(k, i)) * A$$(k, i))"
using asm ih by auto
also have
"\<dots> = conjugate (A$$(i, i)) * A$$(i, i)"
using asm zsi1 by (auto simp add: sum.atLeast_Suc_lessThan)
finally have "(adjoint A * A)$$(i, i) = conjugate (A$$(i, i)) * A$$(i, i)" .
with adaii eq0 have
fsitoneq0: "(\<Sum>k=(Suc i)..<n. conjugate (A$$(i, k)) * A$$(i, k)) = 0" by auto
have "\<And>k. k<n \<Longrightarrow> i < k \<Longrightarrow> conjugate (A$$(i, k)) * A$$(i, k) = 0"
by (rule sum_nonneg_0[of "{(Suc i)..<n}"], auto, subst mult.commute,
rule conjugate_square_positive, rule fsitoneq0)
then have "\<And>k. k<n \<Longrightarrow> i<k \<Longrightarrow> A $$ (i, k) = 0" by auto
with zsi0 show "\<And>j. j<n \<Longrightarrow> i \<noteq> j \<Longrightarrow> A $$ (i, j) = 0"
by (metis linorder_neqE_nat)
qed
}
with case0 show "\<And>i ia.
0 < n \<Longrightarrow>
i < n \<Longrightarrow>
ia < n \<Longrightarrow>
(\<And>k. k < ia \<Longrightarrow> \<forall>j<n. k \<noteq> j \<longrightarrow> A $$ (k, j) = 0) \<Longrightarrow>
\<forall>j<n. ia \<noteq> j \<longrightarrow> A $$ (ia, j) = 0" by auto
qed
qed
lemma normal_complex_mat_has_spectral_decomposition:
assumes A: "(A::complex mat) \<in> carrier_mat n n"
and normal: "A * adjoint A = adjoint A * A"
and c: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
and B: "unitary_schur_decomposition A es = (B,P,Q)"
shows "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es \<and> unitary P"
proof -
have smw: "similar_mat_wit A B P (adjoint P)"
and ut: "upper_triangular B"
and uP: "unitary P"
and dB: "diag_mat B = es"
and "(Q = adjoint P)"
using assms by (auto simp add: unitary_schur_decomposition)
from smw have dimP: "P \<in> carrier_mat n n" and dimB: "B \<in> carrier_mat n n"
and dimaP: "adjoint P \<in> carrier_mat n n"
unfolding similar_mat_wit_def using A by auto
have dimaB: "adjoint B \<in> carrier_mat n n" using dimB by auto
note dims = dimP dimB dimaP dimaB
have "inverts_mat P (adjoint P)" using unitary_def uP dims by auto
then have iaPP: "inverts_mat (adjoint P) P" using inverts_mat_symm using dims by auto
have aPP: "adjoint P * P = 1\<^sub>m n" using dims iaPP unfolding inverts_mat_def by auto
from smw have A: "A = P * B * (adjoint P)" unfolding similar_mat_wit_def Let_def by auto
then have aA: "adjoint A = P * adjoint B * adjoint P"
by (insert A dimP dimB dimaP, auto simp add: adjoint_mult[of _ n n _ n] adjoint_adjoint)
have "A * adjoint A = (P * B * adjoint P) * (P * adjoint B * adjoint P)" using A aA by auto
also have "\<dots> = P * B * (adjoint P * P) * (adjoint B * adjoint P)" using dims by (mat_assoc n)
also have "\<dots> = P * B * 1\<^sub>m n * (adjoint B * adjoint P)" using dims aPP by (auto)
also have "\<dots> = P * B * adjoint B * adjoint P" using dims by (mat_assoc n)
finally have "A * adjoint A = P * B * adjoint B * adjoint P".
then have "adjoint P * (A * adjoint A) * P = (adjoint P * P) * B * adjoint B * (adjoint P * P)"
using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n])
also have "\<dots> = 1\<^sub>m n * B * adjoint B * 1\<^sub>m n" using aPP by auto
also have "\<dots> = B * adjoint B" using dims by auto
finally have eq0: "adjoint P * (A * adjoint A) * P = B * adjoint B".
have "adjoint A * A = (P * adjoint B * adjoint P) * (P * B * adjoint P)" using A aA by auto
also have "\<dots> = P * adjoint B * (adjoint P * P) * (B * adjoint P)" using dims by (mat_assoc n)
also have "\<dots> = P * adjoint B * 1\<^sub>m n * (B * adjoint P)" using dims aPP by (auto)
also have "\<dots> = P * adjoint B * B * adjoint P" using dims by (mat_assoc n)
finally have "adjoint A * A = P * adjoint B * B * adjoint P" by auto
then have "adjoint P * (adjoint A * A) * P = (adjoint P * P) * adjoint B * B * (adjoint P * P)"
using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n])
also have "\<dots> = 1\<^sub>m n * adjoint B * B * 1\<^sub>m n" using aPP by auto
also have "\<dots> = adjoint B * B" using dims by auto
finally have eq1: "adjoint P * (adjoint A * A) * P = adjoint B * B".
from normal have "adjoint P * (adjoint A * A) * P = adjoint P * (A * adjoint A) * P" by auto
with eq0 eq1 have "B * adjoint B = adjoint B * B" by auto
with ut dims have "diagonal_mat B" using normal_upper_triangular_matrix_is_diagonal by auto
with smw uP dB show "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es \<and> unitary P" by auto
qed
lemma complex_mat_has_jordan_nf:
fixes A :: "complex mat"
assumes "A \<in> carrier_mat n n"
shows "\<exists>n_as. jordan_nf A n_as"
proof -
have "\<exists>as. char_poly A = (\<Prod> a \<leftarrow> as. [:- a, 1:]) \<and> length as = n"
using assms by (simp add: complex_mat_char_poly_factorizable)
then show ?thesis using assms
by (auto simp add: jordan_nf_iff_linear_factorization)
qed
lemma hermitian_is_normal:
assumes "hermitian A"
shows "A * adjoint A = adjoint A * A"
using assms by (auto simp add: hermitian_def)
lemma hermitian_eigenvalue_real:
assumes dim: "(A::complex mat) \<in> carrier_mat n n"
and hA: "hermitian A"
and c: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
and B: "unitary_schur_decomposition A es = (B,P,Q)"
shows "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es
\<and> unitary P \<and> (\<forall>i < n. B$$(i, i) \<in> Reals)"
proof -
have normal: "A * adjoint A = adjoint A * A" using hA hermitian_is_normal by auto
then have schur: "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es \<and> unitary P"
using normal_complex_mat_has_spectral_decomposition[OF dim normal c B] by (simp)
then have "similar_mat_wit A B P (adjoint P)"
and uP: "unitary P" and dB: "diag_mat B = es"
using assms by auto
then have A: "A = P * B * (adjoint P)"
and dimB: "B \<in> carrier_mat n n" and dimP: "P \<in> carrier_mat n n"
unfolding similar_mat_wit_def Let_def using dim by auto
then have dimaB: "adjoint B \<in> carrier_mat n n" by auto
have "adjoint A = adjoint (adjoint P) * adjoint (P * B)"
apply (subst A)
apply (subst adjoint_mult[of "P * B" n n "adjoint P" n])
apply (insert dimB dimP, auto)
done
also have "\<dots> = P * adjoint (P * B)" by (auto simp add: adjoint_adjoint)
also have "\<dots> = P * (adjoint B * adjoint P)" using dimB dimP by (auto simp add: adjoint_mult)
also have "\<dots> = P * adjoint B * adjoint P" using dimB dimP by (subst assoc_mult_mat[symmetric, of P n n "adjoint B" n "adjoint P" n], auto)
finally have aA: "adjoint A = P * adjoint B * adjoint P" .
have "A = adjoint A" using hA hermitian_def[of A] by auto
then have "P * B * adjoint P = P * adjoint B * adjoint P" using A aA by auto
then have BaB: "B = adjoint B" using unitary_elim[OF dimB dimaB dimP] uP by auto
{
fix i
assume "i < n"
then have "B$$(i, i) = conjugate (B$$(i, i))"
apply (subst BaB)
by (insert dimB, simp add: adjoint_eval)
then have "B$$(i, i) \<in> Reals" unfolding conjugate_complex_def
using Reals_cnj_iff by auto
}
then have "\<forall>i<n. B$$(i, i) \<in> Reals" by auto
with schur show ?thesis by auto
qed
lemma hermitian_inner_prod_real:
assumes dimA: "(A::complex mat) \<in> carrier_mat n n"
and dimv: "v \<in> carrier_vec n"
and hA: "hermitian A"
shows "inner_prod v (A *\<^sub>v v) \<in> Reals"
proof -
obtain es where es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
using complex_mat_char_poly_factorizable dimA by auto
obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)"
by (cases "unitary_schur_decomposition A es", auto)
then have "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es
\<and> unitary P \<and> (\<forall>i < n. B$$(i, i) \<in> Reals)"
using hermitian_eigenvalue_real dimA es hA by auto
then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B"
and Bii: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<in> Reals"
and dimB: "B \<in> carrier_mat n n" and dimP: "P \<in> carrier_mat n n" and dimaP: "adjoint P \<in> carrier_mat n n"
unfolding similar_mat_wit_def Let_def using dimA by auto
define w where "w = (adjoint P) *\<^sub>v v"
then have dimw: "w \<in> carrier_vec n" using dimaP dimv by auto
from A have "inner_prod v (A *\<^sub>v v) = inner_prod v ((P * B * (adjoint P)) *\<^sub>v v)" by auto
also have "\<dots> = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv
by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto)
also have "\<dots> = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP
by (subst assoc_mult_mat_vec[of _ n n "B" n], auto)
also have "\<dots> = inner_prod w (B *\<^sub>v w)" unfolding w_def
apply (rule adjoint_def_alter[OF _ _ dimP])
apply (insert mult_mat_vec_carrier[OF dimB mult_mat_vec_carrier[OF dimaP dimv]], auto simp add: dimv)
done
also have "\<dots> = (\<Sum>i=0..<n. (\<Sum>j=0..<n.
conjugate (w$i) * B$$(i, j) * w$j))" unfolding scalar_prod_def using dimw dimB
apply (simp add: scalar_prod_def sum_distrib_right)
apply (rule sum.cong, auto, rule sum.cong, auto)
done
also have "\<dots> = (\<Sum>i=0..<n. B$$(i, i) * conjugate (w$i) * w$i)"
apply (rule sum.cong, auto)
apply (simp add: sum.remove)
apply (insert dB[unfolded diagonal_mat_def] dimB, auto)
done
finally have sum: "inner_prod v (A *\<^sub>v v) = (\<Sum>i=0..<n. B$$(i, i) * conjugate (w$i) * w$i)" .
have "\<And>i. i < n \<Longrightarrow> B$$(i, i) * conjugate (w$i) * w$i \<in> Reals" using Bii by (simp add: Reals_cnj_iff)
then have "(\<Sum>i=0..<n. B$$(i, i) * conjugate (w$i) * w$i) \<in> Reals" by auto
then show ?thesis using sum by auto
qed
lemma unit_vec_bracket:
fixes A :: "complex mat"
assumes dimA: "A \<in> carrier_mat n n" and i: "i < n"
shows "inner_prod (unit_vec n i) (A *\<^sub>v (unit_vec n i)) = A$$(i, i)"
proof -
define w where "(w::complex vec) = unit_vec n i"
have "A *\<^sub>v w = col A i" using i dimA w_def by auto
then have 1: "inner_prod w (A *\<^sub>v w) = inner_prod w (col A i)" using w_def by auto
have "conjugate w = w" unfolding w_def unit_vec_def conjugate_vec_def using i by auto
then have 2: "inner_prod w (col A i) = A$$(i, i)" using i dimA w_def by auto
from 1 2 show "inner_prod w (A *\<^sub>v w) = A$$(i, i)" by auto
qed
lemma spectral_decomposition_extract_diag:
fixes P B :: "complex mat"
assumes dimP: "P \<in> carrier_mat n n" and dimB: "B \<in> carrier_mat n n"
and uP: "unitary P" and dB: "diagonal_mat B" and i: "i < n"
shows "inner_prod (col P i) (P * B * (adjoint P) *\<^sub>v (col P i)) = B$$(i, i)"
proof -
have dimaP: "adjoint P\<in> carrier_mat n n" using dimP by auto
have uaP: "unitary (adjoint P)" using unitary_adjoint uP dimP by auto
then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint)
then have iv: "(adjoint P) * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto
define v where "v = col P i"
then have dimv: "v \<in> carrier_vec n" using dimP by auto
define w where "(w::complex vec) = unit_vec n i"
then have dimw: "w \<in> carrier_vec n" by auto
have BaPv: "B *\<^sub>v (adjoint P *\<^sub>v v) \<in> carrier_vec n" using dimB dimaP dimv by auto
have "(adjoint P) *\<^sub>v v = (col (adjoint P * P) i)"
by (simp add: col_mult2[OF dimaP dimP i, symmetric] v_def)
then have aPv: "(adjoint P) *\<^sub>v v = w"
by (auto simp add: iv i w_def)
have "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv
by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto)
also have "\<dots> = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP
by (subst assoc_mult_mat_vec[of _ n n "B" n], auto)
also have "\<dots> = inner_prod (adjoint P *\<^sub>v v) (B *\<^sub>v (adjoint P *\<^sub>v v))"
by (simp add: adjoint_def_alter[OF dimv BaPv dimP])
also have "\<dots> = inner_prod w (B *\<^sub>v w)" using aPv by auto
also have "\<dots> = B$$(i, i)" using w_def unit_vec_bracket dimB i by auto
finally show "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = B$$(i, i)".
qed
lemma hermitian_inner_prod_zero:
fixes A :: "complex mat"
assumes dimA: "A \<in> carrier_mat n n" and hA: "hermitian A"
and zero: "\<forall>v\<in>carrier_vec n. inner_prod v (A *\<^sub>v v) = 0"
shows "A = 0\<^sub>m n n"
proof -
obtain es where es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
using complex_mat_char_poly_factorizable dimA by auto
obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)"
by (cases "unitary_schur_decomposition A es", auto)
then have "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es
\<and> unitary P \<and> (\<forall>i < n. B$$(i, i) \<in> Reals)"
using hermitian_eigenvalue_real dimA es hA by auto
then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B"
and Bii: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<in> Reals"
and dimB: "B \<in> carrier_mat n n" and dimP: "P \<in> carrier_mat n n" and dimaP: "adjoint P \<in> carrier_mat n n"
and uP: "unitary P"
unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto
then have uaP: "unitary (adjoint P)" using unitary_adjoint by auto
then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint)
then have iv: "adjoint P * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto
have "B = 0\<^sub>m n n"
proof-
{
fix i assume i: "i < n"
define v where "v = col P i"
then have dimv: "v \<in> carrier_vec n" using v_def dimP by auto
have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def
using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto
moreover have "inner_prod v (A *\<^sub>v v) = 0" using dimv zero by auto
ultimately have "B$$(i, i) = 0" by auto
}
note zB = this
show "B = 0\<^sub>m n n" by (insert zB dB dimB, rule eq_matI, auto simp add: diagonal_mat_def)
qed
then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto
qed
lemma complex_mat_decomposition_to_hermitian:
fixes A :: "complex mat"
assumes dim: "A \<in> carrier_mat n n"
shows "\<exists>B C. hermitian B \<and> hermitian C \<and> A = B + \<i> \<cdot>\<^sub>m C \<and> B \<in> carrier_mat n n \<and> C \<in> carrier_mat n n"
proof -
obtain B C where B: "B = (1 / 2) \<cdot>\<^sub>m (A + adjoint A)"
and C: "C = (-\<i> / 2) \<cdot>\<^sub>m (A - adjoint A)" by auto
then have dimB: "B \<in> carrier_mat n n" and dimC: "C \<in> carrier_mat n n" using dim by auto
have "hermitian B" unfolding B hermitian_def using dim
by (auto simp add: adjoint_eval)
moreover have "hermitian C" unfolding C hermitian_def using dim
apply (subst eq_matI)
apply (auto simp add: adjoint_eval algebra_simps)
done
moreover have "A = B + \<i> \<cdot>\<^sub>m C" using dim B C
apply (subst eq_matI)
apply (auto simp add: adjoint_eval algebra_simps)
done
ultimately show ?thesis using dimB dimC by auto
qed
subsection \<open>Outer product\<close>
definition outer_prod :: "'a::conjugatable_field vec \<Rightarrow> 'a vec \<Rightarrow> 'a mat" where
"outer_prod v w = mat (dim_vec v) 1 (\<lambda>(i, j). v $ i) * mat 1 (dim_vec w) (\<lambda>(i, j). (conjugate w) $ j)"
lemma outer_prod_dim[simp]:
fixes v w :: "'a::conjugatable_field vec"
assumes v: "v \<in> carrier_vec n" and w: "w \<in> carrier_vec m"
shows "outer_prod v w \<in> carrier_mat n m"
unfolding outer_prod_def using assms mat_of_cols_carrier mat_of_rows_carrier by auto
lemma mat_of_vec_mult_eq_scalar_prod:
fixes v w :: "'a::conjugatable_field vec"
assumes "v \<in> carrier_vec n" and "w \<in> carrier_vec n"
shows "mat 1 (dim_vec v) (\<lambda>(i, j). (conjugate v) $ j) * mat (dim_vec w) 1 (\<lambda>(i, j). w $ i)
= mat 1 1 (\<lambda>k. inner_prod v w)"
apply (rule eq_matI) using assms apply (simp add: scalar_prod_def) apply (rule sum.cong) by auto
lemma one_dim_mat_mult_is_scale:
fixes A B :: "('a::conjugatable_field mat)"
assumes "B \<in> carrier_mat 1 n"
shows "(mat 1 1 (\<lambda>k. a)) * B = a \<cdot>\<^sub>m B"
apply (rule eq_matI) using assms by (auto simp add: scalar_prod_def)
lemma outer_prod_mult_outer_prod:
fixes a b c d :: "'a::conjugatable_field vec"
assumes a: "a \<in> carrier_vec d1" and b: "b \<in> carrier_vec d2"
and c: "c \<in> carrier_vec d2" and d: "d \<in> carrier_vec d3"
shows "outer_prod a b * outer_prod c d = inner_prod b c \<cdot>\<^sub>m outer_prod a d"
proof -
let ?ma = "mat (dim_vec a) 1 (\<lambda>(i, j). a $ i)"
let ?mb = "mat 1 (dim_vec b) (\<lambda>(i, j). (conjugate b) $ j)"
let ?mc = "mat (dim_vec c) 1 (\<lambda>(i, j). c $ i)"
let ?md = "mat 1 (dim_vec d) (\<lambda>(i, j). (conjugate d) $ j)"
have "(?ma * ?mb) * (?mc * ?md) = ?ma * (?mb * (?mc * ?md))"
apply (subst assoc_mult_mat[of "?ma" d1 1 "?mb" d2 "?mc * ?md" d3] )
using assms by auto
also have "\<dots> = ?ma * ((?mb * ?mc) * ?md)"
apply (subst assoc_mult_mat[symmetric, of "?mb" 1 d2 "?mc" 1 "?md" d3])
using assms by auto
also have "\<dots> = ?ma * ((mat 1 1 (\<lambda>k. inner_prod b c)) * ?md)"
apply (subst mat_of_vec_mult_eq_scalar_prod[of b d2 c]) using assms by auto
also have "\<dots> = ?ma * (inner_prod b c \<cdot>\<^sub>m ?md)"
apply (subst one_dim_mat_mult_is_scale) using assms by auto
also have "\<dots> = (inner_prod b c) \<cdot>\<^sub>m (?ma * ?md)" using assms by auto
finally show ?thesis unfolding outer_prod_def by auto
qed
lemma index_outer_prod:
fixes v w :: "'a::conjugatable_field vec"
assumes v: "v \<in> carrier_vec n" and w: "w \<in> carrier_vec m"
and ij: "i < n" "j < m"
shows "(outer_prod v w)$$(i, j) = v $ i * conjugate (w $ j)"
unfolding outer_prod_def using assms by (simp add: scalar_prod_def)
lemma mat_of_vec_mult_vec:
fixes a b c :: "'a::conjugatable_field vec"
assumes a: "a \<in> carrier_vec d" and b: "b \<in> carrier_vec d"
shows "mat 1 d (\<lambda>(i, j). (conjugate a) $ j) *\<^sub>v b = vec 1 (\<lambda>k. inner_prod a b)"
apply (rule eq_vecI)
apply (simp add: scalar_prod_def carrier_vecD[OF a] carrier_vecD[OF b])
apply (rule sum.cong) by auto
lemma mat_of_vec_mult_one_dim_vec:
fixes a b :: "'a::conjugatable_field vec"
assumes a: "a \<in> carrier_vec d"
shows "mat d 1 (\<lambda>(i, j). a $ i) *\<^sub>v vec 1 (\<lambda>k. c) = c \<cdot>\<^sub>v a"
apply (rule eq_vecI)
by (auto simp add: scalar_prod_def carrier_vecD[OF a])
lemma outer_prod_mult_vec:
fixes a b c :: "'a::conjugatable_field vec"
assumes a: "a \<in> carrier_vec d1" and b: "b \<in> carrier_vec d2"
and c: "c \<in> carrier_vec d2"
shows "outer_prod a b *\<^sub>v c = inner_prod b c \<cdot>\<^sub>v a"
proof -
have "outer_prod a b *\<^sub>v c
= mat d1 1 (\<lambda>(i, j). a $ i)
* mat 1 d2 (\<lambda>(i, j). (conjugate b) $ j)
*\<^sub>v c" unfolding outer_prod_def using assms by auto
also have "\<dots> = mat d1 1 (\<lambda>(i, j). a $ i)
*\<^sub>v (mat 1 d2 (\<lambda>(i, j). (conjugate b) $ j)
*\<^sub>v c)" apply (subst assoc_mult_mat_vec) using assms by auto
also have "\<dots> = mat d1 1 (\<lambda>(i, j). a $ i)
*\<^sub>v vec 1 (\<lambda>k. inner_prod b c)" using mat_of_vec_mult_vec[of b] assms by auto
also have "\<dots> = inner_prod b c \<cdot>\<^sub>v a" using mat_of_vec_mult_one_dim_vec assms by auto
finally show ?thesis by auto
qed
lemma trace_outer_prod_right:
fixes A :: "'a::conjugatable_field mat" and v w :: "'a vec"
assumes A: "A \<in> carrier_mat n n"
and v: "v \<in> carrier_vec n" and w: "w \<in> carrier_vec n"
shows "trace (A * outer_prod v w) = inner_prod w (A *\<^sub>v v)" (is "?lhs = ?rhs")
proof -
define B where "B = outer_prod v w"
then have B: "B \<in> carrier_mat n n" using assms by auto
have "trace(A * B) = (\<Sum>i = 0..<n. \<Sum>j = 0..<n. A $$ (i,j) * B $$ (j,i))"
unfolding trace_def using A B by (simp add: scalar_prod_def)
also have "\<dots> = (\<Sum>i = 0..<n. \<Sum>j = 0..<n. A $$ (i,j) * v $ j * conjugate (w $ i))"
unfolding B_def
apply (rule sum.cong, simp, rule sum.cong, simp)
by (insert v w, auto simp add: index_outer_prod)
finally have "?lhs = (\<Sum>i = 0..<n. \<Sum>j = 0..<n. A $$ (i,j) * v $ j * conjugate (w $ i))" using B_def by auto
moreover have "?rhs = (\<Sum>i = 0..<n. \<Sum>j = 0..<n. A $$ (i,j) * v $ j * conjugate (w $ i))" using A v w
by (simp add: scalar_prod_def sum_distrib_right)
ultimately show ?thesis by auto
qed
lemma trace_outer_prod:
fixes v w :: "('a::conjugatable_field vec)"
assumes v: "v \<in> carrier_vec n" and w: "w \<in> carrier_vec n"
shows "trace (outer_prod v w) = inner_prod w v" (is "?lhs = ?rhs")
proof -
have "(1\<^sub>m n) * (outer_prod v w) = outer_prod v w" apply (subst left_mult_one_mat) using outer_prod_dim assms by auto
moreover have "1\<^sub>m n *\<^sub>v v = v" using assms by auto
ultimately show ?thesis using trace_outer_prod_right[of "1\<^sub>m n" n v w] assms by auto
qed
lemma inner_prod_outer_prod:
fixes a b c d :: "'a::conjugatable_field vec"
assumes a: "a \<in> carrier_vec n" and b: "b \<in> carrier_vec n"
and c: "c \<in> carrier_vec m" and d: "d \<in> carrier_vec m"
shows "inner_prod a (outer_prod b c *\<^sub>v d) = inner_prod a b * inner_prod c d" (is "?lhs = ?rhs")
proof -
define P where "P = outer_prod b c"
then have dimP: "P \<in> carrier_mat n m" using assms by auto
have "inner_prod a (P *\<^sub>v d) = (\<Sum>i=0..<n. (\<Sum>j=0..<m. conjugate (a$i) * P$$(i, j) * d$j))" using assms dimP
apply (simp add: scalar_prod_def sum_distrib_right)
apply (rule sum.cong, auto)
apply (rule sum.cong, auto)
done
also have "\<dots> = (\<Sum>i=0..<n. (\<Sum>j=0..<m. conjugate (a$i) * b$i * conjugate(c$j) * d$j))"
using P_def b c by(simp add: index_outer_prod algebra_simps)
finally have eq: "?lhs = (\<Sum>i=0..<n. (\<Sum>j=0..<m. conjugate (a$i) * b$i * conjugate(c$j) * d$j))" using P_def by auto
have "?rhs = (\<Sum>i=0..<n. conjugate (a$i) * b$i) * (\<Sum>j=0..<m. conjugate(c$j) * d$j)" using assms
by (auto simp add: scalar_prod_def algebra_simps)
also have "\<dots> = (\<Sum>i=0..<n. (\<Sum>j=0..<m. conjugate (a$i) * b$i * conjugate(c$j) * d$j))"
using assms by (simp add: sum_product algebra_simps)
finally show "?lhs = ?rhs" using eq by auto
qed
subsection \<open>Semi-definite matrices\<close>
definition positive :: "complex mat \<Rightarrow> bool" where
"positive A \<longleftrightarrow>
A \<in> carrier_mat (dim_col A) (dim_col A) \<and>
(\<forall>v. dim_vec v = dim_col A \<longrightarrow> inner_prod v (A *\<^sub>v v) \<ge> 0)"
lemma positive_iff_normalized_vec:
"positive A \<longleftrightarrow>
A \<in> carrier_mat (dim_col A) (dim_col A) \<and>
(\<forall>v. (dim_vec v = dim_col A \<and> vec_norm v = 1) \<longrightarrow> inner_prod v (A *\<^sub>v v) \<ge> 0)"
proof (rule)
assume "positive A"
then show "A \<in> carrier_mat (dim_col A) (dim_col A) \<and>
(\<forall>v. dim_vec v = dim_col A \<and> vec_norm v = 1 \<longrightarrow> 0 \<le> inner_prod v (A *\<^sub>v v))"
unfolding positive_def by auto
next
define n where "n = dim_col A"
assume "A \<in> carrier_mat (dim_col A) (dim_col A) \<and> (\<forall>v. dim_vec v = dim_col A \<and> vec_norm v = 1 \<longrightarrow> 0 \<le> inner_prod v (A *\<^sub>v v))"
then have A: "A \<in> carrier_mat (dim_col A) (dim_col A)" and geq0: "\<forall>v. dim_vec v = dim_col A \<and> vec_norm v = 1 \<longrightarrow> 0 \<le> inner_prod v (A *\<^sub>v v)" by auto
then have dimA: "A \<in> carrier_mat n n" using n_def[symmetric] by auto
{
fix v assume dimv: "(v::complex vec) \<in> carrier_vec n"
have "0 \<le> inner_prod v (A *\<^sub>v v)"
proof (cases "v = 0\<^sub>v n")
case True
then show "0 \<le> inner_prod v (A *\<^sub>v v)" using dimA by auto
next
case False
then have 1: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto
then have cnv: "cnj (vec_norm v) = vec_norm v" using Reals_cnj_iff complex_is_Real_iff by auto
define w where "w = vec_normalize v"
then have dimw: "w \<in> carrier_vec n" using dimv by auto
have nvw: "v = vec_norm v \<cdot>\<^sub>v w" using w_def vec_eq_norm_smult_normalized by auto
have "vec_norm w = 1" using normalized_vec_norm[OF dimv False] vec_norm_def w_def by auto
then have 2: "0 \<le> inner_prod w (A *\<^sub>v w)" using geq0 dimw dimA by auto
have "inner_prod v (A *\<^sub>v v) = vec_norm v * vec_norm v * inner_prod w (A *\<^sub>v w)" using dimA dimv dimw
apply (subst (1 2) nvw)
apply (subst mult_mat_vec, simp, simp)
apply (subst scalar_prod_smult_left[of "(A *\<^sub>v w)" "conjugate (vec_norm v \<cdot>\<^sub>v w)" "vec_norm v"], simp)
apply (simp add: conjugate_smult_vec cnv)
done
also have "\<dots> \<ge> 0" using 1 2 by auto
finally show "0 \<le> inner_prod v (A *\<^sub>v v)" by auto
qed
}
then have geq: "\<forall>v. dim_vec v = dim_col A \<longrightarrow> 0 \<le> inner_prod v (A *\<^sub>v v)" using dimA by auto
show "positive A" unfolding positive_def
by (rule, simp add: A, rule geq)
qed
lemma positive_is_hermitian:
fixes A :: "complex mat"
assumes pA: "positive A"
shows "hermitian A"
proof -
define n where "n = dim_col A"
then have dimA: "A \<in> carrier_mat n n" using positive_def pA by auto
obtain B C where B: "hermitian B" and C: "hermitian C" and A: "A = B + \<i> \<cdot>\<^sub>m C"
and dimB: "B \<in> carrier_mat n n" and dimC: "C \<in> carrier_mat n n" and dimiC: "\<i> \<cdot>\<^sub>m C \<in> carrier_mat n n"
using complex_mat_decomposition_to_hermitian[OF dimA] by auto
{
fix v :: "complex vec" assume dimv: "v \<in> carrier_vec n"
have dimvA: "dim_vec v = dim_col A" using dimv dimA by auto
have "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + inner_prod v ((\<i> \<cdot>\<^sub>m C) *\<^sub>v v)"
unfolding A using dimB dimiC dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right)
moreover have "inner_prod v ((\<i> \<cdot>\<^sub>m C) *\<^sub>v v) = \<i> * inner_prod v (C *\<^sub>v v)" using dimv dimC
apply (simp add: scalar_prod_def sum_distrib_left cong: sum.cong)
apply (rule sum.cong, auto)
done
ultimately have ABC: "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + \<i> * inner_prod v (C *\<^sub>v v)" by auto
moreover have "inner_prod v (B *\<^sub>v v) \<in> Reals" using B dimB dimv hermitian_inner_prod_real by auto
moreover have "inner_prod v (C *\<^sub>v v) \<in> Reals" using C dimC dimv hermitian_inner_prod_real by auto
moreover have "inner_prod v (A *\<^sub>v v) \<in> Reals" using pA unfolding positive_def
apply (rule)
apply (fold n_def)
apply (simp add: complex_is_Real_iff[of "inner_prod v (A *\<^sub>v v)"])
apply (auto simp add: dimvA)
done
ultimately have "inner_prod v (C *\<^sub>v v) = 0" using of_real_Re by fastforce
}
then have "C = 0\<^sub>m n n" using hermitian_inner_prod_zero dimC C by auto
then have "A = B" using A dimC dimB by auto
then show "hermitian A" using B by auto
qed
lemma positive_eigenvalue_positive:
assumes dimA: "(A::complex mat) \<in> carrier_mat n n"
and pA: "positive A"
and c: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
and B: "unitary_schur_decomposition A es = (B,P,Q)"
shows "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<ge> 0"
proof -
have hA: "hermitian A" using positive_is_hermitian pA by auto
have "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es
\<and> unitary P \<and> (\<forall>i < n. B$$(i, i) \<in> Reals)"
using hermitian_eigenvalue_real dimA hA B c by auto
then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B"
and Bii: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<in> Reals"
and dimB: "B \<in> carrier_mat n n" and dimP: "P \<in> carrier_mat n n" and dimaP: "adjoint P \<in> carrier_mat n n"
and uP: "unitary P"
unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto
{
fix i assume i: "i < n"
define v where "v = col P i"
then have dimv: "v \<in> carrier_vec n" using v_def dimP by auto
have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def
using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto
moreover have "inner_prod v (A *\<^sub>v v) \<ge> 0" using dimv pA dimA positive_def by auto
ultimately show "B$$(i, i) \<ge> 0" by auto
}
qed
lemma diag_mat_mult_diag_mat:
fixes B D :: "'a::semiring_0 mat"
assumes dimB: "B \<in> carrier_mat n n" and dimD: "D \<in> carrier_mat n n"
and dB: "diagonal_mat B" and dD: "diagonal_mat D"
shows "B * D = mat n n (\<lambda>(i,j). (if i = j then (B$$(i, i)) * (D$$(i, i)) else 0))"
proof(rule eq_matI, auto)
have Bij: "\<And>x y. x < n \<Longrightarrow> y < n \<Longrightarrow> x \<noteq> y \<Longrightarrow> B$$(x, y) = 0" using dB diagonal_mat_def dimB by auto
have Dij: "\<And>x y. x < n \<Longrightarrow> y < n \<Longrightarrow> x \<noteq> y \<Longrightarrow> D$$(x, y) = 0" using dD diagonal_mat_def dimD by auto
{
fix i j assume ij: "i < n" "j < n"
have "(B * D) $$ (i, j) = (\<Sum>k=0..<n. (B $$ (i, k)) * (D $$ (k, j)))" using dimB dimD
by (auto simp add: scalar_prod_def ij)
also have "\<dots> = B$$(i, i) * D$$(i, j)"
apply (simp add: sum.remove[of _i] ij)
apply (simp add: Bij Dij ij)
done
finally have "(B * D) $$ (i, j) = B$$(i, i) * D$$(i, j)".
}
note BDij = this
from BDij show "\<And>j. j < n \<Longrightarrow> (B * D) $$ (j, j) = B $$ (j, j) * D $$ (j, j)" by auto
from BDij show "\<And>i j. i < n \<Longrightarrow> j < n \<Longrightarrow> i \<noteq> j \<Longrightarrow> (B * D) $$ (i, j) = 0" using Bij Dij by auto
from assms show "dim_row B = n" "dim_col D = n" by auto
qed
lemma positive_only_if_decomp:
assumes dimA: "A \<in> carrier_mat n n" and pA: "positive A"
shows "\<exists>M \<in> carrier_mat n n. M * adjoint M = A"
proof -
from pA have hA: "hermitian A" using positive_is_hermitian by auto
obtain es where es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
using complex_mat_char_poly_factorizable dimA by auto
obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)"
by (cases "unitary_schur_decomposition A es", auto)
then have "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> diag_mat B = es
\<and> unitary P \<and> (\<forall>i < n. B$$(i, i) \<in> Reals)"
using hermitian_eigenvalue_real dimA es hA by auto
then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B"
and Bii: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<in> Reals"
and dimB: "B \<in> carrier_mat n n" and dimP: "P \<in> carrier_mat n n" and dimaP: "adjoint P \<in> carrier_mat n n"
unfolding similar_mat_wit_def Let_def using dimA by auto
have Bii: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<ge> 0" using pA dimA es schur positive_eigenvalue_positive by auto
define D where "D = mat n n (\<lambda>(i, j). (if (i = j) then csqrt (B$$(i, i)) else 0))"
then have dimD: "D \<in> carrier_mat n n" and dimaD: "adjoint D \<in> carrier_mat n n" using dimB by auto
have dD: "diagonal_mat D" using dB D_def unfolding diagonal_mat_def by auto
then have daD: "diagonal_mat (adjoint D)" by (simp add: adjoint_eval diagonal_mat_def)
have Dii: "\<And>i. i < n \<Longrightarrow> D$$(i, i) = csqrt (B$$(i, i))" using dimD D_def by auto
{
fix i assume i: "i < n"
define c where "c = csqrt (B$$(i, i))"
have c: "c \<ge> 0" using Bii i c_def by auto
then have "conjugate c = c"
using Reals_cnj_iff complex_is_Real_iff by auto
then have "c * cnj c = B$$(i, i)" using c_def c unfolding conjugate_complex_def by (metis power2_csqrt power2_eq_square)
}
note cBii = this
have "D * adjoint D = mat n n (\<lambda>(i,j). (if (i = j) then B$$(i, i) else 0))"
apply (simp add: diag_mat_mult_diag_mat[OF dimD dimaD dD daD])
apply (rule eq_matI, auto simp add: D_def adjoint_eval cBii)
done
also have "\<dots> = B" using dimB dB[unfolded diagonal_mat_def] by auto
finally have DaDB: "D * adjoint D = B".
define M where "M = P * D"
then have dimM: "M \<in> carrier_mat n n" using dimP dimD by auto
have "M * adjoint M = (P * D) * (adjoint D * adjoint P)" using M_def adjoint_mult[OF dimP dimD] by auto
also have "\<dots> = P * (D * adjoint D) * (adjoint P)" using dimP dimD by (mat_assoc n)
also have "\<dots> = P * B * (adjoint P)" using DaDB by auto
finally have "M * adjoint M = A" using A by auto
with dimM show "\<exists>M \<in> carrier_mat n n. M * adjoint M = A" by auto
qed
lemma positive_if_decomp:
assumes dimA: "A \<in> carrier_mat n n" and "\<exists>M. M * adjoint M = A"
shows "positive A"
proof -
from assms obtain M where M: "M * adjoint M = A" by auto
define m where "m = dim_col M"
have dimM: "M \<in> carrier_mat n m" using M dimA m_def by auto
{
fix v assume dimv: "(v::complex vec) \<in> carrier_vec n"
have dimaM: "adjoint M \<in> carrier_mat m n" using dimM by auto
have dimaMv: "(adjoint M) *\<^sub>v v \<in> carrier_vec m" using dimaM dimv by auto
have "inner_prod v (A *\<^sub>v v) = inner_prod v (M * adjoint M *\<^sub>v v)" using M by auto
also have "\<dots> = inner_prod v (M *\<^sub>v (adjoint M *\<^sub>v v))" using assoc_mult_mat_vec dimM dimaM dimv by auto
also have "\<dots> = inner_prod (adjoint M *\<^sub>v v) (adjoint M *\<^sub>v v)" using adjoint_def_alter[OF dimv dimaMv dimM] by auto
also have "\<dots> \<ge> 0" using self_cscalar_prod_geq_0 by auto
finally have "inner_prod v (A *\<^sub>v v) \<ge> 0".
}
note geq0 = this
from dimA geq0 show "positive A" using positive_def by auto
qed
lemma positive_iff_decomp:
assumes dimA: "A \<in> carrier_mat n n"
shows "positive A \<longleftrightarrow> (\<exists>M\<in>carrier_mat n n. M * adjoint M = A)"
proof
assume pA: "positive A"
then show "\<exists>M\<in>carrier_mat n n. M * adjoint M = A" using positive_only_if_decomp assms by auto
next
assume "\<exists>M\<in>carrier_mat n n. M * adjoint M = A"
then obtain M where M: "M * adjoint M = A" by auto
then show "positive A" using M positive_if_decomp assms by auto
qed
lemma positive_dim_eq:
assumes "positive A"
shows "dim_row A = dim_col A"
using carrier_matD(1)[of A "dim_col A" "dim_col A"] assms[unfolded positive_def] by simp
lemma positive_zero:
"positive (0\<^sub>m n n)"
by (simp add: positive_def zero_mat_def mult_mat_vec_def scalar_prod_def)
lemma positive_one:
"positive (1\<^sub>m n)"
proof (rule positive_if_decomp)
show "1\<^sub>m n \<in> carrier_mat n n" by auto
have "adjoint (1\<^sub>m n) = 1\<^sub>m n" using hermitian_one hermitian_def by auto
then have "1\<^sub>m n * adjoint (1\<^sub>m n) = 1\<^sub>m n" by auto
then show "\<exists>M. M * adjoint M = 1\<^sub>m n" by fastforce
qed
lemma positive_antisym:
assumes pA: "positive A" and pnA: "positive (-A)"
shows "A = 0\<^sub>m (dim_col A) (dim_col A)"
proof -
define n where "n = dim_col A"
from pA have dimA: "A \<in> carrier_mat n n" and dimnA: "-A \<in> carrier_mat n n"
using positive_def n_def by auto
from pA have hA: "hermitian A" using positive_is_hermitian by auto
obtain es where es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> es. [:- e, 1:])"
using complex_mat_char_poly_factorizable dimA by auto
obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)"
by (cases "unitary_schur_decomposition A es", auto)
then have "similar_mat_wit A B P (adjoint P) \<and> diagonal_mat B \<and> unitary P"
using hermitian_eigenvalue_real dimA es hA by auto
then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and uP: "unitary P"
and dimB: "B \<in> carrier_mat n n" and dimnB: "-B \<in> carrier_mat n n"
and dimP: "P \<in> carrier_mat n n" and dimaP: "adjoint P \<in> carrier_mat n n"
unfolding similar_mat_wit_def Let_def using dimA by auto
from es schur have geq0: "\<And>i. i < n \<Longrightarrow> B$$(i, i) \<ge> 0" using positive_eigenvalue_positive dimA pA by auto
from A have nA: "-A = P * (-B) * (adjoint P)" using mult_smult_assoc_mat dimB dimP dimaP by auto
from dB have dnB: "diagonal_mat (-B)" by (simp add: diagonal_mat_def)
{
fix i assume i: "i < n"
define v where "v = col P i"
then have dimv: "v \<in> carrier_vec n" using v_def dimP by auto
have "inner_prod v ((-A) *\<^sub>v v) = (-B)$$(i, i)" unfolding nA v_def
using spectral_decomposition_extract_diag[OF dimP dimnB uP dnB i] by auto
moreover have "inner_prod v ((-A) *\<^sub>v v) \<ge> 0" using dimv pnA dimnA positive_def by auto
ultimately have "B$$(i, i) \<le> 0" using dimB i by auto
moreover have "B$$(i, i) \<ge> 0" using i geq0 by auto
ultimately have "B$$(i, i) = 0" by (metis no_atp(10))
}
then have "B = 0\<^sub>m n n" using dimB dB[unfolded diagonal_mat_def]
by (subst eq_matI, auto)
then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto
qed
lemma positive_add:
assumes pA: "positive A" and pB: "positive B"
and dimA: "A \<in> carrier_mat n n" and dimB: "B \<in> carrier_mat n n"
shows "positive (A + B)"
unfolding positive_def
proof
have dimApB: "A + B \<in> carrier_mat n n" using dimA dimB by auto
then show "A + B \<in> carrier_mat (dim_col (A + B)) (dim_col (A + B))" using carrier_matD[of "A+B"] by auto
{
fix v assume dimv: "(v::complex vec) \<in> carrier_vec n"
have 1: "inner_prod v (A *\<^sub>v v) \<ge> 0" using dimv pA[unfolded positive_def] dimA by auto
have 2: "inner_prod v (B *\<^sub>v v) \<ge> 0" using dimv pB[unfolded positive_def] dimB by auto
have "inner_prod v ((A + B) *\<^sub>v v) = inner_prod v (A *\<^sub>v v) + inner_prod v (B *\<^sub>v v)"
using dimA dimB dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right)
also have "\<dots> \<ge> 0" using 1 2 by auto
finally have "inner_prod v ((A + B) *\<^sub>v v) \<ge> 0".
}
note geq0 = this
then have "\<And>v. dim_vec v = n \<Longrightarrow> 0 \<le> inner_prod v ((A + B) *\<^sub>v v)" by auto
then show "\<forall>v. dim_vec v = dim_col (A + B) \<longrightarrow> 0 \<le> inner_prod v ((A + B) *\<^sub>v v)" using dimApB by auto
qed
lemma positive_trace:
assumes "A \<in> carrier_mat n n" and "positive A"
shows "trace A \<ge> 0"
using assms positive_iff_decomp trace_adjoint_positive by auto
lemma positive_close_under_left_right_mult_adjoint:
fixes M A :: "complex mat"
assumes dM: "M \<in> carrier_mat n n" and dA: "A \<in> carrier_mat n n"
and pA: "positive A"
shows "positive (M * A * adjoint M)"
unfolding positive_def
proof (rule, simp add: mult_carrier_mat[OF mult_carrier_mat[OF dM dA] adjoint_dim[OF dM]] carrier_matD[OF dM], rule, rule)
have daM: "adjoint M \<in> carrier_mat n n" using dM by auto
fix v::"complex vec" assume "dim_vec v = dim_col (M * A * adjoint M)"
then have dv: "v \<in> carrier_vec n" using assms by auto
then have "adjoint M *\<^sub>v v \<in> carrier_vec n" using daM by auto
have assoc: "M * A * adjoint M *\<^sub>v v = M *\<^sub>v (A *\<^sub>v (adjoint M *\<^sub>v v))"
using dA dM daM dv by (auto simp add: assoc_mult_mat_vec[of _ n n _ n])
have "inner_prod v (M * A * adjoint M *\<^sub>v v) = inner_prod (adjoint M *\<^sub>v v) (A *\<^sub>v (adjoint M *\<^sub>v v))"
apply (subst assoc)
apply (subst adjoint_def_alter[where ?A = "M"])
by (auto simp add: dv dA daM dM carrier_matD[OF dM] mult_mat_vec_carrier[of _ n n])
also have "\<dots> \<ge> 0" using dA dv daM pA positive_def by auto
finally show "inner_prod v (M * A * adjoint M *\<^sub>v v) \<ge> 0" by auto
qed
lemma positive_same_outer_prod:
fixes v w :: "complex vec"
assumes v: "v \<in> carrier_vec n"
shows "positive (outer_prod v v)"
proof -
have d1: "adjoint (mat (dim_vec v) 1 (\<lambda>(i, j). v $ i)) \<in> carrier_mat 1 n" using assms by auto
have d2: "mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y) \<in> carrier_mat 1 n" using assms by auto
have dv: "dim_vec v = n" using assms by auto
have "mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y) = adjoint (mat (dim_vec v) 1 (\<lambda>(i, j). v $ i))" (is "?r = adjoint ?l")
apply (rule eq_matI)
subgoal for i j by (simp add: dv adjoint_eval)
using d1 d2 by auto
then have "outer_prod v v = ?l * adjoint ?l" unfolding outer_prod_def by auto
then have "\<exists>M. M * adjoint M = outer_prod v v" by auto
then show "positive (outer_prod v v)" using positive_if_decomp[OF outer_prod_dim[OF v v]] by auto
qed
lemma smult_smult_mat:
fixes k :: complex and l :: complex
assumes "A \<in> carrier_mat nr n"
shows "k \<cdot>\<^sub>m (l \<cdot>\<^sub>m A) = (k * l) \<cdot>\<^sub>m A" by auto
lemma positive_smult:
assumes "A \<in> carrier_mat n n"
and "positive A"
and "c \<ge> 0"
shows "positive (c \<cdot>\<^sub>m A)"
proof -
have sc: "csqrt c \<ge> 0" using assms(3) by fastforce
obtain M where dimM: "M \<in> carrier_mat n n" and A: "M * adjoint M = A" using assms(1-2) positive_iff_decomp by auto
have "c \<cdot>\<^sub>m A = c \<cdot>\<^sub>m (M * adjoint M)" using A by auto
have ccsq: "conjugate (csqrt c) = (csqrt c)" using sc Reals_cnj_iff[of "csqrt c"] complex_is_Real_iff by auto
have MM: "(M * adjoint M) \<in> carrier_mat n n" using A assms by fastforce
have leftd: "c \<cdot>\<^sub>m (M * adjoint M) \<in> carrier_mat n n" using A assms by fastforce
have rightd: "(csqrt c \<cdot>\<^sub>m M) * (adjoint (csqrt c \<cdot>\<^sub>m M))\<in> carrier_mat n n" using A assms by fastforce
have "(csqrt c \<cdot>\<^sub>m M) * (adjoint (csqrt c \<cdot>\<^sub>m M)) = (csqrt c \<cdot>\<^sub>m M) * ((conjugate (csqrt c)) \<cdot>\<^sub>m adjoint M)"
using adjoint_scale assms(1) by (metis adjoint_scale)
also have "\<dots> = (csqrt c \<cdot>\<^sub>m M) * (csqrt c \<cdot>\<^sub>m adjoint M)" using sc ccsq by fastforce
also have "\<dots> = csqrt c \<cdot>\<^sub>m (M * (csqrt c \<cdot>\<^sub>m adjoint M))"
using mult_smult_assoc_mat index_smult_mat(2,3) by fastforce
also have "\<dots> = csqrt c \<cdot>\<^sub>m ((csqrt c) \<cdot>\<^sub>m (M * adjoint M))"
using mult_smult_distrib by fastforce
also have "\<dots> = c \<cdot>\<^sub>m (M * adjoint M)"
using smult_smult_mat[of "M * adjoint M" n n "(csqrt c)" "(csqrt c)"] MM sc
by (metis power2_csqrt power2_eq_square )
also have "\<dots> = c \<cdot>\<^sub>m A" using A by auto
finally have "(csqrt c \<cdot>\<^sub>m M) * (adjoint (csqrt c \<cdot>\<^sub>m M)) = c \<cdot>\<^sub>m A" by auto
moreover have "c \<cdot>\<^sub>m A \<in> carrier_mat n n" using assms(1) by auto
moreover have "csqrt c \<cdot>\<^sub>m M \<in> carrier_mat n n" using dimM by auto
ultimately show ?thesis using positive_iff_decomp by auto
qed
text \<open>Version of previous theorem for real numbers\<close>
lemma positive_scale:
fixes c :: real
assumes "A \<in> carrier_mat n n"
and "positive A"
and "c \<ge> 0"
shows "positive (c \<cdot>\<^sub>m A)"
apply (rule positive_smult) using assms by auto
subsection \<open>L\"{o}wner partial order\<close>
definition lowner_le :: "complex mat \<Rightarrow> complex mat \<Rightarrow> bool" (infix "\<le>\<^sub>L" 50) where
"A \<le>\<^sub>L B \<longleftrightarrow> dim_row A = dim_row B \<and> dim_col A = dim_col B \<and> positive (B - A)"
lemma lowner_le_refl:
assumes "A \<in> carrier_mat n n"
shows "A \<le>\<^sub>L A"
unfolding lowner_le_def
apply (simp add: minus_r_inv_mat[OF assms])
by (rule positive_zero)
lemma lowner_le_antisym:
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n"
and L1: "A \<le>\<^sub>L B" and L2: "B \<le>\<^sub>L A"
shows "A = B"
proof -
from L1 have P1: "positive (B - A)" by (simp add: lowner_le_def)
from L2 have P2: "positive (A - B)" by (simp add: lowner_le_def)
have "A - B = - (B - A)" using A B by auto
then have P3: "positive (- (B - A))" using P2 by auto
have BA: "B - A \<in> carrier_mat n n" using A B by auto
have "B - A = 0\<^sub>m n n" using BA by (subst positive_antisym[OF P1 P3], auto)
then have "B + (-A) + A = 0\<^sub>m n n + A" using A B minus_add_uminus_mat[OF B A] by auto
then have "B + (-A + A) = 0\<^sub>m n n + A" using A B by auto
then show "A = B" using A B BA uminus_l_inv_mat[OF A] by auto
qed
lemma lowner_le_inner_prod_le:
fixes A B :: "complex mat" and v :: "complex vec"
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n"
and v: "v \<in> carrier_vec n"
and "A \<le>\<^sub>L B"
shows "inner_prod v (A *\<^sub>v v) \<le> inner_prod v (B *\<^sub>v v)"
proof -
from assms have "positive (B-A)" by (auto simp add: lowner_le_def)
with assms have geq: "inner_prod v ((B-A) *\<^sub>v v) \<ge> 0"
unfolding positive_def by auto
have "inner_prod v ((B-A) *\<^sub>v v) = inner_prod v (B *\<^sub>v v) - inner_prod v (A *\<^sub>v v)"
unfolding minus_add_uminus_mat[OF B A]
by (subst add_mult_distrib_mat_vec[OF B _ v], insert A B v, auto simp add: inner_prod_distrib_right[OF v])
then show ?thesis using geq by auto
qed
lemma lowner_le_trans:
fixes A B C :: "complex mat"
assumes A: "A \<in> carrier_mat n n" and B: "B \<in> carrier_mat n n" and C: "C \<in> carrier_mat n n"
and L1: "A \<le>\<^sub>L B" and L2: "B \<le>\<^sub>L C"
shows "A \<le>\<^sub>L C"
unfolding lowner_le_def
proof (auto simp add: carrier_matD[OF A] carrier_matD[OF C])
have dim: "C - A \<in> carrier_mat n n" using A C by auto
{
fix v assume v: "(v::complex vec) \<in> carrier_vec n"
from L1 have "inner_prod v (A *\<^sub>v v) \<le> inner_prod v (B *\<^sub>v v)" using lowner_le_inner_prod_le A B v by auto
also from L2 have "\<dots> \<le> inner_prod v (C *\<^sub>v v)" using lowner_le_inner_prod_le B C v by auto
finally have "inner_prod v (A *\<^sub>v v) \<le> inner_prod v (C *\<^sub>v v)".
then have "inner_prod v (C *\<^sub>v v) - inner_prod v (A *\<^sub>v v) \<ge> 0" by auto
then have "inner_prod v ((C - A) *\<^sub>v v) \<ge> 0" using A C v
apply (subst minus_add_uminus_mat[OF C A])
apply (subst add_mult_distrib_mat_vec[OF C _ v], simp)
apply (simp add: inner_prod_distrib_right[OF v])
done
}
note leq = this
show "positive (C - A)" unfolding positive_def
apply (rule, simp add: carrier_matD[OF A] dim)
apply (subst carrier_matD[OF dim], insert leq, auto)
done
qed
lemma lowner_le_imp_trace_le:
assumes "A \<in> carrier_mat n n" and "B \<in> carrier_mat n n"
and "A \<le>\<^sub>L B"
shows "trace A \<le> trace B"
proof -
have "positive (B - A)" using assms lowner_le_def by auto
moreover have "B - A \<in> carrier_mat n n" using assms by auto
ultimately have "trace (B - A) \<ge> 0" using positive_trace by auto
moreover have "trace (B - A) = trace B - trace A" using trace_minus_linear assms by auto
ultimately have "trace B - trace A \<ge> 0" by auto
then show "trace A \<le> trace B" by auto
qed
lemma lowner_le_add:
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "C \<in> carrier_mat n n" "D \<in> carrier_mat n n"
and "A \<le>\<^sub>L B" "C \<le>\<^sub>L D"
shows "A + C \<le>\<^sub>L B + D"
proof -
have "B + D - (A + C) = B - A + (D - C) " using assms by auto
then have "positive (B + D - (A + C))" using assms unfolding lowner_le_def using positive_add
by (metis minus_carrier_mat)
then show "A + C \<le>\<^sub>L B + D" unfolding lowner_le_def using assms by fastforce
qed
lemma lowner_le_swap:
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n"
and "A \<le>\<^sub>L B"
shows "-B \<le>\<^sub>L -A"
proof -
have "positive (B - A)" using assms lowner_le_def by fastforce
moreover have "B - A = (-A) - (-B)" using assms by fastforce
ultimately have "positive ((-A) - (-B))" by auto
then show ?thesis using lowner_le_def assms by fastforce
qed
lemma lowner_le_minus:
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "C \<in> carrier_mat n n" "D \<in> carrier_mat n n"
and "A \<le>\<^sub>L B" "C \<le>\<^sub>L D"
shows "A - D \<le>\<^sub>L B - C"
proof -
have "positive (D - C)" using assms lowner_le_def by auto
then have "-D \<le>\<^sub>L -C" using lowner_le_swap assms by auto
then have "A + (-D) \<le>\<^sub>L B + (-C)" using lowner_le_add[of "A" n "B"] assms by auto
moreover have "A + (-D) = A - D" and "B + (-C) = B - C" by auto
ultimately show ?thesis by auto
qed
lemma outer_prod_le_one:
assumes "v \<in> carrier_vec n"
and "inner_prod v v \<le> 1"
shows "outer_prod v v \<le>\<^sub>L 1\<^sub>m n"
proof -
let ?o = "outer_prod v v"
have do: "?o \<in> carrier_mat n n" using assms by auto
{
fix u :: "complex vec" assume "dim_vec u = n"
then have du: "u \<in> carrier_vec n" by auto
have r: "inner_prod u u \<in> Reals" apply (simp add: scalar_prod_def carrier_vecD[OF du])
using complex_In_mult_cnj_zero complex_is_Real_iff by blast
have geq0: "inner_prod u u \<ge> 0"
using self_cscalar_prod_geq_0 by auto
have "inner_prod u (?o *\<^sub>v u) = inner_prod u v * inner_prod v u"
apply (subst inner_prod_outer_prod)
using du assms by auto
also have "\<dots> \<le> inner_prod u u * inner_prod v v" using Cauchy_Schwarz_complex_vec du assms by auto
also have "\<dots> \<le> inner_prod u u" using assms(2) r geq0
by (simp add: mult_right_le_one_le)
finally have le: "inner_prod u (?o *\<^sub>v u) \<le> inner_prod u u".
have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) = inner_prod u ((1\<^sub>m n *\<^sub>v u) - ?o *\<^sub>v u)"
apply (subst minus_mult_distrib_mat_vec) using do du by auto
also have "\<dots> = inner_prod u u - inner_prod u (?o *\<^sub>v u)"
apply (subst inner_prod_minus_distrib_right)
using du do by auto
also have "\<dots> \<ge> 0" using le by auto
finally have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) \<ge> 0" by auto
}
then have "positive (1\<^sub>m n - outer_prod v v)"
unfolding positive_def using do by auto
then show ?thesis unfolding lowner_le_def using do by auto
qed
lemma zero_lowner_le_positiveD:
fixes A :: "complex mat"
assumes dA: "A \<in> carrier_mat n n" and le: "0\<^sub>m n n \<le>\<^sub>L A"
shows "positive A"
using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto)
lemma zero_lowner_le_positiveI:
fixes A :: "complex mat"
assumes dA: "A \<in> carrier_mat n n" and le: "positive A"
shows "0\<^sub>m n n \<le>\<^sub>L A"
using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto)
lemma lowner_le_trans_positiveI:
fixes A B :: "complex mat"
assumes dA: "A \<in> carrier_mat n n" and pA: "positive A" and le: "A \<le>\<^sub>L B"
shows "positive B"
proof -
have dB: "B \<in> carrier_mat n n" using le dA lowner_le_def by auto
have "0\<^sub>m n n \<le>\<^sub>L A" using zero_lowner_le_positiveI dA pA by auto
then have "0\<^sub>m n n \<le>\<^sub>L B" using dA dB le by (simp add: lowner_le_trans[of _ n A B])
then show ?thesis using dB zero_lowner_le_positiveD by auto
qed
lemma lowner_le_keep_under_measurement:
fixes M A B :: "complex mat"
assumes dM: "M \<in> carrier_mat n n" and dA: "A \<in> carrier_mat n n" and dB: "B \<in> carrier_mat n n"
and le: "A \<le>\<^sub>L B"
shows "adjoint M * A * M \<le>\<^sub>L adjoint M * B * M"
unfolding lowner_le_def
proof (rule conjI, fastforce)+
have daM: "adjoint M \<in> carrier_mat n n" using dM by auto
have dBmA: "B - A \<in> carrier_mat n n" using dB dA by fastforce
have "positive (B - A)" using le lowner_le_def by auto
then have p: "positive (adjoint M * (B - A) * M)"
using positive_close_under_left_right_mult_adjoint[OF daM dBmA] adjoint_adjoint[of M] by auto
moreover have e: "adjoint M * (B - A) * M = adjoint M * B * M - adjoint M * A * M" using dM dB dA by (mat_assoc n)
ultimately show "positive (adjoint M * B * M - adjoint M * A * M)" by auto
qed
lemma smult_distrib_left_minus_mat:
fixes A B :: "'a::comm_ring_1 mat"
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n"
shows "c \<cdot>\<^sub>m (B - A) = c \<cdot>\<^sub>m B - c \<cdot>\<^sub>m A"
using assms by (auto simp add: minus_add_uminus_mat add_smult_distrib_left_mat)
lemma lowner_le_smultc:
fixes c :: complex
assumes "c \<ge> 0" "A \<le>\<^sub>L B" "A \<in> carrier_mat n n" "B \<in> carrier_mat n n"
shows "c \<cdot>\<^sub>m A \<le>\<^sub>L c \<cdot>\<^sub>m B"
proof -
have eqBA: "c \<cdot>\<^sub>m (B - A) = c \<cdot>\<^sub>m B - c \<cdot>\<^sub>m A"
using assms by (auto simp add: smult_distrib_left_minus_mat)
have "positive (B - A)" using assms(2) unfolding lowner_le_def by auto
then have "positive (c \<cdot>\<^sub>m (B - A))" using positive_smult[of "B-A" n c] assms by fastforce
moreover have "c \<cdot>\<^sub>m A \<in> carrier_mat n n" using index_smult_mat(2,3) assms(3) by auto
moreover have "c \<cdot>\<^sub>m B \<in> carrier_mat n n" using index_smult_mat(2,3) assms(4) by auto
ultimately show ?thesis unfolding lowner_le_def using eqBA by fastforce
qed
lemma lowner_le_smult:
fixes c :: real
assumes "c \<ge> 0" "A \<le>\<^sub>L B" "A \<in> carrier_mat n n" "B \<in> carrier_mat n n"
shows "c \<cdot>\<^sub>m A \<le>\<^sub>L c \<cdot>\<^sub>m B"
apply (rule lowner_le_smultc) using assms by auto
lemma minus_smult_vec_distrib:
fixes w :: "'a::comm_ring_1 vec"
shows "(a - b) \<cdot>\<^sub>v w = a \<cdot>\<^sub>v w - b \<cdot>\<^sub>v w"
apply (rule eq_vecI)
by (auto simp add: scalar_prod_def algebra_simps)
lemma smult_mat_mult_mat_vec_assoc:
fixes A :: "'a::comm_ring_1 mat"
assumes A: "A \<in> carrier_mat n m" and w: "w \<in> carrier_vec m"
shows "a \<cdot>\<^sub>m A *\<^sub>v w = a \<cdot>\<^sub>v (A *\<^sub>v w)"
apply (rule eq_vecI)
apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w])
apply (subst sum_distrib_left) apply (rule sum.cong, simp)
by auto
lemma mult_mat_vec_smult_vec_assoc:
fixes A :: "'a::comm_ring_1 mat"
assumes A: "A \<in> carrier_mat n m" and w: "w \<in> carrier_vec m"
shows "A *\<^sub>v (a \<cdot>\<^sub>v w) = a \<cdot>\<^sub>v (A *\<^sub>v w)"
apply (rule eq_vecI)
apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w])
apply (subst sum_distrib_left) apply (rule sum.cong, simp)
by auto
lemma outer_prod_left_right_mat:
fixes A B :: "complex mat"
assumes du: "u \<in> carrier_vec d2" and dv: "v \<in> carrier_vec d3"
and dA: "A \<in> carrier_mat d1 d2" and dB: "B \<in> carrier_mat d3 d4"
shows "A * (outer_prod u v) * B = (outer_prod (A *\<^sub>v u) (adjoint B *\<^sub>v v))"
unfolding outer_prod_def
proof -
have eq1: "A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i)) = mat (dim_vec (A *\<^sub>v u)) 1 (\<lambda>(i, j). (A *\<^sub>v u) $ i)"
apply (rule eq_matI)
by (auto simp add: dA du scalar_prod_def)
have conj: "conjugate a * b = conjugate ((a::complex) * conjugate b) " for a b by auto
have eq2: "mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y) * B = mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\<lambda>(i, y). conjugate (adjoint B *\<^sub>v v) $ y)"
apply (rule eq_matI)
apply (auto simp add: carrier_matD[OF dB] carrier_vecD[OF dv] scalar_prod_def adjoint_def conjugate_vec_def sum_conjugate )
apply (rule sum.cong)
by (auto simp add: conj)
have "A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i) * mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y)) * B =
(A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i))) *(mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y)) * B"
using dA du dv dB assoc_mult_mat[OF dA, of "mat (dim_vec u) 1 (\<lambda>(i, j). u $ i)" 1 "mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y)"] by fastforce
also have "\<dots> = (A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i))) *((mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y)) * B)"
using dA du dv dB assoc_mult_mat[OF _ _ dB, of "(A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i)))" d1 1] by fastforce
finally show "A * (mat (dim_vec u) 1 (\<lambda>(i, j). u $ i) * mat 1 (dim_vec v) (\<lambda>(i, y). conjugate v $ y)) * B =
mat (dim_vec (A *\<^sub>v u)) 1 (\<lambda>(i, j). (A *\<^sub>v u) $ i) * mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\<lambda>(i, y). conjugate (adjoint B *\<^sub>v v) $ y)"
using eq1 eq2 by auto
qed
subsection \<open>Density operators\<close>
definition density_operator :: "complex mat \<Rightarrow> bool" where
"density_operator A \<longleftrightarrow> positive A \<and> trace A = 1"
definition partial_density_operator :: "complex mat \<Rightarrow> bool" where
"partial_density_operator A \<longleftrightarrow> positive A \<and> trace A \<le> 1"
lemma pure_state_self_outer_prod_is_partial_density_operator:
fixes v :: "complex vec"
assumes dimv: "v \<in> carrier_vec n" and nv: "vec_norm v = 1"
shows "partial_density_operator (outer_prod v v)"
unfolding partial_density_operator_def
proof
have dimov: "outer_prod v v \<in> carrier_mat n n" using dimv by auto
show "positive (outer_prod v v)" unfolding positive_def
proof (rule, simp add: carrier_matD(2)[OF dimov] dimov, rule allI, rule impI)
fix w assume "dim_vec (w::complex vec) = dim_col (outer_prod v v)"
then have dimw: "w \<in> carrier_vec n" using dimov carrier_vecI by auto
then have "inner_prod w ((outer_prod v v) *\<^sub>v w) = inner_prod w v * inner_prod v w"
using inner_prod_outer_prod dimw dimv by auto
also have "\<dots> = inner_prod w v * conjugate (inner_prod w v)" using dimw dimv
apply (subst conjugate_scalar_prod[of v "conjugate w"], simp)
apply (subst conjugate_vec_sprod_comm[of "conjugate v" _ "conjugate w"], auto)
apply (rule carrier_vec_conjugate[OF dimv])
apply (rule carrier_vec_conjugate[OF dimw])
done
also have "\<dots> \<ge> 0" by auto
finally show "inner_prod w ((outer_prod v v) *\<^sub>v w) \<ge> 0".
qed
have eq: "trace (outer_prod v v) = (\<Sum>i=0..<n. v$i * conjugate(v$i))" unfolding trace_def
apply (subst carrier_matD(1)[OF dimov])
apply (simp add: index_outer_prod[OF dimv dimv])
done
have "vec_norm v = csqrt (\<Sum>i=0..<n. v$i * conjugate(v$i))" unfolding vec_norm_def using dimv
by (simp add: scalar_prod_def)
then have "(\<Sum>i=0..<n. v$i * conjugate(v$i)) = 1" using nv by auto
with eq show "trace (outer_prod v v) \<le> 1" by auto
qed
(* Lemma 2.1 *)
lemma lowner_le_trace:
assumes A: "A \<in> carrier_mat n n"
and B: "B \<in> carrier_mat n n"
shows "A \<le>\<^sub>L B \<longleftrightarrow> (\<forall>\<rho>\<in>carrier_mat n n. partial_density_operator \<rho> \<longrightarrow> trace (A * \<rho>) \<le> trace (B * \<rho>))"
proof (rule iffI)
have dimBmA: "B - A \<in> carrier_mat n n" using A B by auto
{
assume "A \<le>\<^sub>L B"
then have pBmA: "positive (B - A)" using lowner_le_def by auto
moreover have "B - A \<in> carrier_mat n n" using assms by auto
ultimately have "\<exists>M\<in>carrier_mat n n. M * adjoint M = B - A" using positive_iff_decomp[of "B - A"] by auto
then obtain M where dimM: "M \<in> carrier_mat n n" and M: "M * adjoint M = B - A" by auto
{
fix \<rho> assume dimr: "\<rho> \<in> carrier_mat n n" and pdr: "partial_density_operator \<rho>"
have eq: "trace(B * \<rho>) - trace(A * \<rho>) = trace((B - A) * \<rho>)" using A B dimr
apply (subst minus_mult_distrib_mat, auto)
apply (subst trace_minus_linear, auto)
done
have pr: "positive \<rho>" using pdr partial_density_operator_def by auto
then have "\<exists>P\<in>carrier_mat n n. \<rho> = P * adjoint P" using positive_iff_decomp dimr by auto
then obtain P where dimP: "P \<in> carrier_mat n n" and P: "\<rho> = P * adjoint P" by auto
have "trace((B - A) * \<rho>) = trace(M * adjoint M * (P * adjoint P))" using P M by auto
also have "\<dots> = trace((adjoint P * M) * adjoint (adjoint P * M))" using dimM dimP by (mat_assoc n)
also have "\<dots> \<ge> 0" using trace_adjoint_positive by auto
finally have "trace((B - A) * \<rho>) \<ge> 0".
with eq have " trace (B * \<rho>) - trace (A * \<rho>) \<ge> 0" by auto
}
then show "\<forall>\<rho>\<in>carrier_mat n n. partial_density_operator \<rho> \<longrightarrow> trace (A * \<rho>) \<le> trace (B * \<rho>)" by auto
}
{
assume asm: "\<forall>\<rho>\<in>carrier_mat n n. partial_density_operator \<rho> \<longrightarrow> trace (A * \<rho>) \<le> trace (B * \<rho>)"
have "positive (B - A)"
proof -
{
fix v assume "dim_vec (v::complex vec) = dim_col (B - A) \<and> vec_norm v = 1"
then have dimv: "v \<in> carrier_vec n" and nv: "vec_norm v = 1"
using carrier_matD[OF dimBmA] by (auto intro: carrier_vecI)
have dimov: "outer_prod v v \<in> carrier_mat n n" using dimv by auto
then have "partial_density_operator (outer_prod v v)"
using dimv nv pure_state_self_outer_prod_is_partial_density_operator by auto
then have leq: "trace(A * (outer_prod v v)) \<le> trace(B * (outer_prod v v))" using asm dimov by auto
have "trace((B - A) * (outer_prod v v)) = trace(B * (outer_prod v v)) - trace(A * (outer_prod v v))" using A B dimov
apply (subst minus_mult_distrib_mat, auto)
apply (subst trace_minus_linear, auto)
done
then have "trace((B - A) * (outer_prod v v)) \<ge> 0" using leq by auto
then have "inner_prod v ((B - A) *\<^sub>v v) \<ge> 0" using trace_outer_prod_right[OF dimBmA dimv dimv] by auto
}
then show "positive (B - A)" using positive_iff_normalized_vec[of "B - A"] dimBmA A by simp
qed
then show "A \<le>\<^sub>L B" using lowner_le_def A B by auto
}
qed
lemma lowner_le_traceI:
assumes "A \<in> carrier_mat n n"
and "B \<in> carrier_mat n n"
and "\<And>\<rho>. \<rho> \<in> carrier_mat n n \<Longrightarrow> partial_density_operator \<rho> \<Longrightarrow> trace (A * \<rho>) \<le> trace (B * \<rho>)"
shows "A \<le>\<^sub>L B"
using lowner_le_trace assms by auto
lemma trace_pdo_eq_imp_eq:
assumes A: "A \<in> carrier_mat n n"
and B: "B \<in> carrier_mat n n"
and teq: "\<And>\<rho>. \<rho> \<in> carrier_mat n n \<Longrightarrow> partial_density_operator \<rho> \<Longrightarrow> trace (A * \<rho>) = trace (B * \<rho>)"
shows "A = B"
proof -
from teq have "A \<le>\<^sub>L B" using lowner_le_trace[OF A B] teq by auto
moreover from teq have "B \<le>\<^sub>L A" using lowner_le_trace[OF B A] teq by auto
ultimately show "A = B" using lowner_le_antisym A B by auto
qed
lemma lowner_le_traceD:
assumes "A \<in> carrier_mat n n" "B \<in> carrier_mat n n" "\<rho> \<in> carrier_mat n n"
and "A \<le>\<^sub>L B"
and "partial_density_operator \<rho>"
shows "trace (A * \<rho>) \<le> trace (B * \<rho>)"
using lowner_le_trace assms by blast
lemma sum_only_one_neq_0:
assumes "finite A" and "j \<in> A" and "\<And>i. i \<in> A \<Longrightarrow> i \<noteq> j \<Longrightarrow> g i = 0"
shows "sum g A = g j"
proof -
have "{j} \<subseteq> A" using assms by auto
moreover have "\<forall>i\<in>A - {j}. g i = 0" using assms by simp
ultimately have "sum g A = sum g {j}" using assms
by (auto simp add: comm_monoid_add_class.sum.mono_neutral_right[of A "{j}" g])
moreover have "sum g {j} = g j" by simp
ultimately show ?thesis by auto
qed
end
diff --git a/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy b/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy
--- a/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy
+++ b/thys/Relational-Incorrectness-Logic/RelationalIncorrectness.thy
@@ -1,834 +1,834 @@
theory RelationalIncorrectness
imports "HOL-IMP.Big_Step"
begin
(* Author: Toby Murray *)
section "Under-Approximate Relational Judgement"
-text {*
+text \<open>
This is the relational analogue of OHearn's~\cite{OHearn_19} and de Vries \& Koutavas'~\cite{deVries_Koutavas_11}
judgements.
Note that in our case it doesn't really make sense to talk about ``erroneous'' states: the
presence of an error can be seen only by the violation of a relation. Unlike O'Hearn, we cannot
encode it directly into the semantics of our programs, without giving them a relational semantics.
We use the standard big step semantics of IMP unchanged.
-*}
+\<close>
type_synonym rassn = "state \<Rightarrow> state \<Rightarrow> bool"
definition
ir_valid :: "rassn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> rassn \<Rightarrow> bool"
where
"ir_valid P c c' Q \<equiv> (\<forall> t t'. Q t t' \<longrightarrow> (\<exists>s s'. P s s' \<and> (c,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t'))"
section "Rules of the Logic"
definition
flip :: "rassn \<Rightarrow> rassn"
where
"flip P \<equiv> \<lambda>s s'. P s' s"
inductive
ir_hoare :: "rassn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> rassn \<Rightarrow> bool"
where
ir_Skip: "(\<And>t t'. Q t t' \<Longrightarrow> \<exists>s'. P t s' \<and> (c',s') \<Rightarrow> t') \<Longrightarrow>
ir_hoare P SKIP c' Q" |
ir_If_True: "ir_hoare (\<lambda>s s'. P s s' \<and> bval b s) c\<^sub>1 c' Q \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" |
ir_If_False: "ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s) c\<^sub>2 c' Q \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q" |
ir_Seq1: "ir_hoare P c c' Q \<Longrightarrow> ir_hoare Q d SKIP R \<Longrightarrow> ir_hoare P (Seq c d) c' R" |
ir_Assign: "ir_hoare (\<lambda>t t'. \<exists> v. P (t(x := v)) t' \<and> (t x) = aval e (t(x := v))) SKIP c' Q \<Longrightarrow>
ir_hoare P (Assign x e) c' Q" |
ir_While_False: "ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s) SKIP c' Q \<Longrightarrow>
ir_hoare P (WHILE b DO c) c' Q" |
ir_While_True: "ir_hoare (\<lambda>s s'. P s s' \<and> bval b s) (c;; WHILE b DO c) c' Q \<Longrightarrow>
ir_hoare P (WHILE b DO c) c' Q" |
ir_While_backwards_frontier: "(\<And>n. ir_hoare (\<lambda> s s'. P n s s' \<and> bval b s) c SKIP (P (Suc n))) \<Longrightarrow>
ir_hoare (\<lambda>s s'. \<exists>n. P n s s') (WHILE b DO c) c' Q \<Longrightarrow>
ir_hoare (P 0) (WHILE b DO c) c' Q" |
ir_conseq: "ir_hoare P c c' Q \<Longrightarrow> (\<And>s s'. P s s' \<Longrightarrow> P' s s') \<Longrightarrow> (\<And>s s'. Q' s s' \<Longrightarrow> Q s s') \<Longrightarrow>
ir_hoare P' c c' Q'" |
ir_disj: "ir_hoare P\<^sub>1 c c' Q\<^sub>1 \<Longrightarrow> ir_hoare P\<^sub>2 c c' Q\<^sub>2 \<Longrightarrow>
ir_hoare (\<lambda>s s'. P\<^sub>1 s s' \<or> P\<^sub>2 s s') c c' (\<lambda> t t'. Q\<^sub>1 t t' \<or> Q\<^sub>2 t t')" |
ir_sym: "ir_hoare (flip P) c c' (flip Q) \<Longrightarrow> ir_hoare P c' c Q"
section "Simple Derived Rules"
lemma meh_simp[simp]: "(SKIP, s') \<Rightarrow> t' = (s' = t')"
by auto
lemma ir_pre: "ir_hoare P c c' Q \<Longrightarrow> (\<And>s s'. P s s' \<Longrightarrow> P' s s') \<Longrightarrow>
ir_hoare P' c c' Q"
by(erule ir_conseq, blast+)
lemma ir_post: "ir_hoare P c c' Q \<Longrightarrow> (\<And>s s'. Q' s s' \<Longrightarrow> Q s s') \<Longrightarrow>
ir_hoare P c c' Q'"
by(erule ir_conseq, blast+)
section "Soundness"
lemma Skip_ir_valid[intro]:
"(\<And>t t'. Q t t' \<Longrightarrow> \<exists>s'. P t s' \<and> (c', s') \<Rightarrow> t') \<Longrightarrow> ir_valid P SKIP c' Q"
by(auto simp: ir_valid_def)
lemma sym_ir_valid[intro]:
"ir_valid (flip P) c' c (flip Q) \<Longrightarrow> ir_valid P c c' Q"
by(fastforce simp: ir_valid_def flip_def)
lemma If_True_ir_valid[intro]:
"ir_valid (\<lambda>a c. P a c \<and> bval b a) c\<^sub>1 c' Q \<Longrightarrow>
ir_valid P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q"
by(fastforce simp: ir_valid_def)
lemma If_False_ir_valid[intro]:
"ir_valid (\<lambda>a c. P a c \<and> \<not> bval b a) c\<^sub>2 c' Q \<Longrightarrow>
ir_valid P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) c' Q"
by(fastforce simp: ir_valid_def)
lemma Seq1_ir_valid[intro]:
"ir_valid P c c' Q \<Longrightarrow> ir_valid Q d SKIP R \<Longrightarrow> ir_valid P (c;; d) c' R"
by(fastforce simp: ir_valid_def)
lemma Seq2_ir_valid[intro]:
"ir_valid P c SKIP Q \<Longrightarrow> ir_valid Q d c' R \<Longrightarrow> ir_valid P (c;; d) c' R"
by(fastforce simp: ir_valid_def)
lemma Seq_ir_valid[intro]:
"ir_valid P c c' Q \<Longrightarrow> ir_valid Q d d' R \<Longrightarrow> ir_valid P (c;; d) (c';; d') R"
by(fastforce simp: ir_valid_def)
lemma Assign_blah[intro]:
"t x = aval e (t(x := v))
\<Longrightarrow> (x ::= e, t(x := v)) \<Rightarrow> t"
using Assign
by (metis fun_upd_triv fun_upd_upd)
lemma Assign_ir_valid[intro]:
"ir_valid (\<lambda>t t'. \<exists> v. P (t(x := v)) t' \<and> (t x) = aval e (t(x := v))) SKIP c' Q \<Longrightarrow> ir_valid P (Assign x e) c' Q"
by(fastforce simp: ir_valid_def)
lemma While_False_ir_valid[intro]:
"ir_valid (\<lambda>s s'. P s s' \<and> \<not> bval b s) SKIP c' Q \<Longrightarrow>
ir_valid P (WHILE b DO c) c' Q"
by(fastforce simp: ir_valid_def)
lemma While_True_ir_valid[intro]:
"ir_valid (\<lambda>s s'. P s s' \<and> bval b s) (Seq c (WHILE b DO c)) c' Q \<Longrightarrow>
ir_valid P (WHILE b DO c) c' Q"
by(clarsimp simp: ir_valid_def, blast)
lemma While_backwards_frontier_ir_valid':
assumes asm: "\<And>n. \<forall>t t'. P (k + Suc n) t t' \<longrightarrow>
(\<exists>s. P (k + n) s t' \<and> bval b s \<and> (c, s) \<Rightarrow> t)"
assumes last: "\<forall>t t'. Q t t' \<longrightarrow> (\<exists>s s'. (\<exists>n. P (k + n) s s') \<and> (WHILE b DO c, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t')"
assumes post: "Q t t'"
shows "\<exists>s s'. P k s s' \<and> (WHILE b DO c, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t'"
proof -
from post last obtain s s' n where
"P (k + n) s s'" "(WHILE b DO c, s) \<Rightarrow> t" "(c', s') \<Rightarrow> t'"
by blast
with asm show ?thesis
proof(induction n arbitrary: k t t')
case 0
then show ?case
by (metis WhileFalse WhileTrue add.right_neutral)
next
case (Suc n)
from Suc
obtain r r' where final_iteration: "P (Suc k) r r'" "(WHILE b DO c, r) \<Rightarrow> t" "(c', r') \<Rightarrow> t'"
by (metis add_Suc_shift)
from final_iteration(1) obtain q q' where
"P k q r' \<and> bval b q \<and> (c, q) \<Rightarrow> r"
by (metis Nat.add_0_right Suc.prems(1) plus_1_eq_Suc semiring_normalization_rules(24))
with final_iteration show ?case by blast
qed
qed
lemma While_backwards_frontier_ir_valid[intro]:
"(\<And>n. ir_valid (\<lambda> s s'. P n s s' \<and> bval b s) c SKIP (P (Suc n))) \<Longrightarrow>
ir_valid (\<lambda>s s'. \<exists>n. P n s s') (WHILE b DO c) c' Q \<Longrightarrow>
ir_valid (P 0) (WHILE b DO c) c' Q"
by(auto simp: meh_simp ir_valid_def intro: While_backwards_frontier_ir_valid')
lemma conseq_ir_valid:
"ir_valid P c c' Q \<Longrightarrow> (\<And>s s'. P s s' \<Longrightarrow> P' s s') \<Longrightarrow> (\<And>s s'. Q' s s' \<Longrightarrow> Q s s') \<Longrightarrow>
ir_valid P' c c' Q'"
by(clarsimp simp: ir_valid_def, blast)
lemma disj_ir_valid[intro]:
"ir_valid P\<^sub>1 c c' Q\<^sub>1 \<Longrightarrow> ir_valid P\<^sub>2 c c' Q\<^sub>2 \<Longrightarrow>
ir_valid (\<lambda>s s'. P\<^sub>1 s s' \<or> P\<^sub>2 s s') c c' (\<lambda> t t'. Q\<^sub>1 t t' \<or> Q\<^sub>2 t t')"
by(fastforce simp: ir_valid_def)
theorem soundness:
"ir_hoare P c c' Q \<Longrightarrow> ir_valid P c c' Q"
apply(induction rule: ir_hoare.induct)
apply(blast intro!: Skip_ir_valid)
by (blast intro: conseq_ir_valid)+
section "Completeness"
lemma ir_Skip_Skip[intro]:
"ir_hoare P SKIP SKIP P"
by(rule ir_Skip, simp)
lemma ir_hoare_Skip_Skip[simp]:
"ir_hoare P SKIP SKIP Q = (\<forall>s s'. Q s s' \<longrightarrow> P s s')"
using soundness ir_valid_def meh_simp ir_conseq ir_Skip by metis
lemma ir_valid_Seq1:
"ir_valid P (c1;; c2) c' Q \<Longrightarrow> ir_valid P c1 c' (\<lambda>t t'. \<exists>s s'. P s s' \<and> (c1,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t' \<and> (\<exists>u. (c2,t) \<Rightarrow> u \<and> Q u t'))"
by(auto simp: ir_valid_def)
lemma ir_valid_Seq1':
"ir_valid P (c1;; c2) c' Q \<Longrightarrow> ir_valid (\<lambda>t t'. \<exists>s s'. P s s' \<and> (c1,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t' \<and> (\<exists>u. (c2,t) \<Rightarrow> u \<and> Q u t')) c2 SKIP Q"
by(clarsimp simp: ir_valid_def, meson SeqE)
lemma ir_valid_track_history:
"ir_valid P c c' Q \<Longrightarrow>
ir_valid P c c' (\<lambda>t t'. Q s s' \<and> (\<exists>s s'. P s s' \<and> (c,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t'))"
by(auto simp: ir_valid_def)
lemma ir_valid_If:
"ir_valid (\<lambda>s s'. P s s') (IF b THEN c1 ELSE c2) c' Q \<Longrightarrow>
ir_valid (\<lambda>s s'. P s s' \<and> bval b s) c1 c' (\<lambda>t t'. Q t t' \<and> (\<exists>s s'. P s s' \<and> (c1,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t' \<and> bval b s)) \<and>
ir_valid (\<lambda>s s'. P s s' \<and> \<not> bval b s) c2 c' (\<lambda>t t'. Q t t' \<and> (\<exists>s s'. P s s' \<and> (c2,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t' \<and> \<not> bval b s))"
by(clarsimp simp: ir_valid_def, blast)
-text {*
+text \<open>
Inspired by the
``@{text "p(n) = {\<sigma> | you can get back from \<sigma> to some state in p by executing C backwards n times}"}''
part of OHearn~\cite{OHearn_19}.
-*}
+\<close>
primrec get_back where
"get_back P b c 0 = (\<lambda>t t'. P t t')" |
"get_back P b c (Suc n) = (\<lambda>t t'. \<exists>s. (c,s) \<Rightarrow> t \<and> bval b s \<and> get_back P b c n s t')"
(* Currently not used anywhere *)
lemma ir_valid_get_back:
"ir_valid (get_back P b c (Suc k)) (WHILE b DO c) c' Q \<Longrightarrow>
ir_valid (get_back P b c k) (WHILE b DO c) c' Q"
proof(induct k)
case 0
then show ?case by(clarsimp simp: ir_valid_def, blast)
next
case (Suc k)
then show ?case using WhileTrue get_back.simps(2) ir_valid_def by smt
qed
(* both this an the next one get used in the completeness proof *)
lemma ir_valid_While1:
"ir_valid (get_back P b c k) (WHILE b DO c) c' Q \<Longrightarrow>
(ir_valid (\<lambda>s s'. get_back P b c k s s' \<and> bval b s) c SKIP (\<lambda>t t'. get_back P b c (Suc k) t t' \<and> (\<exists>u u'. (WHILE b DO c,t) \<Rightarrow> u \<and> (c',t') \<Rightarrow> u' \<and> Q u u')))"
proof (subst ir_valid_def, clarsimp)
fix t t' s u u'
assume "ir_valid (get_back P b c k) (WHILE b DO c) c' Q"
"(WHILE b DO c, t) \<Rightarrow> u"
"(c, s) \<Rightarrow> t"
"(c', t') \<Rightarrow> u'"
"Q u u'"
"bval b s"
"get_back P b c k s t'"
thus "\<exists>s. get_back P b c k s t' \<and> bval b s \<and> (c, s) \<Rightarrow> t"
proof(induction k arbitrary: t t' s u u')
case 0
then show ?case
by auto
next
case (Suc k)
show ?case
using Suc.prems(3) Suc.prems(6) Suc.prems(7) by blast
qed
qed
lemma ir_valid_While3:
"ir_valid (get_back P b c k) (WHILE b DO c) c' Q \<Longrightarrow>
(ir_valid (\<lambda>s s'. get_back P b c k s s' \<and> bval b s) c c' (\<lambda>t t'. (\<exists>s'. (c',s') \<Rightarrow> t' \<and> get_back P b c (Suc k) t s' \<and> (\<exists>u. (WHILE b DO c,t) \<Rightarrow> u \<and> Q u t'))))"
apply(subst ir_valid_def, clarsimp)
proof -
fix t t' s' s u
assume "ir_valid (get_back P b c k) (WHILE b DO c) c' Q"
"(WHILE b DO c, t) \<Rightarrow> u"
"(c, s) \<Rightarrow> t"
"(c', s') \<Rightarrow> t'"
"Q u t'"
"bval b s"
"get_back P b c k s s'"
thus "\<exists>s s'. get_back P b c k s s' \<and> bval b s \<and> (c, s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t'"
proof(induction k arbitrary: t t' s' s u)
case 0
then show ?case
by auto
next
case (Suc k)
show ?case
using Suc.prems(3) Suc.prems(4) Suc.prems(6) Suc.prems(7) by blast
qed
qed
(* not used anywhere *)
lemma ir_valid_While2:
"ir_valid P (WHILE b DO c) c' Q \<Longrightarrow>
ir_valid (\<lambda>s s'. P s s' \<and> \<not> bval b s) SKIP c' (\<lambda>t t'. Q t t' \<and> (\<exists>s'. (c',s') \<Rightarrow> t' \<and> P t s' \<and> \<not> bval b t))"
by(clarsimp simp: ir_valid_def, blast)
lemma assign_upd_blah:
"(\<lambda>a. if a = x1 then s x1 else (s(x1 := aval x2 s)) a) = s"
by(rule ext, auto)
lemma Assign_complete:
assumes v: "ir_valid P (x1 ::= x2) c' Q"
assumes q: "Q t t'"
shows "\<exists>s'. (\<exists>v. P (t(x1 := v)) s' \<and> t x1 = aval x2 (t(x1 := v))) \<and> (c', s') \<Rightarrow> t'"
proof -
from v and q obtain s s' where a: "P s s' \<and> (x1 ::= x2,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t'"
using ir_valid_def by smt
hence "P (\<lambda>a. if a = x1 then s x1 else (s(x1 := aval x2 s)) a) s' \<and> aval x2 s = aval x2 (s(x1 := s x1))"
using assign_upd_blah
by simp
thus ?thesis
using assign_upd_blah a
by (metis AssignE fun_upd_same fun_upd_triv fun_upd_upd)
qed
lemmas ir_Skip_sym = ir_sym[OF ir_Skip, simplified flip_def]
theorem completeness:
"ir_valid P c c' Q \<Longrightarrow> ir_hoare P c c' Q"
proof(induct c arbitrary: P c' Q)
case SKIP
show ?case
apply(rule ir_Skip)
using SKIP by(fastforce simp: ir_valid_def)
next
case (Assign x1 x2)
show ?case
apply(rule ir_Assign)
apply(rule ir_Skip)
using Assign_complete Assign by blast
next
case (Seq c1 c2)
have a: "ir_hoare P c1 c' (\<lambda>t t'. \<exists>s s'. P s s' \<and> (c1, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t' \<and> (\<exists>u. (c2, t) \<Rightarrow> u \<and> Q u t'))"
using ir_valid_Seq1 Seq by blast
show ?case
apply(rule ir_Seq1)
apply (blast intro: a)
apply(rule ir_Skip_sym)
by(metis SeqE ir_valid_def Seq)
next
case (If x1 c1 c2)
have t: "ir_hoare (\<lambda>s s'. P s s' \<and> bval x1 s) c1 c'
(\<lambda>t t'. Q t t' \<and> (\<exists>s s'. P s s' \<and> (c1, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t' \<and> bval x1 s))" and
f: " ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval x1 s) c2 c'
(\<lambda>t t'. Q t t' \<and> (\<exists>s s'. P s s' \<and> (c2, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t' \<and> \<not> bval x1 s))"
using ir_valid_If If by blast+
show ?case
(* consider both cases of the if via conseq, disj, then _True and _False *)
apply(rule ir_conseq)
apply(rule ir_disj)
apply(rule ir_If_True,fastforce intro: t)
apply(rule ir_If_False, fastforce intro: f)
apply blast
by (smt IfE ir_valid_def If)
next
case (While x1 c)
have a: "\<And>n. ir_hoare (\<lambda>s s'. get_back P x1 c n s s' \<and> bval x1 s) c SKIP (get_back P x1 c (Suc n))"
using ir_valid_While1 While
by (smt get_back.simps(2) ir_valid_def meh_simp)
have b: "ir_hoare (\<lambda>s s'. P s s' \<and> bval x1 s) c c'
(\<lambda>t t'. \<exists>s'. (c', s') \<Rightarrow> t' \<and> (\<exists>s. (c, s) \<Rightarrow> t \<and> bval x1 s \<and> P s s') \<and>
(\<exists>u. (WHILE x1 DO c, t) \<Rightarrow> u \<and> Q u t'))"
using ir_valid_While3[where k=0,simplified] While by blast
have gb: "\<And>t t'. Q t t' \<and> (\<exists>s'. (c', s') \<Rightarrow> t' \<and> P t s' \<and> \<not> bval x1 t) \<Longrightarrow>
\<exists>s'. ((\<exists>n. get_back P x1 c n t s') \<and> \<not> bval x1 t) \<and> (c', s') \<Rightarrow> t'"
by (meson get_back.simps(1))
show ?case
(* use the frontier rule much as in OHearn POPL *)
apply(rule ir_conseq)
apply(rule_tac P="get_back P x1 c" in ir_While_backwards_frontier)
apply(blast intro: a)
(* consider both cases of the While via conseq, disj, then _True and _False *)
apply(rule ir_conseq)
apply(rule ir_disj)
apply(rule_tac P="\<lambda>s s'. \<exists>n. get_back P x1 c n s s'" and Q="(\<lambda>t t'. Q t t' \<and> (\<exists>s'. (c', s') \<Rightarrow> t' \<and> P t s' \<and> \<not> bval x1 t))" in ir_While_False)
apply(rule ir_Skip, blast intro: gb)
apply(rule ir_While_True)
apply(rule ir_Seq1[OF b])
apply(rule ir_Skip_sym)
apply(fastforce)
apply (metis get_back.simps(1))
apply assumption
apply simp
by (metis While.prems WhileE ir_valid_def)
qed
section "A Decomposition Principle: Proofs via Under-Approximate Hoare Logic"
-text {*
+text \<open>
We show the under-approximate analogue holds for Beringer's~\cite{Beringer_11} decomposition
principle for over-approximate relational logic.
-*}
+\<close>
definition
decomp :: "rassn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> rassn \<Rightarrow> rassn" where
"decomp P c c' Q \<equiv> \<lambda>t s'. \<exists>s t'. P s s' \<and> (c,s) \<Rightarrow> t \<and> (c',s') \<Rightarrow> t' \<and> Q t t'"
lemma ir_valid_decomp1:
"ir_valid P c c' Q \<Longrightarrow> ir_valid P c SKIP (decomp P c c' Q) \<and> ir_valid (decomp P c c' Q) SKIP c' Q"
by(fastforce simp: ir_valid_def meh_simp decomp_def)
lemma ir_valid_decomp2:
"ir_valid P c SKIP R \<and> ir_valid R SKIP c' Q \<Longrightarrow> ir_valid P c c' Q"
by(fastforce simp: ir_valid_def meh_simp decomp_def)
lemma ir_valid_decomp:
"ir_valid P c c' Q = (ir_valid P c SKIP (decomp P c c' Q) \<and> ir_valid (decomp P c c' Q) SKIP c' Q)"
using ir_valid_decomp1 ir_valid_decomp2 by blast
-text {*
+text \<open>
Completeness with soundness means we can prove proof rules about @{term ir_hoare} in terms
of @{term ir_valid}.
-*}
+\<close>
lemma ir_to_Skip:
"ir_hoare P c c' Q = (ir_hoare P c SKIP (decomp P c c' Q) \<and> ir_hoare (decomp P c c' Q) SKIP c' Q)"
using soundness completeness ir_valid_decomp
by meson
-text {*
+text \<open>
O'Hearn's under-approximate Hoare triple, for the ``ok'' case (since that is the only case we have)
This is also likely the same as from the "Reverse Hoare Logic" paper (SEFM).
-*}
+\<close>
type_synonym assn = "state \<Rightarrow> bool"
definition
ohearn :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"
where
"ohearn P c Q \<equiv> (\<forall>t. Q t \<longrightarrow> (\<exists>s. P s \<and> (c,s) \<Rightarrow> t))"
lemma fold_ohearn1:
"ir_valid P c SKIP Q = (\<forall>t'. ohearn (\<lambda>t. P t t') c (\<lambda>t. Q t t'))"
by(fastforce simp add: ir_valid_def ohearn_def)
lemma fold_ohearn2:
"ir_valid P SKIP c' Q = (\<forall>t. ohearn (P t) c' (Q t))"
by(simp add: ir_valid_def ohearn_def)
theorem relational_via_hoare:
"ir_hoare P c c' Q = ((\<forall>t'. ohearn (\<lambda>t. P t t') c (\<lambda>t. decomp P c c' Q t t')) \<and> (\<forall>t. ohearn (decomp P c c' Q t) c' (Q t)))"
proof -
have a: "\<And>P c c' Q. ir_hoare P c c' Q = ir_valid P c c' Q"
using soundness completeness by auto
show ?thesis
using ir_to_Skip a fold_ohearn1 fold_ohearn2 by metis
qed
section "Deriving Proof Rules from Completeness"
-text {*
+text \<open>
Note that we can more easily derive proof rules sometimes by appealing to the
corresponding properties of @{term ir_valid} than from the proof rules directly.
We see more examples of this later on when we consider examples.
-*}
+\<close>
lemma ir_Seq2:
"ir_hoare P c SKIP Q \<Longrightarrow> ir_hoare Q d c' R \<Longrightarrow> ir_hoare P (Seq c d) c' R"
using soundness completeness Seq2_ir_valid by metis
lemma ir_Seq:
"ir_hoare P c c' Q \<Longrightarrow> ir_hoare Q d d' R \<Longrightarrow> ir_hoare P (Seq c d) (Seq c' d') R"
using soundness completeness Seq_ir_valid by metis
section "Examples"
subsection "Some Derived Proof Rules"
-text {*
+text \<open>
First derive some proof rules -- here not by appealing to completeness but just using
the existing rules
-*}
+\<close>
lemma ir_If_True_False:
"ir_hoare (\<lambda>s s'. P s s' \<and> bval b s \<and> \<not> bval b' s') c\<^sub>1 c\<^sub>2' Q \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') Q"
apply(rule ir_If_True)
apply(rule ir_sym)
apply(rule ir_If_False)
apply(rule ir_sym)
by(simp add: flip_def)
lemma ir_Assign_Assign:
"ir_hoare P (x ::= e) (x' ::= e') (\<lambda>t t'. \<exists>v v'. P (t(x := v)) (t'(x' := v')) \<and> t x = aval e (t(x := v)) \<and> t' x' = aval e' (t'(x' := v')))"
apply(rule ir_Assign)
apply(rule ir_sym)
apply(rule ir_Assign)
by(simp add: flip_def, auto)
subsection "prog1"
-text {*
+text \<open>
A tiny insecure program. Note that we only need to reason on one path through this program to
detect that it is insecure.
-*}
+\<close>
abbreviation low_eq :: rassn where "low_eq s s' \<equiv> s ''low'' = s' ''low''"
abbreviation low_neq :: rassn where "low_neq s s' \<equiv> \<not> low_eq s s'"
definition prog1 :: com where
"prog1 \<equiv> (IF (Less (N 0) (V ''x'')) THEN (Assign ''low'' (N 1)) ELSE (Assign ''low'' (N 0)))"
-text {*
+text \<open>
We prove that @{term prog1} is definitely insecure. To do that, we need to find some non-empty
post-relation that implies insecurity. The following property encodes the idea that the
post-relation is non-empty, i.e. represents a feasible pair of execution paths.
-*}
+\<close>
definition
nontrivial :: "rassn \<Rightarrow> bool"
where
"nontrivial Q \<equiv> (\<exists>t t'. Q t t')"
-text {*
+text \<open>
Note the property we prove here explicitly encodes the fact that the postcondition can be anything
that implies insecurity, i.e. implies @{term low_neq}. In particular we should not necessarily
expect it to cover the entirely of all states that satisfy @{term low_neq}.
Also note that we also have to prove that the postcondition is non-trivial. This is necessary to
make sure that the violation we find is not an infeasible path.
-*}
+\<close>
lemma prog1:
"\<exists>Q. ir_hoare low_eq prog1 prog1 Q \<and> (\<forall>s s'. Q s s' \<longrightarrow> low_neq s s') \<and> nontrivial Q"
apply(rule exI)
apply(rule conjI)
apply(simp add: prog1_def)
apply(rule ir_If_True_False)
apply(rule ir_Assign_Assign)
apply(rule conjI)
apply auto[1]
apply(clarsimp simp: nontrivial_def)
apply(rule_tac x="\<lambda>v. 1" in exI)
apply simp
apply(rule_tac x="\<lambda>v. 0" in exI)
by auto
subsection "More Derived Proof Rules for Examples"
definition BEq :: "aexp \<Rightarrow> aexp \<Rightarrow> bexp" where
"BEq a b \<equiv> And (Less a (Plus b (N 1))) (Less b (Plus a (N 1)))"
lemma BEq_aval[simp]:
"bval (BEq a b) s = ((aval a s) = (aval b s))"
by(auto simp add: BEq_def)
lemma ir_If_True_True:
"ir_hoare (\<lambda>s s'. P s s' \<and> bval b s \<and> bval b' s') c\<^sub>1 c\<^sub>1' Q\<^sub>1 \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\<lambda>t t'. Q\<^sub>1 t t')"
by(simp add: ir_If_True ir_sym flip_def)
lemma ir_If_False_False:
"ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s \<and> \<not> bval b' s') c\<^sub>2 c\<^sub>2' Q\<^sub>2 \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\<lambda>t t'. Q\<^sub>2 t t')"
by(simp add: ir_If_False ir_sym flip_def)
lemma ir_If':
"ir_hoare (\<lambda>s s'. P s s' \<and> bval b s \<and> bval b' s') c\<^sub>1 c\<^sub>1' Q\<^sub>1 \<Longrightarrow>
ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s \<and> \<not> bval b' s') c\<^sub>2 c\<^sub>2' Q\<^sub>2 \<Longrightarrow>
ir_hoare P (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (IF b' THEN c\<^sub>1' ELSE c\<^sub>2') (\<lambda>t t'. Q\<^sub>1 t t' \<or> Q\<^sub>2 t t')"
apply(rule ir_pre)
apply(rule ir_disj)
apply(rule ir_If_True_True)
apply assumption
apply(rule ir_If_False_False)
apply assumption
apply blast
done
lemma ir_While_triv:
"ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s \<and> \<not> bval b' s') SKIP SKIP Q\<^sub>2 \<Longrightarrow>
ir_hoare P (WHILE b DO c) (WHILE b' DO c') (\<lambda>s s'. (Q\<^sub>2 s s'))"
by(simp add: ir_While_False ir_sym flip_def)
lemma ir_While':
"ir_hoare (\<lambda>s s'. P s s' \<and> bval b s \<and> bval b' s') (c;;WHILE b DO c) (c';;WHILE b' DO c') Q\<^sub>1 \<Longrightarrow>
ir_hoare (\<lambda>s s'. P s s' \<and> \<not> bval b s \<and> \<not> bval b' s') SKIP SKIP Q\<^sub>2 \<Longrightarrow>
ir_hoare P (WHILE b DO c) (WHILE b' DO c') (\<lambda>s s'. (Q\<^sub>1 s s' \<or> Q\<^sub>2 s s'))"
apply(rule ir_pre)
apply(rule ir_disj)
apply(rule ir_While_True)
apply(rule ir_sym)
apply(simp add: flip_def)
apply(rule ir_While_True)
apply(rule ir_sym)
apply(simp add: flip_def)
apply(rule ir_While_triv)
apply assumption
apply simp
done
subsection "client0"
definition low_eq_strong where
"low_eq_strong s s' \<equiv> (s ''high'' \<noteq> s' ''high'') \<and> low_eq s s'"
lemma low_eq_strong_upd[simp]:
"var \<noteq> ''high'' \<and> var \<noteq> ''low'' \<Longrightarrow> low_eq_strong(s(var := v)) (s'(var := v')) = low_eq_strong s s'"
by(auto simp: low_eq_strong_def)
-text {*
+text \<open>
A variation on client0 from O'Hearn~\cite{OHearn_19}: how to reason about loops via a single unfolding
-*}
+\<close>
definition client0 :: com where
"client0 \<equiv> (Assign ''x'' (N 0);;
(While (Less (N 0) (V ''n''))
((Assign ''x'' (Plus (V ''x'') (V ''n'')));;
(Assign ''n'' (V ''nondet''))));;
(If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP))"
lemma client0:
"\<exists>Q. ir_hoare low_eq client0 client0 Q \<and> (\<forall>s s'. Q s s' \<longrightarrow> low_neq s s') \<and> nontrivial Q"
apply(rule exI, rule conjI, simp add: client0_def)
apply(rule_tac P=low_eq_strong in ir_pre)
apply(rule ir_Seq)
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_While')
apply clarsimp
apply(rule ir_Seq)
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_While_triv)
apply clarsimp
apply assumption
apply clarsimp
apply assumption
apply(rule ir_If_True_True)
apply(rule ir_Assign_Assign)
apply(fastforce simp: low_eq_strong_def)
apply(rule conjI)
apply(clarsimp simp: low_eq_strong_def split: if_splits)
(* ugh having to manually do constraint solving here... *)
apply(clarsimp simp: low_eq_strong_def nontrivial_def)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 2000000 else if v = ''high'' then 1 else if v = ''n'' then -1 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 2000000 else if v = ''high'' then 0 else if v = ''n'' then -1 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI)
apply clarsimp
done
(* Not needed? *)
lemma ir_While_backwards:
"(\<And>n. ir_hoare (\<lambda> s s'. P n s s' \<and> bval b s) c SKIP (P (Suc n))) \<Longrightarrow>
ir_hoare (\<lambda>s s'. \<exists>n. P n s s' \<and> \<not> bval b s) SKIP c' Q \<Longrightarrow>
ir_hoare (P 0) (WHILE b DO c) c' Q"
apply(rule ir_While_backwards_frontier)
apply assumption
apply(rule ir_While_False)
apply auto
done
subsection "Derive a variant of the backwards variant rule"
-text {* Here we appeal to completeness again to derive this rule from the corresponding
- property about @{term ir_valid}. *}
+text \<open>Here we appeal to completeness again to derive this rule from the corresponding
+ property about @{term ir_valid}.\<close>
subsection "A variant of the frontier rule"
-text {*
+text \<open>
Agin we derive this rule by appealing to completeness and the corresponding property of
@{term ir_valid}
-*}
+\<close>
lemma While_backwards_frontier_both_ir_valid':
assumes asm: "\<And>n. \<forall>t t'. P (k + Suc n) t t' \<longrightarrow>
(\<exists>s s'. P (k + n) s s' \<and> bval b s \<and> bval b' s' \<and> (c, s) \<Rightarrow> t \<and> (c', s') \<Rightarrow> t')"
assumes last: "\<forall>t t'. Q t t' \<longrightarrow> (\<exists>s s'. (\<exists>n. P (k + n) s s') \<and> (WHILE b DO c, s) \<Rightarrow> t \<and> (WHILE b' DO c', s') \<Rightarrow> t')"
assumes post: "Q t t'"
shows "\<exists>s s'. P k s s' \<and> (WHILE b DO c, s) \<Rightarrow> t \<and> (WHILE b' DO c', s') \<Rightarrow> t'"
proof -
from post last obtain s s' n where
"P (k + n) s s'" "(WHILE b DO c, s) \<Rightarrow> t" "(WHILE b' DO c', s') \<Rightarrow> t'"
by blast
with asm show ?thesis
proof(induction n arbitrary: k t t')
case 0
then show ?case
by (metis WhileFalse WhileTrue add.right_neutral)
next
case (Suc n)
from Suc
obtain r r' where final_iteration: "P (Suc k) r r'" "(WHILE b DO c, r) \<Rightarrow> t" "(WHILE b' DO c', r') \<Rightarrow> t'"
by (metis add_Suc_shift)
from final_iteration(1) obtain q q' where
"P k q q' \<and> bval b q \<and> bval b' q' \<and> (c, q) \<Rightarrow> r \<and> (c', q') \<Rightarrow> r'"
by (metis Nat.add_0_right Suc.prems(1) plus_1_eq_Suc semiring_normalization_rules(24))
with final_iteration show ?case by blast
qed
qed
lemma While_backwards_frontier_both_ir_valid[intro]:
"(\<And>n. ir_valid (\<lambda> s s'. P n s s' \<and> bval b s \<and> bval b' s') c c' (P (Suc n))) \<Longrightarrow>
ir_valid (\<lambda>s s'. \<exists>n. P n s s') (WHILE b DO c) (WHILE b' DO c') Q \<Longrightarrow>
ir_valid (P 0) (WHILE b DO c) (WHILE b' DO c') (\<lambda>s s'. Q s s')"
by(auto simp: ir_valid_def intro: While_backwards_frontier_both_ir_valid')
lemma ir_While_backwards_frontier_both:
"\<lbrakk>\<And>n. ir_hoare (\<lambda>s s'. P n s s' \<and> bval b s \<and> bval b' s') c c' (P (Suc n));
ir_hoare (\<lambda>s s'. \<exists>n. P n s s') (WHILE b DO c) (WHILE b' DO c') Q\<rbrakk>
\<Longrightarrow> ir_hoare (P 0) (WHILE b DO c) (WHILE b' DO c') (\<lambda>s s'. Q s s')"
using soundness completeness While_backwards_frontier_both_ir_valid by auto
-text {*
+text \<open>
The following rule then follows easily as a special case
-*}
+\<close>
lemma ir_While_backwards_both:
"(\<And>n. ir_hoare (\<lambda> s s'. P n s s' \<and> bval b s \<and> bval b' s') c c' (P (Suc n))) \<Longrightarrow>
ir_hoare (P 0) (WHILE b DO c) (WHILE b' DO c') (\<lambda>s s'. \<exists>n. P n s s' \<and> \<not> bval b s \<and> \<not> bval b' s')"
apply(rule ir_While_backwards_frontier_both)
apply blast
by(simp add: ir_While_False ir_sym flip_def)
subsection "client1"
-text {*
+text \<open>
An example roughly equivalent to cient1 from O'Hearn~\cite{OHearn_19}0
In particular we use the backwards variant rule to reason about the loop.
-*}
+\<close>
definition client1 :: com where
"client1 \<equiv> (Assign ''x'' (N 0);;
(While (Less (V ''x'') (V ''n''))
((Assign ''x'' (Plus (V ''x'') (N 1)))));;
(If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP))"
lemma client1:
"\<exists>Q. ir_hoare low_eq client1 client1 Q \<and> (\<forall>s s'. Q s s' \<longrightarrow> low_neq s s') \<and> nontrivial Q"
apply(rule exI, rule conjI, simp add: client1_def)
apply(rule_tac P=low_eq_strong in ir_pre)
apply(rule ir_Seq)
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_pre)
apply(rule ir_While_backwards_both[where P="\<lambda>n s s'. low_eq_strong s s' \<and> s ''x'' = int n \<and> s' ''x'' = int n \<and> int n \<le> s ''n'' \<and> int n \<le> s' ''n''"])
apply(rule ir_post)
apply(rule ir_Assign_Assign)
apply clarsimp
apply clarsimp
apply(rule ir_If_True_True)
apply(rule ir_Assign_Assign)
apply(fastforce simp: low_eq_strong_def)
apply(rule conjI)
apply(clarsimp simp: low_eq_strong_def split: if_splits)
apply clarsimp
(* ugh having to manually do constraint solving here... *)
apply(clarsimp simp: low_eq_strong_def nontrivial_def)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 2000000 else if v = ''high'' then 1 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 2000000 else if v = ''high'' then 0 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI)
apply clarsimp
done
subsection "client2"
-text {*
+text \<open>
An example akin to client2 from O'Hearn~\cite{OHearn_19}.
Note that this example is carefully written to show use of the frontier rule first to
reason up to the broken loop iteration, and then we unfold the loop at that point to
reason about the broken iteration, and then use the plain backwards variant rule to
reason over the remainder of the loop.
-*}
+\<close>
definition client2 :: com where
"client2 \<equiv> (Assign ''x'' (N 0);;
(While (Less (V ''x'') (N 4000000))
((Assign ''x'' (Plus (V ''x'') (N 1)));;
(If (BEq (V ''x'') (N 2000000)) (Assign ''low'' (V ''high'')) SKIP))
)
)"
lemma client2:
"\<exists>Q. ir_hoare low_eq client2 client2 Q \<and> (\<forall>s s'. Q s s' \<longrightarrow> low_neq s s') \<and> nontrivial Q"
apply(rule exI, rule conjI, simp add: client2_def)
apply(rule_tac P=low_eq_strong in ir_pre)
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_pre)
apply(rule ir_While_backwards_frontier_both[where P="\<lambda>n s s'. low_eq_strong s s' \<and> s ''x'' = int n \<and> s' ''x'' = int n \<and> s ''x'' \<ge> 0 \<and> s ''x'' \<le> 2000000 - 1 \<and> s' ''x'' \<ge> 0 \<and> s' ''x'' \<le> 2000000 - 1"])
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_post)
apply(rule ir_If')
apply(rule ir_Assign_Assign)
apply(rule ir_Skip_Skip)
apply clarsimp
apply clarsimp
apply(rule ir_While')
apply clarsimp
apply(rule ir_Seq)
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply(rule ir_If_True_True)
apply(rule ir_Assign_Assign)
apply clarsimp
apply(rule ir_pre)
apply(rule ir_While_backwards_both[where P="\<lambda>n s s'. s ''x'' = 2000000 + int n \<and> s' ''x'' = 2000000 + int n \<and> s ''x'' \<ge> 2000000 \<and> s ''x'' \<le> 4000000 \<and> s' ''x'' \<ge> 2000000 \<and> s' ''x'' \<le> 4000000 \<and> s ''low'' = s ''high'' \<and> s' ''low'' = s' ''high'' \<and> s ''high'' \<noteq> s' ''high''"])
apply(rule ir_Seq)
apply(rule ir_Assign_Assign)
apply(rule ir_If_False_False)
apply fastforce
apply (fastforce simp: low_eq_strong_def)
apply fastforce
apply(fastforce simp: low_eq_strong_def)
apply(fastforce simp: low_eq_strong_def)
apply(rule conjI)
apply(clarsimp simp: low_eq_strong_def split: if_splits)
apply clarsimp
(* ugh having to manually do constraint solving here... *)
apply(clarsimp simp: low_eq_strong_def nontrivial_def)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 4000000 else if v = ''high'' then 1 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 1 else undefined" in exI)
apply(rule_tac x="\<lambda>v. if v = ''x'' then 4000000 else if v = ''high'' then 0 else if v = ''n'' then 2000000 else if v = ''nondet'' then -1 else if v = ''low'' then 0 else undefined" in exI)
apply clarsimp
done
end
diff --git a/thys/Relational_Method/Anonymity.thy b/thys/Relational_Method/Anonymity.thy
--- a/thys/Relational_Method/Anonymity.thy
+++ b/thys/Relational_Method/Anonymity.thy
@@ -1,599 +1,599 @@
(* Title: The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Anonymity properties"
theory Anonymity
imports Authentication
begin
text \<open>
\label{Anonymity}
\<close>
proposition crypts_empty [simp]:
"crypts {} = {}"
by (rule equalityI, rule subsetI, erule crypts.induct, simp_all)
proposition crypts_mono:
"H \<subseteq> H' \<Longrightarrow> crypts H \<subseteq> crypts H'"
by (rule subsetI, erule crypts.induct, auto)
lemma crypts_union_1:
"crypts (H \<union> H') \<subseteq> crypts H \<union> crypts H'"
by (rule subsetI, erule crypts.induct, auto)
lemma crypts_union_2:
"crypts H \<union> crypts H' \<subseteq> crypts (H \<union> H')"
by (rule subsetI, erule UnE, erule_tac [!] crypts.induct, auto)
proposition crypts_union:
"crypts (H \<union> H') = crypts H \<union> crypts H'"
by (rule equalityI, rule crypts_union_1, rule crypts_union_2)
proposition crypts_insert:
"crypts (insert X H) = crypts_msg X \<union> crypts H"
by (simp only: insert_def crypts_union, subst crypts_msg_def, simp)
proposition crypts_msg_num [simp]:
"crypts_msg (Num n) = {Num n}"
by (subst crypts_msg_def, rule equalityI, rule subsetI, erule crypts.induct, simp,
rotate_tac [1-3], erule_tac [1-3] rev_mp, rule_tac [1-3] list.induct, simp_all,
blast)
proposition crypts_msg_agent [simp]:
"crypts_msg (Agent n) = {Agent n}"
by (subst crypts_msg_def, rule equalityI, rule subsetI, erule crypts.induct, simp,
rotate_tac [1-3], erule_tac [1-3] rev_mp, rule_tac [1-3] list.induct, simp_all,
blast)
proposition crypts_msg_pwd [simp]:
"crypts_msg (Pwd n) = {Pwd n}"
by (subst crypts_msg_def, rule equalityI, rule subsetI, erule crypts.induct, simp,
rotate_tac [1-3], erule_tac [1-3] rev_mp, rule_tac [1-3] list.induct, simp_all,
blast)
proposition crypts_msg_key [simp]:
"crypts_msg (Key K) = {Key K}"
by (subst crypts_msg_def, rule equalityI, rule subsetI, erule crypts.induct, simp,
rotate_tac [1-3], erule_tac [1-3] rev_mp, rule_tac [1-3] list.induct, simp_all,
blast)
proposition crypts_msg_mult [simp]:
"crypts_msg (A \<otimes> B) = {A \<otimes> B}"
by (subst crypts_msg_def, rule equalityI, rule subsetI, erule crypts.induct, simp,
rotate_tac [1-3], erule_tac [1-3] rev_mp, rule_tac [1-3] list.induct, simp_all,
blast)
lemma crypts_hash_1:
"crypts {Hash X} \<subseteq> insert (Hash X) (crypts {X})"
by (rule subsetI, erule crypts.induct, simp_all, (erule disjE, rotate_tac, erule rev_mp,
rule list.induct, simp_all, blast, (drule crypts_hash, simp)?)+)
lemma crypts_hash_2:
"insert (Hash X) (crypts {X}) \<subseteq> crypts {Hash X}"
by (rule subsetI, simp, erule disjE, blast, erule crypts.induct, simp,
subst id_apply [symmetric], subst foldr_Nil [symmetric], rule crypts_hash, simp,
blast+)
proposition crypts_msg_hash [simp]:
"crypts_msg (Hash X) = insert (Hash X) (crypts_msg X)"
by (simp add: crypts_msg_def, rule equalityI, rule crypts_hash_1, rule crypts_hash_2)
proposition crypts_comp:
"X \<in> crypts H \<Longrightarrow> Crypt K X \<in> crypts (Crypt K ` H)"
by (erule crypts.induct, blast, (simp only: comp_apply
[symmetric, where f = "Crypt K"] foldr_Cons [symmetric],
(erule crypts_hash | erule crypts_fst | erule crypts_snd))+)
lemma crypts_crypt_1:
"crypts {Crypt K X} \<subseteq> Crypt K ` crypts {X}"
by (rule subsetI, erule crypts.induct, fastforce, rotate_tac [!], erule_tac [!] rev_mp,
rule_tac [!] list.induct, auto)
lemma crypts_crypt_2:
"Crypt K ` crypts {X} \<subseteq> crypts {Crypt K X}"
by (rule subsetI, simp add: image_iff, erule bexE, drule crypts_comp, simp)
proposition crypts_msg_crypt [simp]:
"crypts_msg (Crypt K X) = Crypt K ` crypts_msg X"
by (simp add: crypts_msg_def, rule equalityI, rule crypts_crypt_1, rule crypts_crypt_2)
lemma crypts_mpair_1:
"crypts {\<lbrace>X, Y\<rbrace>} \<subseteq> insert \<lbrace>X, Y\<rbrace> (crypts {X} \<union> crypts {Y})"
by (rule subsetI, erule crypts.induct, simp_all, (erule disjE, rotate_tac, erule rev_mp,
rule list.induct, (simp+, blast)+)+)
lemma crypts_mpair_2:
"insert \<lbrace>X, Y\<rbrace> (crypts {X} \<union> crypts {Y}) \<subseteq> crypts {\<lbrace>X, Y\<rbrace>}"
by (rule subsetI, simp, erule disjE, blast, erule disjE, (erule crypts.induct, simp,
subst id_apply [symmetric], subst foldr_Nil [symmetric], (rule crypts_fst [of _ X] |
rule crypts_snd), rule crypts_used, simp, blast+)+)
proposition crypts_msg_mpair [simp]:
"crypts_msg \<lbrace>X, Y\<rbrace> = insert \<lbrace>X, Y\<rbrace> (crypts_msg X \<union> crypts_msg Y)"
by (simp add: crypts_msg_def, rule equalityI, rule crypts_mpair_1, rule crypts_mpair_2)
proposition foldr_crypt_size:
"size (foldr Crypt KS X) = size X + length KS"
by (induction KS, simp_all)
proposition key_sets_empty [simp]:
"key_sets X {} = {}"
by (simp add: key_sets_def)
proposition key_sets_mono:
"H \<subseteq> H' \<Longrightarrow> key_sets X H \<subseteq> key_sets X H'"
by (auto simp add: key_sets_def)
proposition key_sets_union:
"key_sets X (H \<union> H') = key_sets X H \<union> key_sets X H'"
by (auto simp add: key_sets_def)
proposition key_sets_insert:
"key_sets X (insert Y H) = key_sets_msg X Y \<union> key_sets X H"
by (simp only: insert_def key_sets_union, subst key_sets_msg_def, simp)
proposition key_sets_msg_eq:
"key_sets_msg X X = {{}}"
by (simp add: key_sets_msg_def key_sets_def, rule equalityI, rule subsetI, simp,
erule exE, erule rev_mp, rule list.induct, simp, rule impI, erule conjE,
drule arg_cong [of _ X size], simp_all add: foldr_crypt_size)
proposition key_sets_msg_num [simp]:
"key_sets_msg X (Num n) = (if X = Num n then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_msg_agent [simp]:
"key_sets_msg X (Agent n) = (if X = Agent n then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_msg_pwd [simp]:
"key_sets_msg X (Pwd n) = (if X = Pwd n then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_msg_key [simp]:
"key_sets_msg X (Key K) = (if X = Key K then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_msg_mult [simp]:
"key_sets_msg X (A \<otimes> B) = (if X = A \<otimes> B then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_msg_hash [simp]:
"key_sets_msg X (Hash Y) = (if X = Hash Y then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
lemma key_sets_crypt_1:
"X \<noteq> Crypt K Y \<Longrightarrow>
key_sets X {Crypt K Y} \<subseteq> insert (InvKey K) ` key_sets X {Y}"
by (rule subsetI, simp add: key_sets_def, erule exE, rotate_tac, erule rev_mp,
rule list.induct, auto)
lemma key_sets_crypt_2:
"insert (InvKey K) ` key_sets X {Y} \<subseteq> key_sets X {Crypt K Y}"
by (rule subsetI, simp add: key_sets_def image_iff, (erule exE, erule conjE)+,
drule arg_cong [where f = "Crypt K"], simp only: comp_apply
[symmetric, of "Crypt K"] foldr_Cons [symmetric], subst conj_commute,
rule exI, rule conjI, assumption, simp)
proposition key_sets_msg_crypt [simp]:
"key_sets_msg X (Crypt K Y) = (if X = Crypt K Y then {{}} else
insert (InvKey K) ` key_sets_msg X Y)"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def, rule impI,
rule equalityI, erule key_sets_crypt_1 [simplified],
rule key_sets_crypt_2 [simplified])
proposition key_sets_msg_mpair [simp]:
"key_sets_msg X \<lbrace>Y, Z\<rbrace> = (if X = \<lbrace>Y, Z\<rbrace> then {{}} else {})"
by (simp add: key_sets_msg_eq, simp add: key_sets_msg_def key_sets_def, rule impI,
rule allI, rule list.induct, simp_all)
proposition key_sets_range:
"U \<in> key_sets X H \<Longrightarrow> U \<subseteq> range Key"
by (simp add: key_sets_def, blast)
proposition key_sets_crypts_hash:
"key_sets (Hash X) (crypts H) \<subseteq> key_sets X (crypts H)"
by (simp add: key_sets_def, blast)
proposition key_sets_crypts_fst:
"key_sets \<lbrace>X, Y\<rbrace> (crypts H) \<subseteq> key_sets X (crypts H)"
by (simp add: key_sets_def, blast)
proposition key_sets_crypts_snd:
"key_sets \<lbrace>X, Y\<rbrace> (crypts H) \<subseteq> key_sets Y (crypts H)"
by (simp add: key_sets_def, blast)
lemma log_spied_1:
"\<lbrakk>s \<turnstile> s';
Log X \<in> parts (used s) \<longrightarrow> Log X \<in> spied s;
Log X \<in> parts (used s')\<rbrakk> \<Longrightarrow>
Log X \<in> spied s'"
by (simp add: rel_def, ((erule disjE)?, ((erule exE)+)?, simp add: parts_insert,
((subst (asm) disj_assoc [symmetric])?, erule disjE, (drule parts_dec |
drule parts_enc | drule parts_sep | drule parts_con), simp+)?)+)
proposition log_spied [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
Log X \<in> parts (used s) \<longrightarrow>
Log X \<in> spied s"
by (erule rtrancl_induct, subst parts_init, simp add: Range_iff image_def, rule impI,
rule log_spied_1)
proposition log_dec:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; s' = insert (Spy, X) s \<and> (Spy, Crypt K X) \<in> s \<and>
(Spy, Key (InvK K)) \<in> s\<rbrakk> \<Longrightarrow>
key_sets Y (crypts {Y. Log Y = X}) \<subseteq> key_sets Y (crypts (Log -` spied s))"
by (rule key_sets_mono, rule crypts_mono, rule subsetI, simp, drule parts_dec
[where Y = X], drule_tac [!] sym, simp_all, rule log_spied [simplified])
proposition log_sep:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; s' = insert (Spy, X) (insert (Spy, Y) s) \<and> (Spy, \<lbrace>X, Y\<rbrace>) \<in> s\<rbrakk> \<Longrightarrow>
key_sets Z (crypts {Z. Log Z = X}) \<subseteq> key_sets Z (crypts (Log -` spied s)) \<and>
key_sets Z (crypts {Z. Log Z = Y}) \<subseteq> key_sets Z (crypts (Log -` spied s))"
by (rule conjI, (rule key_sets_mono, rule crypts_mono, rule subsetI, simp,
frule parts_sep [where Z = X], drule_tac [2] parts_sep [where Z = Y],
simp_all add: parts_msg_def, blast+, drule sym, simp,
rule log_spied [simplified], assumption+)+)
lemma idinfo_spied_1:
"\<lbrakk>s \<turnstile> s';
\<langle>n, X\<rangle> \<in> parts (used s) \<longrightarrow> \<langle>n, X\<rangle> \<in> spied s;
\<langle>n, X\<rangle> \<in> parts (used s')\<rbrakk> \<Longrightarrow>
\<langle>n, X\<rangle> \<in> spied s'"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp add: parts_insert,
((subst (asm) disj_assoc [symmetric])?, erule disjE, (drule parts_dec |
drule parts_enc | drule parts_sep | drule parts_con), simp+)?)+,
auto simp add: parts_insert)
proposition idinfo_spied [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
\<langle>n, X\<rangle> \<in> parts (used s) \<longrightarrow>
\<langle>n, X\<rangle> \<in> spied s"
by (erule rtrancl_induct, subst parts_init, simp add: Range_iff image_def, rule impI,
rule idinfo_spied_1)
proposition idinfo_dec:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; s' = insert (Spy, X) s \<and> (Spy, Crypt K X) \<in> s \<and>
(Spy, Key (InvK K)) \<in> s; \<langle>n, Y\<rangle> = X\<rbrakk> \<Longrightarrow>
\<langle>n, Y\<rangle> \<in> spied s"
by (drule parts_dec [where Y = "\<langle>n, Y\<rangle>"], drule sym, simp, rule idinfo_spied)
proposition idinfo_sep:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; s' = insert (Spy, X) (insert (Spy, Y) s) \<and> (Spy, \<lbrace>X, Y\<rbrace>) \<in> s;
\<langle>n, Z\<rangle> = X \<or> \<langle>n, Z\<rangle> = Y\<rbrakk> \<Longrightarrow>
\<langle>n, Z\<rangle> \<in> spied s"
by (drule parts_sep [where Z = "\<langle>n, Z\<rangle>"], erule disjE, (drule sym, simp)+,
rule idinfo_spied)
lemma idinfo_msg_1:
assumes A: "s\<^sub>0 \<Turnstile> s"
shows "\<lbrakk>s \<turnstile> s'; \<langle>n, X\<rangle> \<in> spied s \<longrightarrow> X \<in> spied s; \<langle>n, X\<rangle> \<in> spied s'\<rbrakk> \<Longrightarrow>
X \<in> spied s'"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp, ((subst (asm) disj_assoc
[symmetric])?, erule disjE, (drule idinfo_dec [OF A] | drule idinfo_sep [OF A]),
simp+ | erule disjE, (simp add: image_iff)+, blast?)?)+)
proposition idinfo_msg [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
\<langle>n, X\<rangle> \<in> spied s \<longrightarrow>
X \<in> spied s"
by (erule rtrancl_induct, simp, blast, rule impI, rule idinfo_msg_1)
proposition parts_agent:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; n \<notin> bad_agent\<rbrakk> \<Longrightarrow> Agent n \<notin> parts (used s)"
apply (erule rtrancl_induct)
subgoal
apply (subst parts_init)
apply (simp add: Range_iff image_def)
done
subgoal
apply (simp add: rel_def)
apply ((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff,
(rule ccontr, (drule parts_dec | drule parts_enc | drule parts_sep |
drule parts_con), simp+)?)+
done
done
lemma idinfo_init_1 [rule_format]:
assumes A: "s\<^sub>0 \<Turnstile> s"
shows "\<lbrakk>s \<turnstile> s'; n \<notin> bad_id_password \<union> bad_id_pubkey \<union> bad_id_shakey;
\<forall>X. \<langle>n, X\<rangle> \<notin> spied s\<rbrakk> \<Longrightarrow>
\<langle>n, X\<rangle> \<notin> spied s'"
by (rule notI, simp add: rel_def, ((erule disjE)?, (erule exE)+, (blast | simp,
((drule idinfo_dec [OF A] | drule idinfo_sep [OF A]), simp, blast |
(erule conjE)+, drule parts_agent [OF A], blast))?)+)
proposition idinfo_init:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; n \<notin> bad_id_password \<union> bad_id_pubkey \<union> bad_id_shakey\<rbrakk> \<Longrightarrow>
\<langle>n, X\<rangle> \<notin> spied s"
by (induction arbitrary: X rule: rtrancl_induct, simp add: image_def, blast,
rule idinfo_init_1)
lemma idinfo_mpair_1 [rule_format]:
"\<lbrakk>(s, s') \<in> rel_id_hash \<union> rel_id_crypt \<union> rel_id_sep \<union> rel_id_con;
\<forall>X Y. \<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s \<longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s)) \<noteq> {};
\<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s'\<rbrakk> \<Longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s')) \<noteq> {}"
by ((erule disjE)?, clarify?, simp add: image_iff Image_def, (drule subsetD
[OF key_sets_crypts_hash] | drule key_sets_range, blast | (drule spec)+,
drule mp, simp, simp add: ex_in_conv [symmetric], erule exE, frule subsetD
[OF key_sets_crypts_fst], drule subsetD [OF key_sets_crypts_snd])?)+
lemma idinfo_mpair_2 [rule_format]:
assumes A: "s\<^sub>0 \<Turnstile> s"
shows "\<lbrakk>s \<turnstile> s'; (s, s') \<notin> rel_id_hash \<union> rel_id_crypt \<union> rel_id_sep \<union> rel_id_con;
\<forall>X Y. \<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s \<longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s)) \<noteq> {};
\<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s'\<rbrakk> \<Longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s')) \<noteq> {}"
by (simp only: rel_def Un_iff de_Morgan_disj, simp, ((erule disjE)?, (erule exE)+,
simp add: Image_def, (simp only: Collect_disj_eq crypts_union key_sets_union, simp)?,
((subst (asm) disj_assoc [symmetric])?, erule disjE, (drule idinfo_dec [OF A] |
drule idinfo_sep [OF A]), simp+)?)+)
proposition idinfo_mpair [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
\<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s \<longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s)) \<noteq> {}"
proof (induction arbitrary: X Y rule: rtrancl_induct, simp add: image_def,
rule impI)
fix s s' X Y
assume
"s\<^sub>0 \<Turnstile> s" and
"s \<turnstile> s'" and
"\<And>X Y. \<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s \<longrightarrow>
key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s)) \<noteq> {}" and
"\<langle>n, \<lbrace>X, Y\<rbrace>\<rangle> \<in> spied s'"
thus "key_sets \<lbrace>X, Y\<rbrace> (crypts (Log -` spied s')) \<noteq> {}"
by (cases "(s, s') \<in> rel_id_hash \<union> rel_id_crypt \<union> rel_id_sep \<union> rel_id_con",
erule_tac [2] idinfo_mpair_2, erule_tac idinfo_mpair_1, simp_all)
qed
proposition key_sets_pwd_empty:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
key_sets (Hash (Pwd n)) (crypts (Log -` spied s)) = {} \<and>
key_sets \<lbrace>Pwd n, X\<rbrace> (crypts (Log -` spied s)) = {} \<and>
key_sets \<lbrace>X, Pwd n\<rbrace> (crypts (Log -` spied s)) = {}"
(is "_ \<Longrightarrow> key_sets ?X (?H s) = _ \<and> key_sets ?Y _ = _ \<and> key_sets ?Z _ = _")
proof (erule rtrancl_induct, simp add: image_iff Image_def)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "key_sets (Hash (Pwd n)) (?H s) = {} \<and>
key_sets \<lbrace>Pwd n, X\<rbrace> (?H s) = {} \<and>
key_sets \<lbrace>X, Pwd n\<rbrace> (?H s) = {}"
show "key_sets (Hash (Pwd n)) (?H s') = {} \<and>
key_sets \<lbrace>Pwd n, X\<rbrace> (?H s') = {} \<and>
key_sets \<lbrace>X, Pwd n\<rbrace> (?H s') = {}"
by (insert B C, simp add: rel_def, ((erule disjE)?, ((erule exE)+)?, simp add:
image_iff Image_def, (simp only: Collect_disj_eq crypts_union
key_sets_union, simp add: crypts_insert key_sets_insert)?,
(frule log_dec [OF A, where Y = "?X"],
frule log_dec [OF A, where Y = "?Y"], drule log_dec [OF A, where Y = "?Z"] |
frule log_sep [OF A, where Z = "?X"], frule log_sep [OF A, where Z = "?Y"],
drule log_sep [OF A, where Z = "?Z"])?)+)
qed
proposition key_sets_pwd_seskey [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
U \<in> key_sets (Pwd n) (crypts (Log -` spied s)) \<longrightarrow>
(\<exists>SK. U = {SesKey SK} \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s))"
(is "_ \<Longrightarrow> _ \<longrightarrow> ?P s")
proof (erule rtrancl_induct, simp add: image_iff Image_def, rule impI)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "U \<in> key_sets (Pwd n) (crypts (Log -` spied s)) \<longrightarrow> ?P s" and
D: "U \<in> key_sets (Pwd n) (crypts (Log -` spied s'))"
show "?P s'"
by (insert B C D, simp add: rel_def, ((erule disjE)?, ((erule exE)+)?, simp
add: image_iff Image_def, (simp only: Collect_disj_eq crypts_union
key_sets_union, simp add: crypts_insert key_sets_insert split: if_split_asm,
blast?)?, (erule disjE, (drule log_dec [OF A] | drule log_sep [OF A]),
(erule conjE)?, drule subsetD, simp)?)+)
qed
lemma pwd_anonymous_1 [rule_format]:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; n \<notin> bad_id_password\<rbrakk> \<Longrightarrow>
\<langle>n, Pwd n\<rangle> \<in> spied s \<longrightarrow>
(\<exists>SK. SesKey SK \<in> spied s \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s))"
(is "\<lbrakk>_; _\<rbrakk> \<Longrightarrow> _ \<longrightarrow> ?P s")
proof (erule rtrancl_induct, simp add: image_def, rule impI)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "\<langle>n, Pwd n\<rangle> \<in> spied s \<longrightarrow> ?P s" and
D: "\<langle>n, Pwd n\<rangle> \<in> spied s'"
show "?P s'"
by (insert B C D, simp add: rel_def, ((erule disjE)?, (erule exE)+, simp add:
image_iff, blast?, ((subst (asm) disj_assoc [symmetric])?, erule disjE,
(drule idinfo_dec [OF A] | drule idinfo_sep [OF A]), simp, blast+ |
insert key_sets_pwd_empty [OF A], clarsimp)?, (((erule disjE)?, erule
conjE, drule sym, simp, (drule key_sets_pwd_seskey [OF A] | drule
idinfo_mpair [OF A, simplified]), simp)+ | drule key_sets_range, blast)?)+)
qed
theorem pwd_anonymous:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_id_password" and
C: "n \<notin> bad_shakey \<inter> (bad_pwd \<union> bad_prikey) \<inter> (bad_id_pubkey \<union> bad_id_shak)"
shows "\<langle>n, Pwd n\<rangle> \<notin> spied s"
proof
assume D: "\<langle>n, Pwd n\<rangle> \<in> spied s"
hence "n \<in> bad_id_password \<union> bad_id_pubkey \<union> bad_id_shakey"
by (rule contrapos_pp, rule_tac idinfo_init [OF A])
moreover have "\<exists>SK. SesKey SK \<in> spied s \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s)"
(is "\<exists>SK. ?P SK \<and> (?Q SK \<or> ?R SK)")
by (rule pwd_anonymous_1 [OF A B D])
then obtain SK where "?P SK" and "?Q SK \<or> ?R SK" by blast
moreover {
assume "?Q SK"
hence "n \<in> bad_shakey \<inter> bad_prikey"
- by (rule_tac contrapos_pp [OF `?P SK`], rule_tac owner_seskey_secret [OF A])
+ by (rule_tac contrapos_pp [OF \<open>?P SK\<close>], rule_tac owner_seskey_secret [OF A])
}
moreover {
assume "?R SK"
hence "n \<in> bad_shakey \<inter> (bad_pwd \<union> bad_prikey)"
- by (rule_tac contrapos_pp [OF `?P SK`], rule_tac asset_seskey_secret [OF A])
+ by (rule_tac contrapos_pp [OF \<open>?P SK\<close>], rule_tac asset_seskey_secret [OF A])
}
ultimately show False
using B and C by blast
qed
proposition idinfo_pwd_start:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_agent"
shows "\<lbrakk>s \<turnstile> s'; \<exists>X. \<langle>n, X\<rangle> \<in> spied s' \<and> X \<noteq> Pwd n;
\<not> (\<exists>X. \<langle>n, X\<rangle> \<in> spied s \<and> X \<noteq> Pwd n)\<rbrakk> \<Longrightarrow>
\<exists>SK. SesKey SK \<in> spied s \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s)"
proof (simp add: rel_def, insert parts_agent [OF A B], insert key_sets_pwd_empty
[OF A], (erule disjE, (erule exE)+, simp, erule conjE, (subst (asm) disj_assoc
[symmetric])?, (erule disjE)?, (drule idinfo_dec [OF A] | drule idinfo_sep
[OF A] | drule spec, drule mp), simp+)+, auto, rule FalseE, rule_tac [3] FalseE)
fix X U K
assume "\<forall>X. (Spy, \<langle>n, X\<rangle>) \<in> s \<longrightarrow> X = Pwd n" and "(Spy, \<langle>n, K\<rangle>) \<in> s"
hence "K = Pwd n" by simp
moreover assume "U \<in> key_sets X (crypts (Log -` spied s))"
hence "U \<subseteq> range Key"
by (rule key_sets_range)
moreover assume "K \<in> U"
ultimately show False by blast
next
fix X U
assume "\<forall>X. (Spy, \<langle>n, X\<rangle>) \<in> s \<longrightarrow> X = Pwd n" and "(Spy, \<langle>n, X\<rangle>) \<in> s"
hence C: "X = Pwd n" by simp
moreover assume "U \<in> key_sets X (crypts (Log -` spied s))"
ultimately have "U \<in> key_sets (Pwd n) (crypts (Log -` spied s))" by simp
hence "\<exists>SK. U = {SesKey SK} \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s)"
by (rule key_sets_pwd_seskey [OF A])
moreover assume "U \<subseteq> spied s"
ultimately show "\<exists>x U V. (Spy, Key (SesK (x, U, V))) \<in> s \<and>
((Owner n, Crypt (SesK (x, U, V)) X) \<in> s \<or>
(Asset n, Crypt (SesK (x, U, V)) (Num 0)) \<in> s)"
using C by auto
next
fix X U K
assume "\<forall>X. (Spy, \<langle>n, X\<rangle>) \<in> s \<longrightarrow> X = Pwd n" and "(Spy, \<langle>n, K\<rangle>) \<in> s"
hence "K = Pwd n" by simp
moreover assume "U \<in> key_sets X (crypts (Log -` spied s))"
hence "U \<subseteq> range Key"
by (rule key_sets_range)
moreover assume "K \<in> U"
ultimately show False by blast
qed
proposition idinfo_pwd:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; \<exists>X. \<langle>n, X\<rangle> \<in> spied s \<and> X \<noteq> Pwd n;
n \<notin> bad_id_pubkey \<union> bad_id_shakey\<rbrakk> \<Longrightarrow>
\<exists>SK. SesKey SK \<in> spied s \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s)"
by (drule rtrancl_start, assumption, simp, blast, (erule exE)+, (erule conjE)+,
frule idinfo_pwd_start [of _ n], simp+, drule r_into_rtrancl, drule rtrancl_trans,
assumption, (drule state_subset)+, blast)
theorem auth_prikey_anonymous:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_id_prikey" and
C: "n \<notin> bad_shakey \<inter> bad_prikey \<inter> (bad_id_password \<union> bad_id_shak)"
shows "\<langle>n, Auth_PriKey n\<rangle> \<notin> spied s"
proof
assume D: "\<langle>n, Auth_PriKey n\<rangle> \<in> spied s"
hence "n \<in> bad_id_password \<union> bad_id_pubkey \<union> bad_id_shakey"
by (rule contrapos_pp, rule_tac idinfo_init [OF A])
moreover have "Auth_PriKey n \<in> spied s"
by (rule idinfo_msg [OF A D])
hence "n \<in> bad_prikey"
by (rule contrapos_pp, rule_tac auth_prikey_secret [OF A])
moreover from this have E: "n \<notin> bad_id_pubkey"
using B by simp
moreover have "n \<in> bad_shakey"
proof (cases "n \<in> bad_id_shakey", simp)
case False
with D and E have "\<exists>SK. SesKey SK \<in> spied s \<and>
((Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<or>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s)"
(is "\<exists>SK. ?P SK \<and> (?Q SK \<or> ?R SK)")
by (rule_tac idinfo_pwd [OF A], rule_tac exI [of _ "Auth_PriKey n"], simp_all)
then obtain SK where "?P SK" and "?Q SK \<or> ?R SK" by blast
moreover {
assume "?Q SK"
hence "n \<in> bad_shakey \<inter> bad_prikey"
- by (rule_tac contrapos_pp [OF `?P SK`], rule_tac owner_seskey_secret
+ by (rule_tac contrapos_pp [OF \<open>?P SK\<close>], rule_tac owner_seskey_secret
[OF A])
}
moreover {
assume "?R SK"
hence "n \<in> bad_shakey \<inter> (bad_pwd \<union> bad_prikey)"
- by (rule_tac contrapos_pp [OF `?P SK`], rule_tac asset_seskey_secret
+ by (rule_tac contrapos_pp [OF \<open>?P SK\<close>], rule_tac asset_seskey_secret
[OF A])
}
ultimately show ?thesis by blast
qed
ultimately show False
using C by blast
qed
theorem auth_shakey_anonymous:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_id_shakey" and
C: "n \<notin> bad_shakey \<inter> (bad_id_password \<union> bad_id_pubkey)"
shows "\<langle>n, Key (Auth_ShaKey n)\<rangle> \<notin> spied s"
proof
assume D: "\<langle>n, Key (Auth_ShaKey n)\<rangle> \<in> spied s"
hence "n \<in> bad_id_password \<union> bad_id_pubkey \<union> bad_id_shakey"
by (rule contrapos_pp, rule_tac idinfo_init [OF A])
moreover have "Key (Auth_ShaKey n) \<in> spied s"
by (rule idinfo_msg [OF A D])
hence "n \<in> bad_shakey"
by (rule contrapos_pp, rule_tac auth_shakey_secret [OF A])
ultimately show False
using B and C by blast
qed
end
\ No newline at end of file
diff --git a/thys/Relational_Method/Authentication.thy b/thys/Relational_Method/Authentication.thy
--- a/thys/Relational_Method/Authentication.thy
+++ b/thys/Relational_Method/Authentication.thy
@@ -1,1903 +1,1903 @@
(* Title: The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Confidentiality and authenticity properties"
theory Authentication
imports Definitions
begin
text \<open>
\label{Authentication}
\<close>
proposition rtrancl_start [rule_format]:
"(x, y) \<in> r\<^sup>* \<Longrightarrow> P y \<longrightarrow> \<not> P x \<longrightarrow>
(\<exists>u v. (x, u) \<in> r\<^sup>* \<and> (u, v) \<in> r \<and> (v, y) \<in> r\<^sup>* \<and> \<not> P u \<and> P v)"
(is "_ \<Longrightarrow> _ \<longrightarrow> _ \<longrightarrow> (\<exists>u v. ?Q x y u v)")
proof (erule rtrancl_induct, simp, (rule impI)+)
fix y z
assume
A: "(x, y) \<in> r\<^sup>*" and
B: "(y, z) \<in> r" and
C: "P z"
assume "P y \<longrightarrow> \<not> P x \<longrightarrow>(\<exists>u v. ?Q x y u v)" and "\<not> P x"
hence D: "P y \<longrightarrow> (\<exists>u v. ?Q x y u v)" by simp
show "\<exists>u v. ?Q x z u v"
proof (cases "P y")
case True
with D obtain u v where "?Q x y u v" by blast
moreover from this and B have "(v, z) \<in> r\<^sup>*" by auto
ultimately show ?thesis by blast
next
case False
with A and B and C have "?Q x z y z" by simp
thus ?thesis by blast
qed
qed
proposition state_subset:
"s \<Turnstile> s' \<Longrightarrow> s \<subseteq> s'"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition spied_subset:
"s \<Turnstile> s' \<Longrightarrow> spied s \<subseteq> spied s'"
by (rule Image_mono, erule state_subset, simp)
proposition used_subset:
"s \<Turnstile> s' \<Longrightarrow> used s \<subseteq> used s'"
by (rule Range_mono, rule state_subset)
proposition asset_ii_init:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; (Asset n, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s\<rbrakk> \<Longrightarrow>
PriKey A \<notin> spied s\<^sub>0"
by (drule rtrancl_start, assumption, simp add: image_def, (erule exE)+,
erule conjE, rule notI, drule spied_subset, drule subsetD, assumption,
auto simp add: rel_def)
proposition auth_prikey_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow> Auth_PriKey n \<in> used s"
by (drule used_subset, erule subsetD, simp add: Range_iff image_def, blast)
proposition asset_i_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, Crypt (Auth_ShaKey n) (PriKey A)) \<in> s \<longrightarrow>
PriKey A \<in> used s"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition owner_ii_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
PriKey A \<in> used s"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition asset_ii_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<longrightarrow>
PriKey A \<in> used s"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition owner_iii_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, \<lbrace>Num 3, PubKey A\<rbrace>) \<in> s \<longrightarrow>
PriKey A \<in> used s"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition asset_iii_used:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, \<lbrace>Num 4, PubKey A\<rbrace>) \<in> s \<longrightarrow>
PriKey A \<in> used s"
by (erule rtrancl_induct, auto simp add: rel_def image_def)
proposition asset_i_unique [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, Crypt (Auth_ShaKey m) (PriKey A)) \<in> s \<longrightarrow>
(Asset n, Crypt (Auth_ShaKey n) (PriKey A)) \<in> s \<longrightarrow>
m = n"
by (erule rtrancl_induct, simp add: image_def, frule asset_i_used [of _ m A],
drule asset_i_used [of _ n A], auto simp add: rel_def)
proposition owner_ii_unique [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner m, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
m = n"
by (erule rtrancl_induct, simp add: image_def, frule owner_ii_used [of _ m A],
drule owner_ii_used [of _ n A], auto simp add: rel_def)
proposition asset_ii_unique [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<longrightarrow>
(Asset n, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<longrightarrow>
m = n"
by (erule rtrancl_induct, simp add: image_def, frule asset_ii_used [of _ m A],
drule asset_ii_used [of _ n A], auto simp add: rel_def)
proposition auth_prikey_asset_i [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, Crypt (Auth_ShaKey m) (Auth_PriKey n)) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, drule auth_prikey_used [of _ n],
auto simp add: rel_def)
proposition auth_pubkey_owner_ii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner m, \<lbrace>Num 1, Auth_PubKey n\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, drule auth_prikey_used [of _ n],
auto simp add: rel_def)
proposition auth_pubkey_asset_ii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, \<lbrace>Num 2, Auth_PubKey n\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, drule auth_prikey_used [of _ n],
auto simp add: rel_def)
proposition asset_i_owner_ii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, Crypt (Auth_ShaKey m) (PriKey A)) \<in> s \<longrightarrow>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, frule asset_i_used [of _ m A],
drule owner_ii_used [of _ n A], auto simp add: rel_def)
proposition asset_i_asset_ii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, Crypt (Auth_ShaKey m) (PriKey A)) \<in> s \<longrightarrow>
(Asset n, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, frule asset_i_used [of _ m A],
drule asset_ii_used [of _ n A], auto simp add: rel_def)
proposition asset_ii_owner_ii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<longrightarrow>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, frule asset_ii_used [of _ m A],
drule owner_ii_used [of _ n A], auto simp add: rel_def)
proposition asset_iii_owner_iii [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset m, \<lbrace>Num 4, PubKey A\<rbrace>) \<in> s \<longrightarrow>
(Owner n, \<lbrace>Num 3, PubKey A\<rbrace>) \<in> s \<longrightarrow>
False"
by (erule rtrancl_induct, simp add: image_def, frule asset_iii_used [of _ m A],
drule owner_iii_used [of _ n A], auto simp add: rel_def)
proposition asset_iv_state [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, Token n (Auth_PriK n) B C SK) \<in> s \<longrightarrow>
(\<exists>A D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s \<and> (Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
Crypt (SesK SK) (PubKey D) \<in> used s \<and> (Asset n, PubKey B) \<in> s)"
by (erule rtrancl_induct, auto simp add: rel_def)
proposition owner_v_state [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s \<longrightarrow>
(Owner n, SesKey SK) \<in> s \<and>
(\<exists>A B C. Token n A B C SK \<in> used s \<and> B \<in> fst (snd SK))"
by (erule rtrancl_induct, auto simp add: rel_def, blast)
proposition asset_v_state [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, Crypt (SesK SK) (Num 0)) \<in> s \<longrightarrow>
(Asset n, SesKey SK) \<in> s \<and> Crypt (SesK SK) (Pwd n) \<in> used s"
by (erule rtrancl_induct, simp_all add: rel_def image_def,
((erule disjE)?, (erule exE)+, simp add: Range_Un_eq)+)
lemma owner_seskey_nonce_1:
"\<lbrakk>s \<turnstile> s';
(Owner n, SesKey SK) \<in> s \<longrightarrow>
(\<exists>S. fst SK = Some S \<and> Crypt (Auth_ShaKey n) (PriKey S) \<in> used s) \<or>
fst SK = None;
(Owner n, SesKey SK) \<in> s'\<rbrakk> \<Longrightarrow>
(\<exists>S. fst SK = Some S \<and> Crypt (Auth_ShaKey n) (PriKey S) \<in> used s') \<or>
fst SK = None"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp+)+,
split if_split_asm, auto)
proposition owner_seskey_nonce [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, SesKey SK) \<in> s \<longrightarrow>
(\<exists>S. fst SK = Some S \<and> Crypt (Auth_ShaKey n) (PriKey S) \<in> used s) \<or>
fst SK = None"
by (erule rtrancl_induct, simp add: image_def, rule impI, rule owner_seskey_nonce_1)
proposition owner_seskey_other [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, SesKey SK) \<in> s \<longrightarrow>
(\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner n, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner n, Crypt (SesK SK) (PubKey D)) \<in> s)"
by (erule rtrancl_induct, auto simp add: rel_def, blast+)
proposition asset_seskey_nonce [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, SesKey SK) \<in> s \<longrightarrow>
(\<exists>S. fst SK = Some S \<and> (Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s)"
by (erule rtrancl_induct, auto simp add: rel_def)
proposition asset_seskey_other [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Asset n, SesKey SK) \<in> s \<longrightarrow>
(\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s \<and> (Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
(Asset n, Token n (Auth_PriK n) B C SK) \<in> s)"
by (erule rtrancl_induct, auto simp add: rel_def, blast)
declare Range_Un_eq [simp]
proposition used_prod [simp]:
"A \<noteq> {} \<Longrightarrow> used (A \<times> H) = H"
by (simp add: Range_snd)
proposition parts_idem [simp]:
"parts (parts H) = parts H"
by (rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_mono:
"H \<subseteq> H' \<Longrightarrow> parts H \<subseteq> parts H'"
by (rule subsetI, erule parts.induct, auto)
proposition parts_msg_mono:
"X \<in> H \<Longrightarrow> parts_msg X \<subseteq> parts H"
by (subgoal_tac "{X} \<subseteq> H", subst parts_msg_def, erule parts_mono, simp)
lemma parts_union_1:
"parts (H \<union> H') \<subseteq> parts H \<union> parts H'"
by (rule subsetI, erule parts.induct, auto)
lemma parts_union_2:
"parts H \<union> parts H' \<subseteq> parts (H \<union> H')"
by (rule subsetI, erule UnE, erule_tac [!] parts.induct, auto)
proposition parts_union [simp]:
"parts (H \<union> H') = parts H \<union> parts H'"
by (rule equalityI, rule parts_union_1, rule parts_union_2)
proposition parts_insert:
"parts (insert X H) = parts_msg X \<union> parts H"
by (simp only: insert_def parts_union, subst parts_msg_def, simp)
proposition parts_msg_num [simp]:
"parts_msg (Num n) = {Num n}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_msg_pwd [simp]:
"parts_msg (Pwd n) = {Pwd n}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_msg_key [simp]:
"parts_msg (Key K) = {Key K}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_msg_mult [simp]:
"parts_msg (A \<otimes> B) = {A \<otimes> B}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_msg_hash [simp]:
"parts_msg (Hash X) = {Hash X}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
lemma parts_crypt_1:
"parts {Crypt K X} \<subseteq> insert (Crypt K X) (parts {X})"
by (rule subsetI, erule parts.induct, auto)
lemma parts_crypt_2:
"insert (Crypt K X) (parts {X}) \<subseteq> parts {Crypt K X}"
by (rule subsetI, simp, erule disjE, blast, erule parts.induct, auto)
proposition parts_msg_crypt [simp]:
"parts_msg (Crypt K X) = insert (Crypt K X) (parts_msg X)"
by (simp add: parts_msg_def, rule equalityI, rule parts_crypt_1, rule parts_crypt_2)
lemma parts_mpair_1:
"parts {\<lbrace>X, Y\<rbrace>} \<subseteq> insert \<lbrace>X, Y\<rbrace> (parts {X} \<union> parts {Y})"
by (rule subsetI, erule parts.induct, auto)
lemma parts_mpair_2:
"insert \<lbrace>X, Y\<rbrace> (parts {X} \<union> parts {Y}) \<subseteq> parts {\<lbrace>X, Y\<rbrace>}"
by (rule subsetI, simp, erule disjE, blast, erule disjE, erule_tac [!] parts.induct,
auto)
proposition parts_msg_mpair [simp]:
"parts_msg \<lbrace>X, Y\<rbrace> = insert \<lbrace>X, Y\<rbrace> (parts_msg X \<union> parts_msg Y)"
by (simp add: parts_msg_def, rule equalityI, rule parts_mpair_1, rule parts_mpair_2)
proposition parts_msg_idinfo [simp]:
"parts_msg \<langle>n, X\<rangle> = {\<langle>n, X\<rangle>}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_msg_trace [simp]:
"parts_msg (Log X) = {Log X}"
by (subst parts_msg_def, rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_idinfo [simp]:
"parts (IDInfo n ` H) = IDInfo n ` H"
by (rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_trace [simp]:
"parts (Log ` H) = Log ` H"
by (rule equalityI, rule subsetI, erule parts.induct, auto)
proposition parts_dec:
"\<lbrakk>s' = insert (Spy, X) s \<and> (Spy, Crypt K X) \<in> s \<and> (Spy, Key (InvK K)) \<in> s;
Y \<in> parts_msg X\<rbrakk> \<Longrightarrow>
Y \<in> parts (used s)"
by (subgoal_tac "X \<in> parts (used s)", drule parts_msg_mono [of X], auto)
proposition parts_enc:
"\<lbrakk>s' = insert (Spy, Crypt K X) s \<and> (Spy, X) \<in> s \<and> (Spy, Key K) \<in> s;
Y \<in> parts_msg X\<rbrakk> \<Longrightarrow>
Y \<in> parts (used s)"
by (subgoal_tac "X \<in> parts (used s)", drule parts_msg_mono [of X], auto)
proposition parts_sep:
"\<lbrakk>s' = insert (Spy, X) (insert (Spy, Y) s) \<and> (Spy, \<lbrace>X, Y\<rbrace>) \<in> s;
Z \<in> parts_msg X \<or> Z \<in> parts_msg Y\<rbrakk> \<Longrightarrow>
Z \<in> parts (used s)"
by (erule disjE, subgoal_tac "X \<in> parts (used s)", drule parts_msg_mono [of X],
subgoal_tac [3] "Y \<in> parts (used s)", drule_tac [3] parts_msg_mono [of Y], auto)
proposition parts_con:
"\<lbrakk>s' = insert (Spy, \<lbrace>X, Y\<rbrace>) s \<and> (Spy, X) \<in> s \<and> (Spy, Y) \<in> s;
Z \<in> parts_msg X \<or> Z \<in> parts_msg Y\<rbrakk> \<Longrightarrow>
Z \<in> parts (used s)"
by (erule disjE, subgoal_tac "X \<in> parts (used s)", drule parts_msg_mono [of X],
subgoal_tac [3] "Y \<in> parts (used s)", drule_tac [3] parts_msg_mono [of Y], auto)
lemma parts_init_1:
"parts (used s\<^sub>0) \<subseteq> used s\<^sub>0 \<union> range (Hash \<circ> Agent) \<union>
range (Hash \<circ> Auth_PubKey) \<union>
range (\<lambda>n. \<lbrace>Hash (Agent n), Hash (Auth_PubKey n)\<rbrace>)"
by (rule subsetI, erule parts.induct, (clarify | simp add: Range_iff image_def)+)
lemma parts_init_2:
"used s\<^sub>0 \<union> range (Hash \<circ> Agent) \<union> range (Hash \<circ> Auth_PubKey) \<union>
range (\<lambda>n. \<lbrace>Hash (Agent n), Hash (Auth_PubKey n)\<rbrace>) \<subseteq> parts (used s\<^sub>0)"
by (rule subsetI, auto simp add: parts_insert)
proposition parts_init:
"parts (used s\<^sub>0) = used s\<^sub>0 \<union> range (Hash \<circ> Agent) \<union>
range (Hash \<circ> Auth_PubKey) \<union>
range (\<lambda>n. \<lbrace>Hash (Agent n), Hash (Auth_PubKey n)\<rbrace>)"
by (rule equalityI, rule parts_init_1, rule parts_init_2)
proposition parts_crypt_prikey_start:
"\<lbrakk>s \<turnstile> s'; Crypt K (PriKey A) \<in> parts (used s');
Crypt K (PriKey A) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n. K = Auth_ShaKey n \<and>
(Asset n, Crypt (Auth_ShaKey n) (PriKey A)) \<in> s') \<or>
{PriKey A, Key K} \<subseteq> spied s'"
by (simp add: rel_def, erule disjE, (erule exE)+, simp add: parts_insert, blast,
(((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | erule disjE, simp, drule parts_enc |
drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_crypt_prikey:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt K (PriKey A) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n. K = Auth_ShaKey n \<and>
(Asset n, Crypt (Auth_ShaKey n) (PriKey A)) \<in> s) \<or>
{PriKey A, Key K} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_prikey_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_pubkey_start:
"\<lbrakk>s \<turnstile> s'; Crypt (SesK SK) (PubKey C) \<in> parts (used s');
Crypt (SesK SK) (PubKey C) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
C \<in> snd (snd SK) \<and> ((\<exists>n. (Owner n, SesKey SK) \<in> s') \<or>
(\<exists>n B. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s')) \<or>
SesKey SK \<in> spied s'"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp add: parts_insert image_iff)+,
blast, erule disjE, (erule exE)+, simp add: parts_insert image_iff, blast,
(((erule disjE)?, ((erule exE)+)?, simp add: parts_insert image_iff)+,
((drule parts_dec | drule parts_enc | drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_crypt_pubkey:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt (SesK SK) (PubKey C) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
C \<in> snd (snd SK) \<and> ((\<exists>n. (Owner n, SesKey SK) \<in> s) \<or>
(\<exists>n B. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s)) \<or>
SesKey SK \<in> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_pubkey_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_key_start:
"\<lbrakk>s \<turnstile> s'; Crypt K (Key K') \<in> parts (used s');
Crypt K (Key K') \<notin> parts (used s); K' \<notin> range PriK \<union> range PubK\<rbrakk> \<Longrightarrow>
{Key K', Key K} \<subseteq> spied s'"
by (simp add: rel_def, (((erule disjE)?, ((erule exE)+)?, simp add: parts_insert
image_iff)+, ((drule parts_dec | drule parts_enc | drule parts_sep | drule parts_con),
simp+)?)+)
proposition parts_crypt_key:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt K (Key K') \<in> parts (used s);
K' \<notin> range PriK \<union> range PubK\<rbrakk> \<Longrightarrow>
{Key K', Key K} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_key_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_sign_start:
"\<lbrakk>s \<turnstile> s'; Crypt (SesK SK) (Sign n A) \<in> parts (used s');
Crypt (SesK SK) (Sign n A) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
(Asset n, SesKey SK) \<in> s' \<or> SesKey SK \<in> spied s'"
by (simp add: rel_def, (((erule disjE)?, ((erule exE)+)?, simp add: parts_insert
image_iff)+, ((drule parts_dec | drule parts_enc | drule parts_sep | drule parts_con),
simp+)?)+)
proposition parts_crypt_sign:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt (SesK SK) (Sign n A) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
(Asset n, SesKey SK) \<in> s \<or> SesKey SK \<in> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_sign_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_pwd_start:
"\<lbrakk>s \<turnstile> s'; Crypt K (Pwd n) \<in> parts (used s');
Crypt K (Pwd n) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>SK. K = SesK SK \<and> (Owner n, Crypt (SesK SK) (Pwd n)) \<in> s') \<or>
{Pwd n, Key K} \<subseteq> spied s'"
by (simp add: rel_def, (((erule disjE)?, ((erule exE)+)?, simp add: parts_insert
image_iff)+, ((drule parts_dec | drule parts_enc | drule parts_sep | drule parts_con),
simp+)?)+)
proposition parts_crypt_pwd:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt K (Pwd n) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>SK. K = SesK SK \<and> (Owner n, Crypt (SesK SK) (Pwd n)) \<in> s) \<or>
{Pwd n, Key K} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_pwd_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_num_start:
"\<lbrakk>s \<turnstile> s'; Crypt (SesK SK) (Num 0) \<in> parts (used s');
Crypt (SesK SK) (Num 0) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n. (Asset n, Crypt (SesK SK) (Num 0)) \<in> s') \<or> SesKey SK \<in> spied s'"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp add: parts_insert image_iff)+,
blast, (((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | erule disjE, simp, drule parts_enc |
drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_crypt_num:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt (SesK SK) (Num 0) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n. (Asset n, Crypt (SesK SK) (Num 0)) \<in> s) \<or> SesKey SK \<in> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_num_start, assumption+,
(drule state_subset)+, blast)
proposition parts_crypt_mult_start:
"\<lbrakk>s \<turnstile> s'; Crypt (SesK SK) (A \<otimes> B) \<in> parts (used s');
Crypt (SesK SK) (A \<otimes> B) \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
B \<in> fst (snd SK) \<and> (\<exists>n C. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s') \<or>
{A \<otimes> B, SesKey SK} \<subseteq> spied s"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp add: parts_insert image_iff)+,
blast, (((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | erule disjE, simp, drule parts_enc |
drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_crypt_mult:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt (SesK SK) (A \<otimes> B) \<in> parts (used s)\<rbrakk> \<Longrightarrow>
B \<in> fst (snd SK) \<and> (\<exists>n C. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s) \<or>
{A \<otimes> B, SesKey SK} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_crypt_mult_start, assumption+,
drule converse_rtrancl_into_rtrancl, assumption, drule state_subset [of _ s],
drule spied_subset [of _ s], blast)
proposition parts_mult_start:
"\<lbrakk>s \<turnstile> s'; A \<otimes> B \<in> parts (used s'); A \<otimes> B \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n SK. A = Auth_PriK n \<and> (Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s' \<and>
Crypt (SesK SK) (A \<otimes> B) \<in> parts (used s')) \<or>
{PriKey A, PriKey B} \<subseteq> spied s'"
by (simp add: rel_def, (erule disjE, (erule exE)+, simp add: parts_insert image_iff)+,
blast, (((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | drule parts_enc | drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_mult:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; A \<otimes> B \<in> parts (used s)\<rbrakk> \<Longrightarrow>
(\<exists>n. A = Auth_PriK n \<and> (Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s) \<or>
{PriKey A, PriKey B} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, frule parts_mult_start, assumption+,
(drule state_subset)+, blast)
proposition parts_mpair_key_start:
"\<lbrakk>s \<turnstile> s'; \<lbrace>X, Y\<rbrace> \<in> parts (used s'); \<lbrace>X, Y\<rbrace> \<notin> parts (used s);
X = Key K \<or> Y = Key K \<and> K \<notin> range PubK\<rbrakk> \<Longrightarrow>
{X, Y} \<subseteq> spied s'"
by (erule disjE, simp_all add: rel_def,
(((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | drule parts_enc |
drule parts_sep | erule disjE, simp, drule parts_con), simp+)?)+)
proposition parts_mpair_key:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; \<lbrace>X, Y\<rbrace> \<in> parts (used s);
X = Key K \<or> Y = Key K \<and> K \<notin> range PubK\<rbrakk> \<Longrightarrow>
{X, Y} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
blast, (erule exE)+, (erule conjE)+, frule parts_mpair_key_start, assumption+,
(drule state_subset)+, blast)
proposition parts_mpair_pwd_start:
"\<lbrakk>s \<turnstile> s'; \<lbrace>X, Y\<rbrace> \<in> parts (used s'); \<lbrace>X, Y\<rbrace> \<notin> parts (used s);
X = Pwd n \<or> Y = Pwd n\<rbrakk> \<Longrightarrow>
{X, Y} \<subseteq> spied s'"
by (erule disjE, simp_all add: rel_def,
(((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | drule parts_enc |
drule parts_sep | erule disjE, simp, drule parts_con), simp+)?)+)
proposition parts_mpair_pwd:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; \<lbrace>X, Y\<rbrace> \<in> parts (used s); X = Pwd n \<or> Y = Pwd n\<rbrakk> \<Longrightarrow>
{X, Y} \<subseteq> spied s"
by (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
blast, (erule exE)+, (erule conjE)+, frule parts_mpair_pwd_start, assumption+,
(drule state_subset)+, blast)
proposition parts_pubkey_false_start:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "Crypt (SesK SK) (PubKey C) \<in> parts (used s')" and
D: "Crypt (SesK SK) (PubKey C) \<notin> parts (used s)" and
E: "\<forall>n. (Owner n, SesKey SK) \<notin> s'" and
F: "SesKey SK \<notin> spied s'"
shows False
proof -
have "C \<in> snd (snd SK) \<and> ((\<exists>n. (Owner n, SesKey SK) \<in> s') \<or>
(\<exists>n B. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s')) \<or>
SesKey SK \<in> spied s'"
(is "?P C \<and> ((\<exists>n. ?Q n s') \<or> (\<exists>n B. ?R n B C s')) \<or> ?S s'")
by (rule parts_crypt_pubkey_start [OF B C D])
then obtain n B where "?P C" and "?R n B C s'"
using E and F by blast
moreover have "\<not> ?R n B C s"
using D by blast
ultimately have "\<exists>D. Crypt (SesK SK) (PubKey D) \<in> used s"
(is "\<exists>D. ?U D")
using B by (auto simp add: rel_def)
then obtain D where "?U D" ..
hence "?P D \<and> ((\<exists>n. ?Q n s) \<or> (\<exists>n B. ?R n B D s)) \<or> ?S s"
by (rule_tac parts_crypt_pubkey [OF A], blast)
moreover have G: "s \<subseteq> s'"
by (rule state_subset, insert B, simp)
have "\<forall>n. (Owner n, SesKey SK) \<notin> s"
by (rule allI, rule notI, drule subsetD [OF G], insert E, simp)
moreover have H: "spied s \<subseteq> spied s'"
by (rule Image_mono [OF G], simp)
have "SesKey SK \<notin> spied s"
by (rule notI, drule subsetD [OF H], insert F, contradiction)
ultimately obtain n' B' where "?R n' B' D s" by blast
have "\<exists>A' D'. fst (snd SK) = {A', B'} \<and> snd (snd SK) = {D, D'} \<and>
(Asset n', \<lbrace>Num 2, PubKey B'\<rbrace>) \<in> s \<and>
(Asset n', \<lbrace>Num 4, PubKey D'\<rbrace>) \<in> s \<and>
?U D' \<and> (Asset n', PubKey B') \<in> s"
- by (rule asset_iv_state [OF A `?R n' B' D s`])
+ by (rule asset_iv_state [OF A \<open>?R n' B' D s\<close>])
then obtain D' where "snd (snd SK) = {D, D'}" and "?U D'" by blast
hence "Crypt (SesK SK) (PubKey C) \<in> parts (used s)"
- using `?P C` and `?U D` by auto
+ using \<open>?P C\<close> and \<open>?U D\<close> by auto
thus False
using D by contradiction
qed
proposition parts_pubkey_false:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Crypt (SesK SK) (PubKey C) \<in> parts (used s);
\<forall>n. (Owner n, SesKey SK) \<notin> s; SesKey SK \<notin> spied s\<rbrakk> \<Longrightarrow>
False"
proof (drule rtrancl_start, assumption, subst parts_init, simp add: Range_iff image_def,
(erule exE)+, (erule conjE)+, erule parts_pubkey_false_start, assumption+,
rule allI, rule_tac [!] notI)
fix v n
assume "(Owner n, SesKey SK) \<in> v" and "v \<Turnstile> s"
hence "(Owner n, SesKey SK) \<in> s"
by (erule_tac rev_subsetD, rule_tac state_subset)
moreover assume "\<forall>n. (Owner n, SesKey SK) \<notin> s"
ultimately show False by simp
next
fix v
assume "SesKey SK \<in> spied v" and "v \<Turnstile> s"
hence "SesKey SK \<in> spied s"
by (erule_tac rev_subsetD, rule_tac spied_subset)
moreover assume "SesKey SK \<notin> spied s"
ultimately show False by contradiction
qed
proposition asset_ii_spied_start:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "PriKey B \<in> spied s'" and
D: "PriKey B \<notin> spied s" and
E: "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s"
shows "Auth_PriKey n \<in> spied s \<and>
(\<exists>C SK. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s)"
(is "_ \<and> (\<exists>C SK. ?P n C SK s)")
proof -
have "\<exists>A. (A \<otimes> B \<in> spied s \<or> B \<otimes> A \<in> spied s) \<and> PriKey A \<in> spied s"
proof (insert B C D, auto simp add: rel_def, rule_tac [!] FalseE)
assume "Key (PriK B) \<notin> used s"
moreover have "PriKey B \<in> used s"
by (rule asset_ii_used [OF A, THEN mp, OF E])
ultimately show False by simp
next
fix K
assume "(Spy, Crypt K (Key (PriK B))) \<in> s"
hence "Crypt K (PriKey B) \<in> parts (used s)" by auto
hence "(\<exists>m. K = Auth_ShaKey m \<and>
(Asset m, Crypt (Auth_ShaKey m) (PriKey B)) \<in> s) \<or>
{PriKey B, Key K} \<subseteq> spied s"
(is "(\<exists>m. _ \<and> ?P m) \<or> _")
by (rule parts_crypt_prikey [OF A])
then obtain m where "?P m"
using D by blast
thus False
by (rule asset_i_asset_ii [OF A _ E])
next
fix Y
assume "(Spy, \<lbrace>Key (PriK B), Y\<rbrace>) \<in> s"
hence "\<lbrace>PriKey B, Y\<rbrace> \<in> parts (used s)" by auto
hence "{PriKey B, Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "PriK B"], simp)
thus False
using D by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (PriK B)\<rbrace>) \<in> s"
hence "\<lbrace>X, PriKey B\<rbrace> \<in> parts (used s)" by auto
hence "{X, PriKey B} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "PriK B"], simp add: image_def)
thus False
using D by simp
qed
then obtain A where F: "PriKey A \<in> spied s" and
"A \<otimes> B \<in> spied s \<or> B \<otimes> A \<in> spied s"
by blast
hence "A \<otimes> B \<in> parts (used s) \<or> B \<otimes> A \<in> parts (used s)" by blast
moreover have "B \<otimes> A \<notin> parts (used s)"
proof
assume "B \<otimes> A \<in> parts (used s)"
hence "(\<exists>m. B = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s) \<or>
{PriKey B, PriKey A} \<subseteq> spied s"
by (rule parts_mult [OF A])
then obtain m where "B = Auth_PriK m"
using D by blast
hence "(Asset n, \<lbrace>Num 2, Auth_PubKey m\<rbrace>) \<in> s"
using E by simp
thus False
by (rule auth_pubkey_asset_ii [OF A])
qed
ultimately have "A \<otimes> B \<in> parts (used s)" by simp
with A have "\<exists>u v. s\<^sub>0 \<Turnstile> u \<and> u \<turnstile> v \<and> v \<Turnstile> s \<and>
A \<otimes> B \<notin> parts (used u) \<and> A \<otimes> B \<in> parts (used v)"
by (rule rtrancl_start, subst parts_init, simp add: Range_iff image_def)
then obtain u v where G: "u \<turnstile> v" and H: "v \<Turnstile> s" and
I: "A \<otimes> B \<notin> parts (used u)" and "A \<otimes> B \<in> parts (used v)"
by blast
hence "(\<exists>m SK. A = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> v \<and>
Crypt (SesK SK) (A \<otimes> B) \<in> parts (used v)) \<or>
{PriKey A, PriKey B} \<subseteq> spied v"
by (rule_tac parts_mult_start, simp_all)
moreover have "PriKey B \<notin> spied v"
proof
assume "PriKey B \<in> spied v"
hence "PriKey B \<in> spied s"
by (rule rev_subsetD, rule_tac spied_subset [OF H])
thus False
using D by contradiction
qed
ultimately obtain m SK where
J: "Crypt (SesK SK) (A \<otimes> B) \<in> parts (used v)" and
"A = Auth_PriK m" and "(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> v"
by blast
moreover from this have "(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s"
by (erule_tac rev_subsetD, rule_tac state_subset [OF H])
hence "m = n"
by (rule asset_ii_unique [OF A _ E])
ultimately have K: "Auth_PriKey n \<in> spied s"
using F by simp
have "Crypt (SesK SK) (A \<otimes> B) \<notin> parts (used u)"
using I by blast
hence "B \<in> fst (snd SK) \<and> (\<exists>m C. ?P m C SK v) \<or>
{A \<otimes> B, SesKey SK} \<subseteq> spied u"
by (rule parts_crypt_mult_start [OF G J])
moreover have "A \<otimes> B \<notin> spied u"
using I by blast
ultimately obtain m' C where "?P m' C SK v" by blast
hence "?P m' C SK s"
by (rule rev_subsetD, rule_tac state_subset [OF H])
moreover from this have "\<exists>A' D. fst (snd SK) = {A', B} \<and>
snd (snd SK) = {C, D} \<and> (Asset m', \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s \<and>
(Asset m', \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
Crypt (SesK SK) (PubKey D) \<in> used s \<and> (Asset m', PubKey B) \<in> s"
by (rule asset_iv_state [OF A])
hence "(Asset m', \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s" by blast
hence "m' = n"
by (rule asset_ii_unique [OF A _ E])
ultimately show ?thesis
using K by blast
qed
proposition asset_ii_spied:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "PriKey B \<in> spied s" and
C: "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s"
shows "Auth_PriKey n \<in> spied s \<and>
(\<exists>C SK. (Asset n, Token n (Auth_PriK n) B C SK) \<in> s)"
(is "?P s")
proof -
have "\<exists>u v. s\<^sub>0 \<Turnstile> u \<and> u \<turnstile> v \<and> v \<Turnstile> s \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<notin> u \<and> (Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> v"
using A and C by (rule rtrancl_start, auto)
then obtain u v where "v \<Turnstile> s" and "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<notin> u" and
D: "s\<^sub>0 \<Turnstile> u" and E: "u \<turnstile> v" and F: "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> v"
by blast
moreover from this have "PriKey B \<notin> spied v"
by (auto simp add: rel_def)
ultimately have "\<exists>w x. v \<Turnstile> w \<and> w \<turnstile> x \<and> x \<Turnstile> s \<and>
PriKey B \<notin> spied w \<and> PriKey B \<in> spied x"
using B by (rule_tac rtrancl_start, simp_all)
then obtain w x where "PriKey B \<notin> spied w" and "PriKey B \<in> spied x" and
G: "v \<Turnstile> w" and H: "w \<turnstile> x" and I: "x \<Turnstile> s"
by blast
moreover from this have "s\<^sub>0 \<Turnstile> w"
using D and E by simp
moreover have "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> w"
by (rule rev_subsetD [OF F], rule state_subset [OF G])
ultimately have "?P w"
by (rule_tac asset_ii_spied_start, simp_all)
moreover have "w \<subseteq> s"
using H and I by (rule_tac state_subset, simp)
ultimately show ?thesis by blast
qed
proposition asset_iv_unique:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "(Asset m, Token m (Auth_PriK m) B C' SK') \<in> s" and
C: "(Asset n, Token n (Auth_PriK n) B C SK) \<in> s"
(is "?P n C SK s")
shows "m = n \<and> C' = C \<and> SK' = SK"
proof (subst (2) cases_simp [of "m = n", symmetric], simp, rule conjI, rule impI,
rule ccontr)
assume D: "\<not> (C' = C \<and> SK' = SK)" and "m = n"
moreover have "\<exists>u v. s\<^sub>0 \<Turnstile> u \<and> u \<turnstile> v \<and> v \<Turnstile> s \<and>
\<not> (?P m C' SK' u \<and> ?P n C SK u) \<and> ?P m C' SK' v \<and> ?P n C SK v"
using B and C by (rule_tac rtrancl_start [OF A], auto)
ultimately obtain u v where E: "s\<^sub>0 \<Turnstile> u" and F: "u \<turnstile> v" and
G: "?P n C' SK' v" and H: "?P n C SK v" and
"\<not> ?P n C' SK' u \<or> \<not> ?P n C SK u"
by blast
moreover {
assume I: "\<not> ?P n C' SK' u"
hence "?P n C SK u"
by (insert D F G H, auto simp add: rel_def)
hence "\<exists>A D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> u \<and> (Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> u \<and>
Crypt (SesK SK) (PubKey D) \<in> used u \<and> (Asset n, PubKey B) \<in> u"
by (rule asset_iv_state [OF E])
moreover have "(Asset n, PubKey B) \<notin> u"
by (insert F G I, auto simp add: rel_def)
ultimately have False by simp
}
moreover {
assume I: "\<not> ?P n C SK u"
hence "?P n C' SK' u"
by (insert D F G H, auto simp add: rel_def)
hence "\<exists>A D. fst (snd SK') = {A, B} \<and> snd (snd SK') = {C', D} \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> u \<and> (Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> u \<and>
Crypt (SesK SK') (PubKey D) \<in> used u \<and> (Asset n, PubKey B) \<in> u"
by (rule asset_iv_state [OF E])
moreover have "(Asset n, PubKey B) \<notin> u"
by (insert F H I, auto simp add: rel_def)
ultimately have False by simp
}
ultimately show False by blast
next
have "\<exists>A D. fst (snd SK') = {A, B} \<and> snd (snd SK') = {C', D} \<and>
(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s \<and> (Asset m, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
Crypt (SesK SK') (PubKey D) \<in> used s \<and> (Asset m, PubKey B) \<in> s"
(is "?Q m C' SK'")
by (rule asset_iv_state [OF A B])
hence "(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s" by blast
moreover have "?Q n C SK"
by (rule asset_iv_state [OF A C])
hence "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s" by blast
ultimately show "m = n"
by (rule asset_ii_unique [OF A])
qed
theorem sigkey_secret:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow> SigKey \<notin> spied s"
proof (erule rtrancl_induct, simp add: image_def)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "SigKey \<notin> spied s"
show "SigKey \<notin> spied s'"
proof (insert B C, auto simp add: rel_def)
fix K
assume "(Spy, Crypt K SigKey) \<in> s"
hence "Crypt K SigKey \<in> parts (used s)" by blast
hence "{SigKey, Key K} \<subseteq> spied s"
by (rule parts_crypt_key [OF A], simp add: image_def)
with C show False by simp
next
fix Y
assume "(Spy, \<lbrace>SigKey, Y\<rbrace>) \<in> s"
hence "\<lbrace>SigKey, Y\<rbrace> \<in> parts (used s)" by blast
hence "{SigKey, Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "SigK"], simp)
with C show False by simp
next
fix X
assume "(Spy, \<lbrace>X, SigKey\<rbrace>) \<in> s"
hence "\<lbrace>X, SigKey\<rbrace> \<in> parts (used s)" by blast
hence "{X, SigKey} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "SigK"], simp add: image_def)
with C show False by simp
qed
qed
proposition parts_sign_start:
assumes A: "s\<^sub>0 \<Turnstile> s"
shows "\<lbrakk>s \<turnstile> s'; Sign n A \<in> parts (used s'); Sign n A \<notin> parts (used s)\<rbrakk> \<Longrightarrow>
A = Auth_PriK n"
by (simp add: rel_def, (((erule disjE)?, (erule exE)+, simp add: parts_insert image_iff)+,
((drule parts_dec | erule disjE, insert sigkey_secret [OF A], simp, drule parts_enc |
drule parts_sep | drule parts_con), simp+)?)+)
proposition parts_sign:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; Sign n A \<in> parts (used s)\<rbrakk> \<Longrightarrow>
A = Auth_PriK n"
by (rule classical, drule rtrancl_start, assumption, subst parts_init, simp add:
Range_iff image_def, (erule exE)+, (erule conjE)+, drule parts_sign_start)
theorem auth_shakey_secret:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; n \<notin> bad_shakey\<rbrakk> \<Longrightarrow>
Key (Auth_ShaKey n) \<notin> spied s"
proof (erule rtrancl_induct, simp add: image_def)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "Key (Auth_ShaKey n) \<notin> spied s"
show "Key (Auth_ShaKey n) \<notin> spied s'"
proof (insert B C, auto simp add: rel_def)
fix K
assume "(Spy, Crypt K (Key (ShaK (Auth_ShaK n)))) \<in> s"
hence "Crypt K (Key (Auth_ShaKey n)) \<in> parts (used s)" by auto
hence "{Key (Auth_ShaKey n), Key K} \<subseteq> spied s"
by (rule parts_crypt_key [OF A], simp add: image_def)
with C show False by simp
next
fix Y
assume "(Spy, \<lbrace>Key (ShaK (Auth_ShaK n)), Y\<rbrace>) \<in> s"
hence "\<lbrace>Key (Auth_ShaKey n), Y\<rbrace> \<in> parts (used s)" by auto
hence "{Key (Auth_ShaKey n), Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "Auth_ShaKey n"], simp)
with C show False by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (ShaK (Auth_ShaK n))\<rbrace>) \<in> s"
hence "\<lbrace>X, Key (Auth_ShaKey n)\<rbrace> \<in> parts (used s)" by auto
hence "{X, Key (Auth_ShaKey n)} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "Auth_ShaKey n"],
simp add: image_def)
with C show False by simp
qed
qed
theorem auth_prikey_secret:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_prikey"
shows "Auth_PriKey n \<notin> spied s"
proof
assume "Auth_PriKey n \<in> spied s"
moreover have "Auth_PriKey n \<notin> spied s\<^sub>0"
using B by auto
ultimately have "\<exists>u v. s\<^sub>0 \<Turnstile> u \<and> u \<turnstile> v \<and> v \<Turnstile> s \<and>
Auth_PriKey n \<notin> spied u \<and> Auth_PriKey n \<in> spied v"
by (rule rtrancl_start [OF A])
then obtain u v where C: "s\<^sub>0 \<Turnstile> u" and D: "u \<turnstile> v" and
E: "Auth_PriKey n \<notin> spied u" and F: "Auth_PriKey n \<in> spied v"
by blast
have "\<exists>B. (Auth_PriK n \<otimes> B \<in> spied u \<or> B \<otimes> Auth_PriK n \<in> spied u) \<and>
PriKey B \<in> spied u"
proof (insert D E F, auto simp add: rel_def, rule_tac [!] FalseE)
assume "Key (PriK (Auth_PriK n)) \<notin> used u"
moreover have "Auth_PriKey n \<in> used u"
by (rule auth_prikey_used [OF C])
ultimately show False by simp
next
fix K
assume "(Spy, Crypt K (Key (PriK (Auth_PriK n)))) \<in> u"
hence "Crypt K (PriKey (Auth_PriK n)) \<in> parts (used u)" by auto
hence "(\<exists>m. K = Auth_ShaKey m \<and>
(Asset m, Crypt (Auth_ShaKey m) (PriKey (Auth_PriK n))) \<in> u) \<or>
{PriKey (Auth_PriK n), Key K} \<subseteq> spied u"
by (rule parts_crypt_prikey [OF C])
then obtain m where
"(Asset m, Crypt (Auth_ShaKey m) (Auth_PriKey n)) \<in> u"
using E by auto
thus False
by (rule auth_prikey_asset_i [OF C])
next
fix Y
assume "(Spy, \<lbrace>Key (PriK (Auth_PriK n)), Y\<rbrace>) \<in> u"
hence "\<lbrace>Auth_PriKey n, Y\<rbrace> \<in> parts (used u)" by auto
hence "{Auth_PriKey n, Y} \<subseteq> spied u"
by (rule parts_mpair_key [OF C, where K = "PriK (Auth_PriK n)"], simp)
thus False
using E by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (PriK (Auth_PriK n))\<rbrace>) \<in> u"
hence "\<lbrace>X, Auth_PriKey n\<rbrace> \<in> parts (used u)" by auto
hence "{X, Auth_PriKey n} \<subseteq> spied u"
by (rule parts_mpair_key [OF C, where K = "PriK (Auth_PriK n)"], simp
add: image_def)
thus False
using E by simp
qed
then obtain B where G: "PriKey B \<in> spied u" and
"Auth_PriK n \<otimes> B \<in> spied u \<or> B \<otimes> Auth_PriK n \<in> spied u"
by blast
hence "Auth_PriK n \<otimes> B \<in> parts (used u) \<or> B \<otimes> Auth_PriK n \<in> parts (used u)"
by blast
moreover have "B \<otimes> Auth_PriK n \<notin> parts (used u)"
proof
assume "B \<otimes> Auth_PriK n \<in> parts (used u)"
hence "(\<exists>m. B = Auth_PriK m \<and>
(Asset m, \<lbrace>Num 2, PubKey (Auth_PriK n)\<rbrace>) \<in> u) \<or>
{PriKey B, PriKey (Auth_PriK n)} \<subseteq> spied u"
by (rule parts_mult [OF C])
then obtain m where "(Asset m, \<lbrace>Num 2, Auth_PubKey n\<rbrace>) \<in> u"
using E by auto
thus False
by (rule auth_pubkey_asset_ii [OF C])
qed
ultimately have "Auth_PriK n \<otimes> B \<in> parts (used u)" by simp
hence "(\<exists>m. Auth_PriK n = Auth_PriK m \<and>
(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> u) \<or>
{PriKey (Auth_PriK n), PriKey B} \<subseteq> spied u"
by (rule parts_mult [OF C])
then obtain m where "Auth_PriK n = Auth_PriK m" and
"(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> u"
using E by auto
moreover from this have "Auth_PriKey m \<in> spied u \<and>
(\<exists>C SK. (Asset m, Token m (Auth_PriK m) B C SK) \<in> u)"
by (rule_tac asset_ii_spied [OF C G])
ultimately show False
using E by simp
qed
proposition asset_ii_secret:
"\<lbrakk>s\<^sub>0 \<Turnstile> s; n \<notin> bad_prikey; (Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s\<rbrakk> \<Longrightarrow>
PriKey B \<notin> spied s"
by (rule notI, frule asset_ii_spied, assumption+, drule auth_prikey_secret, simp+)
proposition asset_i_secret [rule_format]:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey"
shows "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s \<longrightarrow>
PriKey S \<notin> spied s"
proof (rule rtrancl_induct [OF A], simp add: image_def, rule impI)
fix s s'
assume
C: "s\<^sub>0 \<Turnstile> s" and
D: "s \<turnstile> s'" and
E: "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s \<longrightarrow>
PriKey S \<notin> spied s" and
F: "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s'"
show "PriKey S \<notin> spied s'"
proof (insert D E F, auto simp add: rel_def)
assume "(Asset n, Crypt (ShaK (Auth_ShaK n)) (Key (PriK S))) \<in> s"
hence "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s" by simp
hence "PriKey S \<in> used s"
by (rule asset_i_used [OF C, THEN mp])
moreover assume "Key (PriK S) \<notin> used s"
ultimately show False by simp
next
fix K
assume "(Spy, Crypt K (Key (PriK S))) \<in> s"
hence "Crypt K (PriKey S) \<in> parts (used s)" by auto
hence "(\<exists>m. K = Auth_ShaKey m \<and>
(Asset m, Crypt (Auth_ShaKey m) (PriKey S)) \<in> s) \<or>
{PriKey S, Key K} \<subseteq> spied s"
(is "(\<exists>m. ?P m \<and> ?Q m) \<or> _")
by (rule parts_crypt_prikey [OF C])
moreover assume "(Spy, Key (PriK S)) \<notin> s"
ultimately obtain m where G: "?P m \<and> ?Q m" by auto
hence "?Q m" ..
moreover assume
"(Asset n, Crypt (ShaK (Auth_ShaK n)) (Key (PriK S))) \<in> s"
hence "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s" by simp
ultimately have "m = n"
by (rule asset_i_unique [OF C])
moreover assume "(Spy, Key (InvK K)) \<in> s"
ultimately have "Key (Auth_ShaKey n) \<in> spied s"
using G by simp
moreover have "Key (Auth_ShaKey n) \<notin> spied s"
by (rule auth_shakey_secret [OF C B])
ultimately show False by contradiction
next
fix B
assume "(Spy, S \<otimes> B) \<in> s"
hence "S \<otimes> B \<in> parts (used s)" by blast
hence "(\<exists>m. S = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s) \<or>
{PriKey S, PriKey B} \<subseteq> spied s"
(is "(\<exists>m. ?P m \<and> _) \<or> _")
by (rule parts_mult [OF C])
moreover assume "(Spy, Key (PriK S)) \<notin> s"
ultimately obtain m where "?P m" by auto
moreover assume
"(Asset n, Crypt (ShaK (Auth_ShaK n)) (Key (PriK S))) \<in> s"
ultimately have "(Asset n, Crypt (Auth_ShaKey n) (Auth_PriKey m)) \<in> s"
by simp
thus False
by (rule auth_prikey_asset_i [OF C])
next
fix A
assume "(Spy, A \<otimes> S) \<in> s"
hence "A \<otimes> S \<in> parts (used s)" by blast
hence "(\<exists>m. A = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey S\<rbrace>) \<in> s) \<or>
{PriKey A, PriKey S} \<subseteq> spied s"
(is "(\<exists>m. _ \<and> ?P m) \<or> _")
by (rule parts_mult [OF C])
moreover assume "(Spy, Key (PriK S)) \<notin> s"
ultimately obtain m where "?P m" by auto
assume "(Asset n, Crypt (ShaK (Auth_ShaK n)) (Key (PriK S))) \<in> s"
hence "(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s" by simp
thus False
- by (rule asset_i_asset_ii [OF C _ `?P m`])
+ by (rule asset_i_asset_ii [OF C _ \<open>?P m\<close>])
next
fix Y
assume "(Spy, \<lbrace>Key (PriK S), Y\<rbrace>) \<in> s"
hence "\<lbrace>PriKey S, Y\<rbrace> \<in> parts (used s)" by auto
hence "{PriKey S, Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF C, where K = "PriK S"], simp)
moreover assume "(Spy, Key (PriK S)) \<notin> s"
ultimately show False by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (PriK S)\<rbrace>) \<in> s"
hence "\<lbrace>X, PriKey S\<rbrace> \<in> parts (used s)" by auto
hence "{X, PriKey S} \<subseteq> spied s"
by (rule parts_mpair_key [OF C, where K = "PriK S"], simp add: image_def)
moreover assume "(Spy, Key (PriK S)) \<notin> s"
ultimately show False by simp
qed
qed
proposition owner_ii_secret [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow>
PriKey A \<notin> spied s"
proof (erule rtrancl_induct, simp add: image_def, rule impI)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<longrightarrow> PriKey A \<notin> spied s" and
D: "(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s'"
show "PriKey A \<notin> spied s'"
proof (insert B C D, auto simp add: rel_def)
assume "(Owner n, \<lbrace>Num (Suc 0), Key (PubK A)\<rbrace>) \<in> s"
hence "(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s" by simp
hence "PriKey A \<in> used s"
by (rule owner_ii_used [OF A, THEN mp])
moreover assume "Key (PriK A) \<notin> used s"
ultimately show False by simp
next
fix K
assume "(Spy, Crypt K (Key (PriK A))) \<in> s"
hence "Crypt K (PriKey A) \<in> parts (used s)" by auto
hence "(\<exists>m. K = Auth_ShaKey m \<and>
(Asset m, Crypt (Auth_ShaKey m) (PriKey A)) \<in> s) \<or>
{PriKey A, Key K} \<subseteq> spied s"
(is "(\<exists>m. _ \<and> ?P m) \<or> _")
by (rule parts_crypt_prikey [OF A])
moreover assume "(Spy, Key (PriK A)) \<notin> s"
ultimately obtain m where "?P m" by auto
moreover assume "(Owner n, \<lbrace>Num (Suc 0), Key (PubK A)\<rbrace>) \<in> s"
hence "(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s" by simp
ultimately show False
by (rule asset_i_owner_ii [OF A])
next
fix B
assume "(Spy, A \<otimes> B) \<in> s"
hence "A \<otimes> B \<in> parts (used s)" by blast
hence "(\<exists>m. A = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s) \<or>
{PriKey A, PriKey B} \<subseteq> spied s"
(is "(\<exists>m. ?P m \<and> _) \<or> _")
by (rule parts_mult [OF A])
moreover assume "(Spy, Key (PriK A)) \<notin> s"
ultimately obtain m where "?P m" by auto
moreover assume "(Owner n, \<lbrace>Num (Suc 0), Key (PubK A)\<rbrace>) \<in> s"
ultimately have "(Owner n, \<lbrace>Num 1, Auth_PubKey m\<rbrace>) \<in> s" by simp
thus False
by (rule auth_pubkey_owner_ii [OF A])
next
fix B
assume "(Spy, B \<otimes> A) \<in> s"
hence "B \<otimes> A \<in> parts (used s)" by blast
hence "(\<exists>m. B = Auth_PriK m \<and> (Asset m, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s) \<or>
{PriKey B, PriKey A} \<subseteq> spied s"
(is "(\<exists>m. _ \<and> ?P m) \<or> _")
by (rule parts_mult [OF A])
moreover assume "(Spy, Key (PriK A)) \<notin> s"
ultimately obtain m where "?P m" by auto
moreover assume "(Owner n, \<lbrace>Num (Suc 0), Key (PubK A)\<rbrace>) \<in> s"
hence "(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s" by simp
ultimately show False
by (rule asset_ii_owner_ii [OF A])
next
fix Y
assume "(Spy, \<lbrace>Key (PriK A), Y\<rbrace>) \<in> s"
hence "\<lbrace>PriKey A, Y\<rbrace> \<in> parts (used s)" by auto
hence "{PriKey A, Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "PriK A"], simp)
moreover assume "(Spy, Key (PriK A)) \<notin> s"
ultimately show False by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (PriK A)\<rbrace>) \<in> s"
hence "\<lbrace>X, PriKey A\<rbrace> \<in> parts (used s)" by auto
hence "{X, PriKey A} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "PriK A"], simp add: image_def)
moreover assume "(Spy, Key (PriK A)) \<notin> s"
ultimately show False by simp
qed
qed
proposition seskey_spied [rule_format]:
"s\<^sub>0 \<Turnstile> s \<Longrightarrow>
SesKey SK \<in> spied s \<longrightarrow>
(\<exists>S A C. fst SK = Some S \<and> A \<in> fst (snd SK) \<and> C \<in> snd (snd SK) \<and>
{PriKey S, PriKey A, PriKey C} \<subseteq> spied s)"
(is "_ \<Longrightarrow> _ \<longrightarrow> (\<exists>S A C. ?P S A C s)")
proof (erule rtrancl_induct, simp add: image_def, rule impI)
fix s s'
assume
A: "s\<^sub>0 \<Turnstile> s" and
B: "s \<turnstile> s'" and
C: "SesKey SK \<in> spied s \<longrightarrow> (\<exists>S A C. ?P S A C s)" and
D: "SesKey SK \<in> spied s'"
show "\<exists>S A C. ?P S A C s'"
proof (insert B C D, auto simp add: rel_def, blast, rule_tac [!] FalseE)
fix K
assume "(Spy, Crypt K (Key (SesK SK))) \<in> s"
hence "Crypt K (Key (SesK SK)) \<in> parts (used s)" by blast
hence "{Key (SesK SK), Key K} \<subseteq> spied s"
by (rule parts_crypt_key [OF A], simp add: image_def)
moreover assume "(Spy, Key (SesK SK)) \<notin> s"
ultimately show False by simp
next
fix Y
assume "(Spy, \<lbrace>Key (SesK SK), Y\<rbrace>) \<in> s"
hence "\<lbrace>SesKey SK, Y\<rbrace> \<in> parts (used s)" by auto
hence "{SesKey SK, Y} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "SesK SK"], simp)
moreover assume "(Spy, Key (SesK SK)) \<notin> s"
ultimately show False by simp
next
fix X
assume "(Spy, \<lbrace>X, Key (SesK SK)\<rbrace>) \<in> s"
hence "\<lbrace>X, SesKey SK\<rbrace> \<in> parts (used s)" by auto
hence "{X, SesKey SK} \<subseteq> spied s"
by (rule parts_mpair_key [OF A, where K = "SesK SK"], simp add: image_def)
moreover assume "(Spy, Key (SesK SK)) \<notin> s"
ultimately show False by simp
qed
qed
proposition owner_seskey_shakey:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey" and
C: "(Owner n, SesKey SK) \<in> s"
shows "SesKey SK \<notin> spied s"
proof
have "(\<exists>S. fst SK = Some S \<and> Crypt (Auth_ShaKey n) (PriKey S) \<in> used s) \<or>
fst SK = None"
(is "(\<exists>S. ?P S) \<or> _")
by (rule owner_seskey_nonce [OF A C])
moreover assume "SesKey SK \<in> spied s"
hence D: "\<exists>S A C. fst SK = Some S \<and> A \<in> fst (snd SK) \<and>
C \<in> snd (snd SK) \<and> {PriKey S, PriKey A, PriKey C} \<subseteq> spied s"
by (rule seskey_spied [OF A])
ultimately obtain S where "?P S" by auto
hence "Crypt (Auth_ShaKey n) (PriKey S) \<in> parts (used s)" by blast
hence "(\<exists>m. Auth_ShaKey n = Auth_ShaKey m \<and>
(Asset m, Crypt (Auth_ShaKey m) (PriKey S)) \<in> s) \<or>
{PriKey S, Key (Auth_ShaKey n)} \<subseteq> spied s"
(is "(\<exists>m. ?Q m \<and> ?R m) \<or> _")
by (rule parts_crypt_prikey [OF A])
moreover have "Key (Auth_ShaKey n) \<notin> spied s"
by (rule auth_shakey_secret [OF A B])
ultimately obtain m where "?Q m" and "?R m" by blast
hence "m \<notin> bad_shakey"
using B by simp
hence "PriKey S \<notin> spied s"
- by (rule asset_i_secret [OF A _ `?R m`])
+ by (rule asset_i_secret [OF A _ \<open>?R m\<close>])
moreover have "PriKey S \<in> spied s"
- using D and `?P S` by auto
+ using D and \<open>?P S\<close> by auto
ultimately show False by contradiction
qed
proposition owner_seskey_prikey:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_prikey" and
C: "(Owner m, SesKey SK) \<in> s" and
D: "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s" and
E: "B \<in> fst (snd SK)"
shows "SesKey SK \<notin> spied s"
proof
have "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner m, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner m, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner m, Crypt (SesK SK) (PubKey D)) \<in> s"
(is "\<exists>A B C D. ?P A B \<and> _ \<and> ?Q A \<and> _")
by (rule owner_seskey_other [OF A C])
then obtain A B' where "?P A B'" and "?Q A" by blast
assume "SesKey SK \<in> spied s"
hence "\<exists>S A' C. fst SK = Some S \<and> A' \<in> fst (snd SK) \<and> C \<in> snd (snd SK) \<and>
{PriKey S, PriKey A', PriKey C} \<subseteq> spied s"
(is "\<exists>S A' C. _ \<and> ?R A' \<and> _")
by (rule seskey_spied [OF A])
then obtain A' where "A' \<in> fst (snd SK)" and "PriKey A' \<in> spied s" (is "?R A'")
by blast
hence "{A', A, B} \<subseteq> {A, B'}"
- using E and `?P A B'` by simp
+ using E and \<open>?P A B'\<close> by simp
hence "card {A', A, B} \<le> card {A, B'}"
by (rule_tac card_mono, simp)
also have "\<dots> \<le> Suc (Suc 0)"
by (rule card_insert_le_m1, simp_all)
finally have "card {A', A, B} \<le> Suc (Suc 0)" .
moreover have "card {A', A, B} = Suc (card {A, B})"
proof (rule card_insert_disjoint, simp_all, rule conjI, rule_tac [!] notI)
assume "A' = A"
hence "?R A"
- using `?R A'` by simp
+ using \<open>?R A'\<close> by simp
moreover have "\<not> ?R A"
- by (rule owner_ii_secret [OF A `?Q A`])
+ by (rule owner_ii_secret [OF A \<open>?Q A\<close>])
ultimately show False by contradiction
next
assume "A' = B"
hence "?R B"
- using `?R A'` by simp
+ using \<open>?R A'\<close> by simp
moreover have "\<not> ?R B"
by (rule asset_ii_secret [OF A B D])
ultimately show False by contradiction
qed
moreover have "card {A, B} = Suc (card {B})"
proof (rule card_insert_disjoint, simp_all, rule notI)
assume "A = B"
hence "(Asset n, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s"
using D by simp
thus False
- by (rule asset_ii_owner_ii [OF A _ `?Q A`])
+ by (rule asset_ii_owner_ii [OF A _ \<open>?Q A\<close>])
qed
ultimately show False by simp
qed
proposition asset_seskey_shakey:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey" and
C: "(Asset n, SesKey SK) \<in> s"
shows "SesKey SK \<notin> spied s"
proof
have "\<exists>S. fst SK = Some S \<and>
(Asset n, Crypt (Auth_ShaKey n) (PriKey S)) \<in> s"
(is "\<exists>S. ?P S \<and> ?Q S")
by (rule asset_seskey_nonce [OF A C])
then obtain S where "?P S" and "?Q S" by blast
have "PriKey S \<notin> spied s"
- by (rule asset_i_secret [OF A B `?Q S`])
+ by (rule asset_i_secret [OF A B \<open>?Q S\<close>])
moreover assume "SesKey SK \<in> spied s"
hence "\<exists>S A C. fst SK = Some S \<and> A \<in> fst (snd SK) \<and> C \<in> snd (snd SK) \<and>
{PriKey S, PriKey A, PriKey C} \<subseteq> spied s"
by (rule seskey_spied [OF A])
hence "PriKey S \<in> spied s"
- using `?P S` by auto
+ using \<open>?P S\<close> by auto
ultimately show False by contradiction
qed
theorem owner_seskey_unique:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "(Owner m, Crypt (SesK SK) (Pwd m)) \<in> s" and
C: "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s"
shows "m = n"
proof (rule ccontr)
have D: "(Owner m, SesKey SK) \<in> s \<and>
(\<exists>A B C. Token m A B C SK \<in> used s \<and> B \<in> fst (snd SK))"
(is "?P m \<and> (\<exists>A B C. ?Q m A B C)")
by (rule owner_v_state [OF A B])
hence "?P m" by blast
hence "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner m, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner m, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner m, Crypt (SesK SK) (PubKey D)) \<in> s"
(is "\<exists>A B C D. ?R A B \<and> ?S C D \<and> ?T m A \<and> ?U m C D")
by (rule owner_seskey_other [OF A])
then obtain A B where "?R A B" and "?T m A" by blast
have "?P n \<and> (\<exists>A B C. ?Q n A B C)"
by (rule owner_v_state [OF A C])
hence "?P n" by blast
hence "\<exists>A B C D. ?R A B \<and> ?S C D \<and> ?T n A \<and> ?U n C D"
by (rule owner_seskey_other [OF A])
then obtain A' B' where "?R A' B'" and "?T n A'" by blast
from D obtain A'' B'' C where "?Q m A'' B'' C" by blast
hence "Token m A'' B'' C SK \<in> parts (used s)" by blast
hence "Crypt (SesK SK) (A'' \<otimes> B'') \<in> parts (used s)" by blast
hence "B'' \<in> fst (snd SK) \<and>
(\<exists>i C'. (Asset i, Token i (Auth_PriK i) B'' C' SK) \<in> s) \<or>
{A'' \<otimes> B'', SesKey SK} \<subseteq> spied s"
(is "?V B'' \<and> (\<exists>i C'. ?W i B'' C') \<or> _")
by (rule parts_crypt_mult [OF A])
hence "\<exists>D. ?V D \<and> D \<notin> {A, A'}"
proof (rule disjE, (erule_tac conjE, ((erule_tac exE)+)?)+)
fix i C'
assume "?V B''" and "?W i B'' C'"
have "\<exists>A D. ?R A B'' \<and> ?S C' D \<and>
(Asset i, \<lbrace>Num 2, PubKey B''\<rbrace>) \<in> s \<and> (Asset i, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
Crypt (SesK SK) (PubKey D) \<in> used s \<and> (Asset i, PubKey B'') \<in> s"
(is "\<exists>A D. _ \<and> _ \<and> ?X i B'' \<and> _")
- by (rule asset_iv_state [OF A `?W i B'' C'`])
+ by (rule asset_iv_state [OF A \<open>?W i B'' C'\<close>])
hence "?X i B''" by blast
have "B'' \<noteq> A"
proof
assume "B'' = A"
hence "?X i A"
- using `?X i B''` by simp
+ using \<open>?X i B''\<close> by simp
thus False
- by (rule asset_ii_owner_ii [OF A _ `?T m A`])
+ by (rule asset_ii_owner_ii [OF A _ \<open>?T m A\<close>])
qed
moreover have "B'' \<noteq> A'"
proof
assume "B'' = A'"
hence "?X i A'"
- using `?X i B''` by simp
+ using \<open>?X i B''\<close> by simp
thus False
- by (rule asset_ii_owner_ii [OF A _ `?T n A'`])
+ by (rule asset_ii_owner_ii [OF A _ \<open>?T n A'\<close>])
qed
ultimately show ?thesis
- using `?V B''` by blast
+ using \<open>?V B''\<close> by blast
next
assume "{A'' \<otimes> B'', SesKey SK} \<subseteq> spied s"
hence "SesKey SK \<in> spied s" by simp
hence "\<exists>S D E. fst SK = Some S \<and> ?V D \<and> E \<in> snd (snd SK) \<and>
{PriKey S, PriKey D, PriKey E} \<subseteq> spied s"
by (rule seskey_spied [OF A])
then obtain D where "?V D" and "PriKey D \<in> spied s" (is "?X D")
by blast
moreover have "D \<noteq> A"
proof
assume "D = A"
hence "?X A"
- using `?X D` by simp
+ using \<open>?X D\<close> by simp
moreover have "\<not> ?X A"
- by (rule owner_ii_secret [OF A `?T m A`])
+ by (rule owner_ii_secret [OF A \<open>?T m A\<close>])
ultimately show False by contradiction
qed
moreover have "D \<noteq> A'"
proof
assume "D = A'"
hence "?X A'"
- using `?X D` by simp
+ using \<open>?X D\<close> by simp
moreover have "\<not> ?X A'"
- by (rule owner_ii_secret [OF A `?T n A'`])
+ by (rule owner_ii_secret [OF A \<open>?T n A'\<close>])
ultimately show False by contradiction
qed
ultimately show ?thesis by blast
qed
then obtain D where "?V D" and E: "D \<notin> {A, A'}" by blast
hence "{D, A, A'} \<subseteq> {A, B}"
- using `?R A B` and `?R A' B'` by blast
+ using \<open>?R A B\<close> and \<open>?R A' B'\<close> by blast
hence "card {D, A, A'} \<le> card {A, B}"
by (rule_tac card_mono, simp)
also have "\<dots> \<le> Suc (Suc 0)"
by (rule card_insert_le_m1, simp_all)
finally have "card {D, A, A'} \<le> Suc (Suc 0)" .
moreover have "card {D, A, A'} = Suc (card {A, A'})"
by (rule card_insert_disjoint [OF _ E], simp)
moreover assume "m \<noteq> n"
hence "card {A, A'} = Suc (card {A'})"
proof (rule_tac card_insert_disjoint, simp_all, erule_tac contrapos_nn)
assume "A = A'"
hence "?T n A"
- using `?T n A'` by simp
+ using \<open>?T n A'\<close> by simp
thus "m = n"
- by (rule owner_ii_unique [OF A `?T m A`])
+ by (rule owner_ii_unique [OF A \<open>?T m A\<close>])
qed
ultimately show False by simp
qed
theorem owner_seskey_secret:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> bad_prikey" and
C: "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s"
shows "SesKey SK \<notin> spied s"
proof -
have "(Owner n, SesKey SK) \<in> s \<and>
(\<exists>A B C. Token n A B C SK \<in> used s \<and> B \<in> fst (snd SK))"
(is "?P \<and> (\<exists>A B C. ?Q A B C \<and> ?R B)")
by (rule owner_v_state [OF A C])
then obtain A B C where "?P" and "?Q A B C" and "?R B" by blast
have "n \<in> bad_shakey \<or> n \<notin> bad_shakey" by simp
moreover {
assume "n \<in> bad_shakey"
hence D: "n \<notin> bad_prikey"
using B by simp
hence "Auth_PriKey n \<notin> spied s"
by (rule auth_prikey_secret [OF A])
moreover have "Sign n A \<in> parts (used s)"
- using `?Q A B C` by blast
+ using \<open>?Q A B C\<close> by blast
hence "A = Auth_PriK n"
by (rule parts_sign [OF A])
hence "?Q (Auth_PriK n) B C"
- using `?Q A B C` by simp
+ using \<open>?Q A B C\<close> by simp
hence "Auth_PriK n \<otimes> B \<in> parts (used s)" by blast
hence "(\<exists>m. Auth_PriK n = Auth_PriK m \<and>
(Asset m, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s) \<or>
{PriKey (Auth_PriK n), PriKey B} \<subseteq> spied s"
(is "(\<exists>m. ?S m \<and> ?T m) \<or> _")
by (rule parts_mult [OF A])
ultimately obtain m where "?S m" and "?T m" by auto
hence "m \<notin> bad_prikey"
using D by simp
hence ?thesis
- by (rule owner_seskey_prikey [OF A _ `?P` `?T m` `?R B`])
+ by (rule owner_seskey_prikey [OF A _ \<open>?P\<close> \<open>?T m\<close> \<open>?R B\<close>])
}
moreover {
assume "n \<notin> bad_shakey"
hence ?thesis
- by (rule owner_seskey_shakey [OF A _ `?P`])
+ by (rule owner_seskey_shakey [OF A _ \<open>?P\<close>])
}
ultimately show ?thesis ..
qed
theorem owner_num_genuine:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> bad_prikey" and
C: "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s" and
D: "Crypt (SesK SK) (Num 0) \<in> used s"
shows "(Asset n, Crypt (SesK SK) (Num 0)) \<in> s"
proof -
have "Crypt (SesK SK) (Num 0) \<in> parts (used s)"
using D ..
hence "(\<exists>m. (Asset m, Crypt (SesK SK) (Num 0)) \<in> s) \<or>
SesKey SK \<in> spied s"
by (rule parts_crypt_num [OF A])
moreover have E: "SesKey SK \<notin> spied s"
by (rule owner_seskey_secret [OF A B C])
ultimately obtain m where "(Asset m, Crypt (SesK SK) (Num 0)) \<in> s"
by blast
moreover from this have "(Asset m, SesKey SK) \<in> s \<and>
Crypt (SesK SK) (Pwd m) \<in> used s"
by (rule asset_v_state [OF A])
hence "Crypt (SesK SK) (Pwd m) \<in> parts (used s)" by blast
hence "(\<exists>SK'. SesK SK = SesK SK' \<and>
(Owner m, Crypt (SesK SK') (Pwd m)) \<in> s) \<or>
{Pwd m, Key (SesK SK)} \<subseteq> spied s"
by (rule parts_crypt_pwd [OF A])
hence "(Owner m, Crypt (SesK SK) (Pwd m)) \<in> s"
using E by simp
hence "m = n"
by (rule owner_seskey_unique [OF A _ C])
ultimately show ?thesis by simp
qed
theorem owner_token_genuine:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> bad_prikey" and
C: "(Owner n, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s" and
D: "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s" and
E: "Token n A B C SK \<in> used s"
shows "A = Auth_PriK n \<and> B \<in> fst (snd SK) \<and> C \<in> snd (snd SK) \<and>
(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s \<and> (Asset n, Token n A B C SK) \<in> s"
(is "?P n A \<and> ?Q B \<and> ?R C \<and> ?S n B \<and> _")
proof -
have "Crypt (SesK SK) (Sign n A) \<in> parts (used s)"
using E by blast
hence "(Asset n, SesKey SK) \<in> s \<or> SesKey SK \<in> spied s"
by (rule parts_crypt_sign [OF A])
moreover have "SesKey SK \<notin> spied s"
by (rule owner_seskey_secret [OF A B D])
ultimately have "(Asset n, SesKey SK) \<in> s" by simp
hence "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
?S n B \<and> (Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
(Asset n, Token n (Auth_PriK n) B C SK) \<in> s"
(is "\<exists>A B C D. ?T A B \<and> ?U C D \<and> _ \<and> ?V n D \<and> ?W n B C")
by (rule asset_seskey_other [OF A])
then obtain A' B' C' D where
"?T A' B'" and "?U C' D" and "?S n B'" and "?V n D" and "?W n B' C'"
by blast
have "Sign n A \<in> parts (used s)"
using E by blast
hence "?P n A"
by (rule parts_sign [OF A])
have "Crypt (SesK SK) (A \<otimes> B) \<in> parts (used s)"
using E by blast
hence "?Q B \<and> (\<exists>m C''. ?W m B C'') \<or> {A \<otimes> B, SesKey SK} \<subseteq> spied s"
by (rule parts_crypt_mult [OF A])
moreover have F: "SesKey SK \<notin> spied s"
by (rule owner_seskey_secret [OF A B D])
ultimately obtain m C'' where "?Q B" and "?W m B C''" by blast
have "\<exists>A D. ?T A B \<and> ?U C'' D \<and> ?S m B \<and> ?V m D \<and>
Crypt (SesK SK) (PubKey D) \<in> used s \<and> (Asset m, PubKey B) \<in> s"
- by (rule asset_iv_state [OF A `?W m B C''`])
+ by (rule asset_iv_state [OF A \<open>?W m B C''\<close>])
hence "?S m B" by blast
have "(Owner n, SesKey SK) \<in> s \<and>
(\<exists>A B C. Token n A B C SK \<in> used s \<and> B \<in> fst (snd SK))"
by (rule owner_v_state [OF A D])
hence "(Owner n, SesKey SK) \<in> s" by blast
hence "\<exists>A B C D. ?T A B \<and> ?U C D \<and>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner n, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner n, Crypt (SesK SK) (PubKey D)) \<in> s"
(is "\<exists>A B C D. _ \<and> _ \<and> ?X A \<and> _")
by (rule owner_seskey_other [OF A])
then obtain A'' where "?Q A''" and "?X A''" by blast
have G: "B' = B"
proof (rule ccontr)
have "{A'', B', B} \<subseteq> {A', B'}"
- using `?T A' B'` and `?Q B` and `?Q A''` by simp
+ using \<open>?T A' B'\<close> and \<open>?Q B\<close> and \<open>?Q A''\<close> by simp
hence "card {A'', B', B} \<le> card {A', B'}"
by (rule_tac card_mono, simp)
also have "\<dots> \<le> Suc (Suc 0)"
by (rule card_insert_le_m1, simp_all)
finally have "card {A'', B', B} \<le> Suc (Suc 0)" .
moreover have "A'' \<notin> {B', B}"
proof (simp, rule conjI, rule_tac [!] notI)
assume "A'' = B'"
hence "?S n A''"
- using `?S n B'` by simp
+ using \<open>?S n B'\<close> by simp
thus False
- by (rule asset_ii_owner_ii [OF A _ `?X A''`])
+ by (rule asset_ii_owner_ii [OF A _ \<open>?X A''\<close>])
next
assume "A'' = B"
hence "?S m A''"
- using `?S m B` by simp
+ using \<open>?S m B\<close> by simp
thus False
- by (rule asset_ii_owner_ii [OF A _ `?X A''`])
+ by (rule asset_ii_owner_ii [OF A _ \<open>?X A''\<close>])
qed
hence "card {A'', B', B} = Suc (card {B', B})"
by (rule_tac card_insert_disjoint, simp)
moreover assume "B' \<noteq> B"
hence "card {B', B} = Suc (card {B})"
by (rule_tac card_insert_disjoint, simp_all)
ultimately show False by simp
qed
hence "?S n B"
- using `?S n B'` by simp
+ using \<open>?S n B'\<close> by simp
have "Crypt (SesK SK) (PubKey C) \<in> parts (used s)"
using E by blast
hence "?R C \<and> ((\<exists>n. (Owner n, SesKey SK) \<in> s) \<or> (\<exists>n B. ?W n B C)) \<or>
SesKey SK \<in> spied s"
by (rule parts_crypt_pubkey [OF A])
hence "?R C"
using F by simp
hence "C \<in> {C', D}"
- using `?U C' D` by simp
+ using \<open>?U C' D\<close> by simp
moreover have "C \<noteq> D"
proof
assume "C = D"
hence "?V n C"
- using `?V n D` by simp
+ using \<open>?V n D\<close> by simp
thus False
by (rule asset_iii_owner_iii [OF A _ C])
qed
ultimately have "C = C'" by simp
hence "(Asset n, Token n A B C SK) \<in> s"
- using G and `?P n A` and `?W n B' C'` by simp
+ using G and \<open>?P n A\<close> and \<open>?W n B' C'\<close> by simp
thus ?thesis
- using `?P n A` and `?Q B` and `?R C` and `?S n B` by simp
+ using \<open>?P n A\<close> and \<open>?Q B\<close> and \<open>?R C\<close> and \<open>?S n B\<close> by simp
qed
theorem pwd_secret:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_pwd \<union> bad_shakey \<inter> bad_prikey"
shows "Pwd n \<notin> spied s"
proof (rule rtrancl_induct [OF A], insert B, simp add: image_def)
fix s s'
assume
C: "s\<^sub>0 \<Turnstile> s" and
D: "s \<turnstile> s'" and
E: "Pwd n \<notin> spied s"
show "Pwd n \<notin> spied s'"
proof (insert D E, auto simp add: rel_def)
fix K
assume "(Spy, Crypt K (Pwd n)) \<in> s"
hence "Crypt K (Pwd n) \<in> parts (used s)" by blast
hence "(\<exists>SK. K = SesK SK \<and> (Owner n, Crypt (SesK SK) (Pwd n)) \<in> s) \<or>
{Pwd n, Key K} \<subseteq> spied s"
(is "(\<exists>SK. ?P SK \<and> ?Q SK) \<or> _")
by (rule parts_crypt_pwd [OF C])
then obtain SK where "?P SK" and "?Q SK"
using E by blast
have "n \<notin> bad_shakey \<inter> bad_prikey"
using B by simp
hence "SesKey SK \<notin> spied s"
- by (rule owner_seskey_secret [OF C _ `?Q SK`])
+ by (rule owner_seskey_secret [OF C _ \<open>?Q SK\<close>])
moreover assume "(Spy, Key (InvK K)) \<in> s"
ultimately show False
- using `?P SK` by simp
+ using \<open>?P SK\<close> by simp
next
fix Y
assume "(Spy, \<lbrace>Pwd n, Y\<rbrace>) \<in> s"
hence "\<lbrace>Pwd n, Y\<rbrace> \<in> parts (used s)" by blast
hence "{Pwd n, Y} \<subseteq> spied s"
by (rule parts_mpair_pwd [OF C, where n = n], simp)
with E show False by simp
next
fix X
assume "(Spy, \<lbrace>X, Pwd n\<rbrace>) \<in> s"
hence "\<lbrace>X, Pwd n\<rbrace> \<in> parts (used s)" by blast
hence "{X, Pwd n} \<subseteq> spied s"
by (rule parts_mpair_pwd [OF C, where n = n], simp)
with E show False by simp
qed
qed
theorem asset_seskey_unique:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "(Asset m, Token m (Auth_PriK m) B' C' SK) \<in> s" and
C: "(Asset n, Token n (Auth_PriK n) B C SK) \<in> s"
(is "?P n B C SK s")
shows "m = n \<and> B' = B \<and> C' = C"
proof (subst (2) cases_simp [of "B' = B", symmetric], simp, rule conjI, rule impI,
insert B C, simp only:, drule asset_iv_unique [OF A], simp, simp, rule ccontr)
assume "B' \<noteq> B"
moreover have "\<exists>A D. fst (snd SK) = {A, B'} \<and> snd (snd SK) = {C', D} \<and>
(Asset m, \<lbrace>Num 2, PubKey B'\<rbrace>) \<in> s \<and> (Asset m, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s \<and>
Crypt (SesK SK) (PubKey D) \<in> used s \<and> (Asset m, PubKey B') \<in> s"
(is "?Q m B' C'")
by (rule asset_iv_state [OF A B])
then obtain A where "fst (snd SK) = {A, B'}" and
"(Asset m, \<lbrace>Num 2, PubKey B'\<rbrace>) \<in> s"
by blast
moreover have "?Q n B C"
by (rule asset_iv_state [OF A C])
hence "B \<in> fst (snd SK)" and "(Asset n, \<lbrace>Num 2, PubKey B\<rbrace>) \<in> s"
by auto
ultimately have D: "\<forall>A \<in> fst (snd SK).
\<exists>i C. (Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s \<and> ?P i A C SK s"
using B and C by auto
have "Crypt (SesK SK) (PubKey C) \<in> parts (used s)"
using C by blast
thus False
proof (rule parts_pubkey_false [OF A], rule_tac allI, rule_tac [!] notI)
fix i
assume "(Owner i, SesKey SK) \<in> s"
hence "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner i, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner i, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner i, Crypt (SesK SK) (PubKey D)) \<in> s"
by (rule owner_seskey_other [OF A])
then obtain A where "A \<in> fst (snd SK)" and
E: "(Owner i, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s"
by blast
then obtain j where "(Asset j, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s"
using D by blast
thus False
by (rule asset_ii_owner_ii [OF A _ E])
next
assume "SesKey SK \<in> spied s"
hence "\<exists>S A C. fst SK = Some S \<and> A \<in> fst (snd SK) \<and> C \<in> snd (snd SK) \<and>
{PriKey S, PriKey A, PriKey C} \<subseteq> spied s"
(is "?R s")
by (rule seskey_spied [OF A])
moreover have "\<not> (\<exists>A \<in> fst (snd SK). PriKey A \<in> spied s)"
(is "\<not> ?S s")
proof
assume "?S s"
moreover have "\<not> ?S s\<^sub>0"
by (subst bex_simps, rule ballI, drule bspec [OF D], (erule exE)+,
erule conjE, rule asset_ii_init [OF A])
ultimately have "\<exists>u v. s\<^sub>0 \<Turnstile> u \<and> u \<turnstile> v \<and> v \<Turnstile> s \<and> \<not> ?S u \<and> ?S v"
by (rule rtrancl_start [OF A])
then obtain u v A where E: "s\<^sub>0 \<Turnstile> u" and F: "u \<turnstile> v" and G: "v \<Turnstile> s" and
H: "\<not> ?S u" and I: "A \<in> fst (snd SK)" and J: "PriKey A \<notin> spied u" and
K: "PriKey A \<in> spied v"
by blast
then obtain i where "(Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s"
using D by blast
hence "(Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> v"
proof (rule_tac ccontr, drule_tac rtrancl_start [OF G], simp,
(erule_tac exE)+, (erule_tac conjE)+)
fix w x
assume "w \<turnstile> x" and "(Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<notin> w" and
"(Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> x"
hence "PriKey A \<notin> spied w"
by (auto simp add: rel_def)
moreover assume "v \<Turnstile> w"
hence "PriKey A \<in> spied w"
by (rule_tac rev_subsetD [OF K], rule spied_subset)
ultimately show False by contradiction
qed
hence "(Asset i, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> u"
using F and K by (auto simp add: rel_def)
hence "Auth_PriKey i \<in> spied u \<and> (\<exists>C SK. ?P i A C SK u)"
by (rule asset_ii_spied_start [OF E F K J])
then obtain C' SK' where L: "?P i A C' SK' u" by blast
moreover have M: "u \<Turnstile> s"
using F and G by simp
ultimately have "?P i A C' SK' s"
by (erule_tac rev_subsetD, rule_tac state_subset)
moreover obtain j C where "?P j A C SK s"
using D and I by blast
ultimately have "i = j \<and> C' = C \<and> SK' = SK"
by (rule asset_iv_unique [OF A])
hence "Crypt (SesK SK) (PubKey C) \<in> parts (used u)"
using L by blast
thus False
proof (rule parts_pubkey_false [OF E], rule_tac allI, rule_tac [!] notI)
fix i
assume "(Owner i, SesKey SK) \<in> u"
hence "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner i, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> u \<and>
(Owner i, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> u \<and>
(Owner i, Crypt (SesK SK) (PubKey D)) \<in> u"
by (rule owner_seskey_other [OF E])
then obtain A where "A \<in> fst (snd SK)" and
N: "(Owner i, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> u"
by blast
then obtain j where "(Asset j, \<lbrace>Num 2, PubKey A\<rbrace>) \<in> s"
using D by blast
moreover have "(Owner i, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s"
by (rule rev_subsetD [OF N], rule state_subset [OF M])
ultimately show False
by (rule asset_ii_owner_ii [OF A])
next
assume "SesKey SK \<in> spied u"
hence "?R u"
by (rule seskey_spied [OF E])
thus False
using H by blast
qed
qed
ultimately show False by blast
qed
qed
theorem asset_seskey_secret:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> (bad_pwd \<union> bad_prikey)" and
C: "(Asset n, Crypt (SesK SK) (Num 0)) \<in> s"
shows "SesKey SK \<notin> spied s"
proof -
have D: "(Asset n, SesKey SK) \<in> s \<and> Crypt (SesK SK) (Pwd n) \<in> used s"
by (rule asset_v_state [OF A C])
have "n \<in> bad_shakey \<or> n \<notin> bad_shakey" by simp
moreover {
assume "n \<in> bad_shakey"
hence "Pwd n \<notin> spied s"
using B by (rule_tac pwd_secret [OF A], simp)
moreover have "Crypt (SesK SK) (Pwd n) \<in> parts (used s)"
using D by blast
hence "(\<exists>SK'. SesK SK = SesK SK' \<and>
(Owner n, Crypt (SesK SK') (Pwd n)) \<in> s) \<or>
{Pwd n, Key (SesK SK)} \<subseteq> spied s"
by (rule parts_crypt_pwd [OF A])
ultimately have "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s" by simp
hence ?thesis
using B by (rule_tac owner_seskey_secret [OF A], simp_all)
}
moreover {
assume "n \<notin> bad_shakey"
hence ?thesis
using D by (rule_tac asset_seskey_shakey [OF A], simp_all)
}
ultimately show ?thesis ..
qed
theorem asset_pwd_genuine:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> (bad_pwd \<union> bad_prikey)" and
C: "(Asset n, Crypt (SesK SK) (Num 0)) \<in> s"
shows "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s"
proof -
have "(Asset n, SesKey SK) \<in> s \<and> Crypt (SesK SK) (Pwd n) \<in> used s"
by (rule asset_v_state [OF A C])
hence "Crypt (SesK SK) (Pwd n) \<in> parts (used s)" by blast
hence "(\<exists>SK'. SesK SK = SesK SK' \<and>
(Owner n, Crypt (SesK SK') (Pwd n)) \<in> s) \<or>
{Pwd n, Key (SesK SK)} \<subseteq> spied s"
by (rule parts_crypt_pwd [OF A])
moreover have "SesKey SK \<notin> spied s"
by (rule asset_seskey_secret [OF A B C])
ultimately show ?thesis by simp
qed
theorem asset_token_genuine:
assumes
A: "s\<^sub>0 \<Turnstile> s" and
B: "n \<notin> bad_shakey \<inter> (bad_pwd \<union> bad_prikey)" and
C: "(Asset n, \<lbrace>Num 4, PubKey D\<rbrace>) \<in> s" and
D: "(Asset n, Crypt (SesK SK) (Num 0)) \<in> s" and
E: "D \<in> snd (snd SK)"
shows "(Owner n, Crypt (SesK SK) (PubKey D)) \<in> s"
proof -
have "(Owner n, Crypt (SesK SK) (Pwd n)) \<in> s"
by (rule asset_pwd_genuine [OF A B D])
hence "(Owner n, SesKey SK) \<in> s \<and>
(\<exists>A B C. Token n A B C SK \<in> used s \<and> B \<in> fst (snd SK))"
by (rule owner_v_state [OF A])
hence "(Owner n, SesKey SK) \<in> s" ..
hence "\<exists>A B C D. fst (snd SK) = {A, B} \<and> snd (snd SK) = {C, D} \<and>
(Owner n, \<lbrace>Num 1, PubKey A\<rbrace>) \<in> s \<and>
(Owner n, \<lbrace>Num 3, PubKey C\<rbrace>) \<in> s \<and>
(Owner n, Crypt (SesK SK) (PubKey D)) \<in> s"
(is "\<exists>A B C D. _ \<and> ?P C D \<and> _ \<and> ?Q C \<and> ?R D")
by (rule owner_seskey_other [OF A])
then obtain C D' where "?P C D'" and "?Q C" and "?R D'" by blast
have "D \<noteq> C"
proof
assume "D = C"
hence "?Q D"
- using `?Q C` by simp
+ using \<open>?Q C\<close> by simp
thus False
by (rule asset_iii_owner_iii [OF A C])
qed
hence "D = D'"
- using E and `?P C D'` by simp
+ using E and \<open>?P C D'\<close> by simp
thus ?thesis
- using `?R D'` by simp
+ using \<open>?R D'\<close> by simp
qed
end
\ No newline at end of file
diff --git a/thys/Safe_Distance/Safe_Distance.thy b/thys/Safe_Distance/Safe_Distance.thy
--- a/thys/Safe_Distance/Safe_Distance.thy
+++ b/thys/Safe_Distance/Safe_Distance.thy
@@ -1,1472 +1,1472 @@
\<^marker>\<open>creator "Albert Rizaldi" "Fabian Immler\<close>
section \<open>Safe Distance\<close>
theory Safe_Distance
imports
"HOL-Analysis.Multivariate_Analysis"
"HOL-Decision_Procs.Approximation"
Sturm_Sequences.Sturm
begin
text \<open>
This theory is about formalising the safe distance rule. The safe distance rule is obtained
from Vienna Convention which basically states the following thing.
``The car at all times must maintain a safe distance towards the vehicle in front of it,
such that whenever the vehicle in front and the ego vehicle apply maximum deceleration,
there will not be a collision.''
To formalise this safe distance rule we have to define first what is a safe distance.
To define this safe distance, we have to model the physics of the movement of the vehicle.
The following model is sufficient.
\<open>s = s\<^sub>0 + v\<^sub>0 * t + 1 / 2 * a\<^sub>0 * t\<^sup>2\<close>
Assumptions in this model are :
\<^item> Both vehicles are assumed to be point mass. The exact location of the ego vehicle
is the front-most occupancy of the ego vehicle. Similarly for the other vehicle,
its exact location is the rearmost occupancy of the other vehicle.
\<^item> Both cars can never drive backward.
\<close>
lemmas [simp del] = div_mult_self1 div_mult_self2 div_mult_self3 div_mult_self4
subsection \<open>Quadratic Equations\<close>
lemma discriminant: "a * x\<^sup>2 + b * x + c = (0::real) \<Longrightarrow> 0 \<le> b\<^sup>2 - 4 * a * c"
by (sos "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
lemma quadratic_eq_factoring:
assumes D : "D = b\<^sup>2 - 4 * a * c"
assumes nn: "0 \<le> D"
assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)"
assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)"
assumes a : "a \<noteq> 0"
shows "a * x\<^sup>2 + b * x + c = a * (x - x\<^sub>1) * (x - x\<^sub>2)"
using nn
by (simp add: D x1 x2)
(simp add: assms power2_eq_square power3_eq_cube field_split_simps)
lemma quadratic_eq_zeroes_iff:
assumes D : "D = b\<^sup>2 - 4 * a * c"
assumes x1: "x\<^sub>1 = (-b + sqrt D) / (2 * a)"
assumes x2: "x\<^sub>2 = (-b - sqrt D) / (2 * a)"
assumes a : "a \<noteq> 0"
shows "a * x\<^sup>2 + b * x + c = 0 \<longleftrightarrow> (D \<ge> 0 \<and> (x = x\<^sub>1 \<or> x = x\<^sub>2))" (is "?z \<longleftrightarrow> _")
using quadratic_eq_factoring[OF D _ x1 x2 a, of x] discriminant[of a x b c] a
by (auto simp: D)
subsection \<open>Convexity Condition\<close>
lemma p_convex:
fixes a b c x y z :: real
assumes p_def: "p = (\<lambda>x. a * x\<^sup>2 + b * x + c)"
assumes less : "x < y" "y < z"
and ge : "p x > p y" "p y \<le> p z"
shows "a > 0"
using less ge unfolding p_def
by (sos "((((A<0 * (A<1 * A<2)) * R<1) + (((A<2 * R<1) * (R<1/4 * [y + ~1*z]^2)) +
(((A<=1 * R<1) * (R<1 * [x + ~1*y]^2)) + (((A<=1 * (A<0 * (A<1 * R<1))) * (R<1/4 * [1]^2)) +
(((A<=0 * R<1) * (R<1/4 * [~1*y^2 + x*y + ~1*x*z + y*z]^2)) +
((A<=0 * (A<0 * (A<1 * R<1))) * (R<1 * [x + ~1/2*y + ~1/2*z]^2))))))))")
definition root_in :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool" where
"root_in m M f = (\<exists>x\<in>{m .. M}. f x = 0)"
definition quadroot_in :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"quadroot_in m M a b c = root_in m M (\<lambda>x. a * x^2 + b * x + c)"
lemma card_iff_exists: "0 < card X \<longleftrightarrow> finite X \<and> (\<exists>x. x \<in> X)"
by (auto simp: card_gt_0_iff)
lemma quadroot_in_sturm[code]:
"quadroot_in m M a b c \<longleftrightarrow> (a = 0 \<and> b = 0 \<and> c = 0 \<and> m \<le> M) \<or>
(m \<le> M \<and> poly [:c, b, a:] m = 0) \<or>
count_roots_between [:c, b, a:] m M > 0"
apply (cases "a = 0 \<and> b = 0 \<and> c = 0 \<and> m \<le> M")
apply (force simp: quadroot_in_def root_in_def)
apply (cases "m \<le> M \<and> poly [:c, b, a:] m = 0")
apply (force simp: quadroot_in_def root_in_def algebra_simps power2_eq_square count_roots_between_correct card_iff_exists)
proof -
assume H: "\<not> (a = 0 \<and> b = 0 \<and> c = 0 \<and> m \<le> M)" "\<not> (m \<le> M \<and> poly [:c, b, a:] m = 0)"
hence "poly [:c, b, a:] m \<noteq> 0 \<or> m > M"
by auto
then have "quadroot_in m M a b c \<longleftrightarrow> 0 < count_roots_between [:c, b, a:] m M"
proof (rule disjE)
assume pnz: "poly [:c, b, a:] m \<noteq> 0"
then have nz: "[:c, b, a:] \<noteq> 0" by auto
show ?thesis
unfolding count_roots_between_correct card_iff_exists
apply safe
apply (rule finite_subset[where B="{x. poly [:c, b, a:] x = 0}"])
apply force
apply (rule poly_roots_finite)
apply (rule nz)
using pnz
apply (auto simp add: count_roots_between_correct quadroot_in_def root_in_def card_iff_exists
algebra_simps power2_eq_square)
apply (case_tac "x = m")
apply (force simp: algebra_simps)
apply force
done
qed (auto simp: quadroot_in_def count_roots_between_correct root_in_def card_eq_0_iff)
then show "quadroot_in m M a b c = (a = 0 \<and> b = 0 \<and> c = 0 \<and> m \<le> M \<or>
m \<le> M \<and> poly [:c, b, a:] m = 0 \<or>
0 < count_roots_between [:c, b, a:] m M)"
using H by metis
qed
lemma check_quadroot_linear:
fixes a b c :: real
assumes "a = 0"
shows "\<not> quadroot_in m M a b c \<longleftrightarrow>
((b = 0 \<and> c = 0 \<and> M < m) \<or> (b = 0 \<and> c \<noteq> 0) \<or>
(b \<noteq> 0 \<and> (let x = - c / b in m > x \<or> x > M)))"
proof -
have "quadroot_in m M a b c \<longleftrightarrow> (b = 0 \<longrightarrow> quadroot_in m M a b c) \<and> (b \<noteq> 0 \<longrightarrow> quadroot_in m M a b c)"
by auto
also have "(b = 0 \<longrightarrow> quadroot_in m M a b c) \<longleftrightarrow>
((b = 0 \<longrightarrow> c = 0 \<longrightarrow> m \<le> M) \<and> (b \<noteq> 0 \<or> c = 0))"
by (auto simp: quadroot_in_def Let_def root_in_def assms field_split_simps
intro!: bexI[where x="-c / b"])
also have "(b \<noteq> 0 \<longrightarrow> quadroot_in m M a b c) \<longleftrightarrow> (b = 0 \<or> (let x = -c / b in m \<le> x \<and> x \<le> M))"
apply (auto simp: quadroot_in_def Let_def root_in_def assms field_split_simps
intro!: bexI[where x="-c / b"])
by (metis mult.commute mult_le_cancel_left_neg add_eq_0_iff mult_le_cancel_iff2)+
finally show ?thesis
by (simp add: Let_def not_less not_le)
qed
lemma check_quadroot_nonlinear:
assumes "a \<noteq> 0"
shows "quadroot_in m M a b c =
(let D = b^2 - 4 * a * c in D \<ge> 0 \<and>
((let x = (-b + sqrt D)/(2*a) in m \<le> x \<and> x \<le> M) \<or>
(let x = (-b - sqrt D)/(2*a) in m \<le> x \<and> x \<le> M)))"
by (auto simp: quadroot_in_def Let_def root_in_def quadratic_eq_zeroes_iff[OF refl refl refl assms])
lemma ncheck_quadroot:
shows "\<not>quadroot_in m M a b c \<longleftrightarrow>
(a = 0 \<longrightarrow>\<not>quadroot_in m M a b c) \<and>
(a = 0 \<or> \<not>quadroot_in m M a b c)"
by auto
subsection \<open>Movement\<close>
locale movement =
fixes a v s0 :: real
begin
text \<open>
Function to compute the distance using the equation
\<open>s(t) = s\<^sub>0 + v\<^sub>0 * t + 1 / 2 * a\<^sub>0 * t\<^sup>2\<close>
Input parameters:
\<^item> \<open>s\<^sub>0\<close>: initial distance
\<^item> \<open>v\<^sub>0\<close>: initial velocity (positive means forward direction and the converse is true)
\<^item> \<open>a\<close>: acceleration (positive for increasing and negative for decreasing)
\<^item> \<open>t\<close>: time
For the time \<open>t < 0\<close>, we assume the output of the function is \<open>s\<^sub>0\<close>. Otherwise, the output
is calculated according to the equation above.
\<close>
subsubsection \<open>Continuous Dynamics\<close>
definition p :: "real \<Rightarrow> real" where
"p t = s0 + v * t + 1/2 * a * t\<^sup>2"
lemma p_all_zeroes:
assumes D: "D = v\<^sup>2 - 2 * a * s0"
shows "p t = 0 \<longleftrightarrow> ((a \<noteq> 0 \<and> 0 \<le> D \<and> ((t = (- v + sqrt D) / a) \<or> t = (- v - sqrt D) / a)) \<or>
(a = 0 \<and> v = 0 \<and> s0 = 0) \<or> (a = 0 \<and> v \<noteq> 0 \<and> t = (- s0 / v)))"
using quadratic_eq_zeroes_iff[OF refl refl refl, of "a / 2" t v s0]
by (auto simp: movement.p_def D power2_eq_square field_split_simps)
lemma p_zero[simp]: "p 0 = s0"
by (simp add: p_def)
lemma p_continuous[continuous_intros]: "continuous_on T p"
unfolding p_def by (auto intro!: continuous_intros)
lemma isCont_p[continuous_intros]: "isCont p x"
using p_continuous[of UNIV]
by (auto simp: continuous_on_eq_continuous_at)
definition p' :: "real \<Rightarrow> real" where
"p' t = v + a * t"
lemma p'_zero: "p' 0 = v"
by (simp add: p'_def)
lemma p_has_vector_derivative[derivative_intros]: "(p has_vector_derivative p' t) (at t within s)"
by (auto simp: p_def[abs_def] p'_def has_vector_derivative_def algebra_simps
intro!: derivative_eq_intros)
lemma p_has_real_derivative[derivative_intros]: "(p has_real_derivative p' t) (at t within s)"
using p_has_vector_derivative
by (simp add: has_field_derivative_iff_has_vector_derivative)
definition p'' :: "real \<Rightarrow> real" where
"p'' t = a"
lemma p'_has_vector_derivative[derivative_intros]: "(p' has_vector_derivative p'' t) (at t within s)"
by (auto simp: p'_def[abs_def] p''_def has_vector_derivative_def algebra_simps
intro!: derivative_eq_intros)
lemma p'_has_real_derivative[derivative_intros]: "(p' has_real_derivative p'' t) (at t within s)"
using p'_has_vector_derivative
by (simp add: has_field_derivative_iff_has_vector_derivative)
definition t_stop :: real where
"t_stop = - v / a"
lemma p'_stop_zero: "p' t_stop = (if a = 0 then v else 0)" by (auto simp: p'_def t_stop_def)
lemma p'_pos_iff: "p' x > 0 \<longleftrightarrow> (if a > 0 then x > -v / a else if a < 0 then x < -v / a else v > 0)"
by (auto simp: p'_def field_split_simps)
lemma le_t_stop_iff: "a \<noteq> 0 \<Longrightarrow> x \<le> t_stop \<longleftrightarrow> (if a < 0 then p' x \<ge> 0 else p' x \<le> 0)"
by (auto simp: p'_def field_split_simps t_stop_def)
lemma p'_continuous[continuous_intros]: "continuous_on T p'"
unfolding p'_def by (auto intro: continuous_intros)
lemma isCont_p'[continuous_intros]: "isCont p' x"
using p'_continuous[of UNIV] by (auto simp: continuous_on_eq_continuous_at)
definition p_max :: real where
"p_max = p t_stop"
lemmas p_t_stop = p_max_def[symmetric]
lemma p_max_eq: "p_max = s0 - v\<^sup>2 / a / 2"
by (auto simp: p_max_def p_def t_stop_def field_split_simps power2_eq_square)
subsubsection \<open>Hybrid Dynamics\<close>
definition s :: "real \<Rightarrow> real" where
"s t = ( if t \<le> 0 then s0
else if t \<le> t_stop then p t
else p_max)"
definition q :: "real \<Rightarrow> real" where
"q t = s0 + v * t"
definition q' :: "real \<Rightarrow> real" where
"q' t = v"
lemma init_q: "q 0 = s0" unfolding q_def by auto
lemma q_continuous[continuous_intros]: "continuous_on T q"
unfolding q_def by (auto intro: continuous_intros)
lemma isCont_q[continuous_intros]: "isCont q x"
using q_continuous[of UNIV]
by (auto simp:continuous_on_eq_continuous_at)
lemma q_has_vector_derivative[derivative_intros]: "(q has_vector_derivative q' t) (at t within u)"
by (auto simp: q_def[abs_def] q'_def has_vector_derivative_def algebra_simps
intro!: derivative_eq_intros)
lemma q_has_real_derivative[derivative_intros]: "(q has_real_derivative q' t) (at t within u)"
using q_has_vector_derivative
by (simp add:has_field_derivative_iff_has_vector_derivative)
lemma s_cond_def:
"t \<le> 0 \<Longrightarrow> s t = s0"
"0 \<le> t \<Longrightarrow> t \<le> t_stop \<Longrightarrow> s t = p t"
by (simp_all add: s_def)
end
locale braking_movement = movement +
assumes decel: "a < 0"
assumes nonneg_vel: "v \<ge> 0"
begin
lemma t_stop_nonneg: "0 \<le> t_stop"
using decel nonneg_vel
by (auto simp: t_stop_def divide_simps)
lemma t_stop_pos:
assumes "v \<noteq> 0"
shows "0 < t_stop"
using decel nonneg_vel assms
by (auto simp: t_stop_def divide_simps)
lemma t_stop_zero:
assumes "t_stop = 0"
shows "v = 0"
using assms decel
by (auto simp: t_stop_def)
lemma t_stop_zero_not_moving: "t_stop = 0 \<Longrightarrow> q t = s0"
unfolding q_def using t_stop_zero by auto
abbreviation "s_stop \<equiv> s t_stop"
lemma s_t_stop: "s_stop = p_max"
using t_stop_nonneg
by (auto simp: s_def t_stop_def p_max_def p_def)
lemma s0_le_s_stop: "s0 \<le> s_stop"
proof (rule subst[where t="s_stop" and s="p_max"])
show "p_max = s_stop" by (rule sym[OF s_t_stop])
next
show "s0 \<le> p_max"
proof (rule subst[where t="p_max" and s="s0 - v\<^sup>2 / a / 2"])
show " s0 - v\<^sup>2 / a / 2 = p_max" using p_max_eq by auto
next
have "0 \<le> - v\<^sup>2 / a / 2" using decel zero_le_square[of v]
proof -
have f1: "a \<le> 0"
using \<open>a < 0\<close> by linarith
have "(- 1 * v\<^sup>2 \<le> 0) = (0 \<le> v\<^sup>2)"
by auto
then have "0 \<le> - 1 * v\<^sup>2 / a"
using f1 by (meson zero_le_divide_iff zero_le_power2)
then show ?thesis
by force
qed
thus "s0 \<le> s0 - v\<^sup>2 / a / 2" by auto
qed
qed
lemma p_mono: "x \<le> y \<Longrightarrow> y \<le> t_stop \<Longrightarrow> p x \<le> p y"
using decel
proof -
assume "x \<le> y" and "y \<le> t_stop" and "a < 0"
hence "x + y \<le> - 2 * v / a"
unfolding t_stop_def by auto
hence "-1 / 2 * a * (x + y) \<le> v" (is "?lhs0 \<le> ?rhs0")
- using `a < 0` by (auto simp add: field_simps)
+ using \<open>a < 0\<close> by (auto simp add: field_simps)
hence "?lhs0 * (x- y) \<ge> ?rhs0 * (x - y)"
- using `x \<le> y` by sos
+ using \<open>x \<le> y\<close> by sos
hence "v * x + 1 / 2 * a * x\<^sup>2 \<le> v * y + 1 / 2 * a * y\<^sup>2"
by (auto simp add: field_simps power_def)
thus " p x \<le> p y"
unfolding p_max_def p_def t_stop_def by auto
qed
lemma p_antimono: "x \<le> y \<Longrightarrow> t_stop \<le> x \<Longrightarrow> p y \<le> p x"
using decel
proof -
assume "x \<le> y" and "t_stop \<le> x" and "a < 0"
hence "- 2 * v / a \<le> x + y"
unfolding t_stop_def by auto
hence "v \<le> - 1/ 2 * a * (x + y)"
- using `a < 0` by (auto simp add: field_simps)
+ using \<open>a < 0\<close> by (auto simp add: field_simps)
hence "v * (x - y) \<ge> -1/2 * a * (x\<^sup>2 - y\<^sup>2) "
- using `x \<le> y` by sos
+ using \<open>x \<le> y\<close> by sos
hence "v * y + 1/2 * a * y\<^sup>2 \<le> v * x + 1/2 * a * x\<^sup>2"
by (auto simp add: field_simps)
thus "p y \<le> p x"
unfolding p_max_def p_def t_stop_def by auto
qed
lemma p_strict_mono: "x < y \<Longrightarrow> y \<le> t_stop \<Longrightarrow> p x < p y"
using decel
proof -
assume "x < y" and "y \<le> t_stop" and "a < 0"
hence "x + y < - 2 * v / a"
unfolding t_stop_def by auto
hence "-1 / 2 * a * (x + y) < v" (is "?lhs0 < ?rhs0")
- using `a < 0` by (auto simp add: field_simps)
+ using \<open>a < 0\<close> by (auto simp add: field_simps)
hence "?lhs0 * (x- y) > ?rhs0 * (x - y)"
- using `x < y` by sos
+ using \<open>x < y\<close> by sos
hence "v * x + 1 / 2 * a * x\<^sup>2 < v * y + 1 / 2 * a * y\<^sup>2"
by (auto simp add: field_simps power_def)
thus " p x < p y"
unfolding p_max_def p_def t_stop_def by auto
qed
lemma p_strict_antimono: "x < y \<Longrightarrow> t_stop \<le> x\<Longrightarrow> p y < p x"
using decel
proof -
assume "x < y" and "t_stop \<le> x" and "a < 0"
hence "- 2 * v / a < x + y"
unfolding t_stop_def by auto
hence "v < - 1/ 2 * a * (x + y)"
- using `a < 0` by (auto simp add: field_simps)
+ using \<open>a < 0\<close> by (auto simp add: field_simps)
hence "v * (x - y) > -1/2 * a * (x\<^sup>2 - y\<^sup>2) "
- using `x < y` by sos
+ using \<open>x < y\<close> by sos
hence "v * y + 1/2 * a * y\<^sup>2 < v * x + 1/2 * a * x\<^sup>2"
by (auto simp add: field_simps)
thus "p y < p x"
unfolding p_max_def p_def t_stop_def by auto
qed
lemma p_max: "p x \<le> p_max"
unfolding p_max_def
by (cases "x \<le> t_stop") (auto intro: p_mono p_antimono)
lemma continuous_on_s[continuous_intros]: "continuous_on T s"
unfolding s_def[abs_def]
using t_stop_nonneg
by (intro continuous_on_subset[where t=T and s = "{.. 0}\<union>({0 .. t_stop} \<union> {t_stop ..})"] continuous_on_If)
(auto simp: p_continuous p_max_def antisym_conv[where x=0])
lemma isCont_s[continuous_intros]: "isCont s x"
using continuous_on_s[of UNIV]
by (auto simp: continuous_on_eq_continuous_at)
definition s' :: "real \<Rightarrow> real" where
"s' t = (if t \<le> t_stop then p' t else 0)"
lemma s_has_real_derivative:
assumes "t \<ge> 0" "v / a \<le> 0" "a \<noteq> 0"
shows "(s has_real_derivative s' t) (at t within {0..})"
proof -
from assms have *: "t \<le> t_stop \<longleftrightarrow> t \<in> {0 .. t_stop}" by simp
from assms have "0 \<le> t_stop" by (auto simp: t_stop_def)
have "((\<lambda>t. if t \<in> {0 .. t_stop} then p t else p_max) has_real_derivative
(if t \<in> {0..t_stop} then p' t else 0)) (at t within {0..})"
unfolding s_def[abs_def] s'_def
has_field_derivative_iff_has_vector_derivative
apply (rule has_vector_derivative_If_within_closures[where T = "{t_stop ..}"])
using \<open>0 \<le> t_stop\<close> \<open>a \<noteq> 0\<close>
by (auto simp: assms p'_stop_zero p_t_stop max_def insert_absorb
intro!: p_has_vector_derivative)
from _ _ this show ?thesis
unfolding has_vector_derivative_def has_field_derivative_iff_has_vector_derivative
s'_def s_def[abs_def] *
by (rule has_derivative_transform)
(auto simp: assms s_def p_max_def t_stop_def)
qed
lemma s_has_vector_derivative[derivative_intros]:
assumes "t \<ge> 0" "v / a \<le> 0" "a \<noteq> 0"
shows "(s has_vector_derivative s' t) (at t within {0..})"
using s_has_real_derivative[OF assms]
by (simp add: has_field_derivative_iff_has_vector_derivative)
lemma s_has_field_derivative[derivative_intros]:
assumes "t \<ge> 0" "v / a \<le> 0" "a \<noteq> 0"
shows "(s has_field_derivative s' t) (at t within {0..})"
using s_has_vector_derivative[OF assms]
by(simp add: has_field_derivative_iff_has_vector_derivative)
lemma s_has_real_derivative_at:
assumes "0 < x" "0 \<le> v" "a < 0"
shows "(s has_real_derivative s' x) (at x)"
proof -
from assms have "(s has_real_derivative s' x) (at x within {0 ..})"
by (intro s_has_real_derivative) (auto intro!: divide_nonneg_nonpos)
then have "(s has_real_derivative s' x) (at x within {0<..})"
by (rule DERIV_subset) auto
then show "(s has_real_derivative s' x) (at x)" using assms
by (subst (asm) at_within_open) auto
qed
lemma s_delayed_has_field_derivative[derivative_intros]:
assumes "\<delta> < t" "0 \<le> v" "a < 0"
shows "((\<lambda>x. s (x - \<delta>)) has_field_derivative s' (t - \<delta>)) (at t within {\<delta><..})"
proof -
from assms have "((\<lambda>x. s (x + - \<delta>)) has_real_derivative s' (t - \<delta>)) (at t)"
using DERIV_shift[of "s" "(s' (t - \<delta>))" t "-\<delta>"] s_has_real_derivative_at
by auto
thus "((\<lambda>x. s (x - \<delta>)) has_field_derivative s' (t - \<delta>)) (at t within {\<delta><..})"
using has_field_derivative_at_within by auto
qed
lemma s_delayed_has_vector_derivative[derivative_intros]:
assumes "\<delta> < t" "0 \<le> v" "a < 0"
shows "((\<lambda>x. s (x - \<delta>)) has_vector_derivative s' (t - \<delta>)) (at t within {\<delta><..})"
using s_delayed_has_field_derivative[OF assms]
by(simp add:has_field_derivative_iff_has_vector_derivative)
lemma s'_nonneg: "0 \<le> v \<Longrightarrow> a \<le> 0 \<Longrightarrow> 0 \<le> s' x"
by (auto simp: s'_def p'_def t_stop_def field_split_simps)
lemma s'_pos: "0 \<le> x \<Longrightarrow> x < t_stop \<Longrightarrow> 0 \<le> v \<Longrightarrow> a \<le> 0 \<Longrightarrow> 0 < s' x"
by (intro le_neq_trans s'_nonneg)
(auto simp: s'_def p'_def t_stop_def field_split_simps)
subsubsection \<open>Monotonicity of Movement\<close>
lemma s_mono:
assumes "t \<ge> u" "u \<ge> 0"
shows "s t \<ge> s u"
using p_mono[of u t] assms p_max[of u] by (auto simp: s_def)
lemma s_strict_mono:
assumes "u < t" "t \<le> t_stop" "u \<ge> 0"
shows "s u < s t"
using p_strict_mono[of u t] assms p_max[of u] by (auto simp: s_def)
lemma s_antimono:
assumes "x \<le> y"
assumes "t_stop \<le> x"
shows "s y \<le> s x"
proof -
from assms have "t_stop \<le> y" by auto
hence "s y \<le> p_max" unfolding s_def p_max_eq
using p_max_def p_max_eq s0_le_s_stop s_t_stop by auto
also have "... \<le> s x"
using \<open>t_stop \<le> x\<close> s_mono s_t_stop t_stop_nonneg by fastforce
ultimately show "s y \<le> s x" by auto
qed
lemma init_s : "t \<le> 0 \<Longrightarrow> s t = s0"
using decel nonneg_vel by (simp add: divide_simps s_def)
lemma q_min: "0 \<le> t \<Longrightarrow> s0 \<le> q t"
unfolding q_def using nonneg_vel by auto
lemma q_mono: "x \<le> y \<Longrightarrow> q x \<le> q y"
unfolding q_def using nonneg_vel by (auto simp: mult_left_mono)
subsubsection \<open>Maximum at Stopping Time\<close>
lemma s_max: "s x \<le> s_stop"
using p_max[of x] p_max[of 0] unfolding s_t_stop by (auto simp: s_def)
lemma s_eq_s_stop: "NO_MATCH t_stop x \<Longrightarrow> x \<ge> t_stop \<Longrightarrow> s x = s_stop"
using t_stop_nonneg by (auto simp: s_def p_max_def)
end
subsection \<open>Safe Distance\<close>
locale safe_distance =
fixes a\<^sub>e v\<^sub>e s\<^sub>e :: real
fixes a\<^sub>o v\<^sub>o s\<^sub>o :: real
assumes nonneg_vel_ego : "0 \<le> v\<^sub>e"
assumes nonneg_vel_other : "0 \<le> v\<^sub>o"
assumes decelerate_ego : "a\<^sub>e < 0"
assumes decelerate_other : "a\<^sub>o < 0"
assumes in_front : "s\<^sub>e < s\<^sub>o"
begin
lemmas hyps =
nonneg_vel_ego
nonneg_vel_other
decelerate_ego
decelerate_other
in_front
sublocale ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; rule hyps)
sublocale other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; rule hyps)
sublocale ego_other: movement "a\<^sub>o - a\<^sub>e" "v\<^sub>o - v\<^sub>e" "s\<^sub>o - s\<^sub>e" by unfold_locales
subsubsection \<open>Collision\<close>
definition collision :: "real set \<Rightarrow> bool" where
"collision time_set \<equiv> (\<exists>t \<in> time_set. ego.s t = other.s t )"
abbreviation no_collision :: "real set \<Rightarrow> bool" where
"no_collision time_set \<equiv> \<not> collision time_set"
lemma no_collision_initially : "no_collision {.. 0}"
using decelerate_ego nonneg_vel_ego
using decelerate_other nonneg_vel_other in_front
by (auto simp: divide_simps collision_def ego.s_def other.s_def)
lemma no_collisionI:
"(\<And>t. t \<in> S \<Longrightarrow> ego.s t \<noteq> other.s t) \<Longrightarrow> no_collision S"
unfolding collision_def by blast
theorem cond_1: "ego.s_stop < s\<^sub>o \<Longrightarrow> no_collision {0..}"
proof (rule no_collisionI, simp)
fix t::real
assume "t \<ge> 0"
have "ego.s t \<le> ego.s_stop"
by (rule ego.s_max)
also assume "\<dots> < s\<^sub>o"
also have "\<dots> = other.s 0"
by (simp add: other.init_s)
also have "\<dots> \<le> other.s t"
using \<open>0 \<le> t\<close> hyps
by (intro other.s_mono) auto
finally show "ego.s t \<noteq> other.s t"
by simp
qed
lemma ego_other_strict_ivt:
assumes "ego.s t > other.s t"
shows "collision {0 ..< t}"
proof cases
have [simp]: "s\<^sub>e < s\<^sub>o \<Longrightarrow> ego.s 0 \<le> other.s 0"
by (simp add: movement.s_def)
assume "0 \<le> t"
with assms in_front
have "\<exists>x\<ge>0. x \<le> t \<and> other.s x - ego.s x = 0"
by (intro IVT2, auto simp: continuous_diff other.isCont_s ego.isCont_s)
then show ?thesis
using assms
by (auto simp add: algebra_simps collision_def Bex_def order.order_iff_strict)
qed (insert assms hyps, auto simp: collision_def ego.init_s other.init_s intro!: bexI[where x=0])
lemma collision_subset: "collision s \<Longrightarrow> s \<subseteq> t \<Longrightarrow> collision t"
by (auto simp: collision_def)
lemma ego_other_ivt:
assumes "ego.s t \<ge> other.s t"
shows "collision {0 .. t}"
proof cases
assume "ego.s t > other.s t"
from ego_other_strict_ivt[OF this]
show ?thesis
by (rule collision_subset) auto
qed (insert hyps assms; cases "t \<ge> 0"; force simp: collision_def ego.init_s other.init_s)
theorem cond_2:
assumes "ego.s_stop \<ge> other.s_stop"
shows "collision {0 ..}"
using assms
apply (intro collision_subset[where t="{0 ..}" and s = "{0 .. max ego.t_stop other.t_stop}"])
apply (intro ego_other_ivt[where t = "max ego.t_stop other.t_stop"])
apply (auto simp: ego.s_eq_s_stop other.s_eq_s_stop)
done
abbreviation D2 :: "real" where
"D2 \<equiv> (v\<^sub>o - v\<^sub>e)^2 - 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e)"
abbreviation t\<^sub>D' :: "real" where
"t\<^sub>D' \<equiv> sqrt (2 * (ego.s_stop - other.s_stop) / a\<^sub>o)"
lemma pos_via_half_dist:
"dist a b < b / 2 \<Longrightarrow> b > 0 \<Longrightarrow> a > 0"
by (auto simp: dist_real_def abs_real_def split: if_splits)
lemma collision_within_p:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "collision {0..} \<longleftrightarrow> (\<exists>t\<ge>0. ego.p t = other.p t \<and> t < ego.t_stop \<and> t < other.t_stop)"
proof (auto simp: collision_def, goal_cases)
case (2 t)
then show ?case
by (intro bexI[where x = t]) (auto simp: ego.s_def other.s_def)
next
case (1 t)
then show ?case using assms hyps ego.t_stop_nonneg other.t_stop_nonneg
apply (auto simp: ego.s_def other.s_def ego.s_t_stop other.s_t_stop ego.p_t_stop other.p_t_stop not_le
split: if_splits)
defer
proof goal_cases
case 1
from 1 have le: "ego.t_stop \<le> other.t_stop" by auto
from 1 have "ego.t_stop < t" by simp
from other.s_strict_mono[OF this] 1
have "other.s ego.t_stop < other.s t"
by auto
also have "\<dots> = ego.s ego.t_stop"
using ego.s_t_stop ego.t_stop_nonneg 1 other.s_def by auto
finally have "other.s ego.t_stop < ego.s ego.t_stop" .
from ego_other_strict_ivt[OF this] le in_front
show ?case
by (auto simp add: collision_def) (auto simp: movement.s_def split: if_splits)
next
case 2
from 2 have "other.p_max = ego.p t" by simp
also have "\<dots> \<le> ego.p ego.t_stop"
using 2
by (intro ego.p_mono) auto
also have "\<dots> = ego.p_max"
by (simp add: ego.p_t_stop)
also note \<open>\<dots> < other.p_max\<close>
finally show ?case by arith
next
case 3
thus "\<exists>t\<ge>0. ego.p t = other.p t \<and> t < ego.t_stop \<and> t < other.t_stop"
apply (cases "t = other.t_stop")
apply (simp add: other.p_t_stop )
apply (metis (no_types) ego.p_max not_le)
apply (cases "t = ego.t_stop")
apply (simp add: ego.p_t_stop)
defer
apply force
proof goal_cases
case (1)
let ?d = "\<lambda>t. other.p' t - ego.p' t"
define d' where "d' = ?d ego.t_stop / 2"
have d_cont: "isCont ?d ego.t_stop"
unfolding ego.t_stop_def other.p'_def ego.p'_def by simp
have "?d ego.t_stop > 0"
using 1
by (simp add: ego.p'_stop_zero other.p'_pos_iff) (simp add: ego.t_stop_def other.t_stop_def)
then have "d' > 0" by (auto simp: d'_def)
from d_cont[unfolded continuous_at_eps_delta, THEN spec, rule_format, OF \<open>d' > 0\<close>]
obtain e where e: "e > 0" "\<And>x. dist x ego.t_stop < e \<Longrightarrow> ?d x > 0"
unfolding d'_def
using \<open>?d ego.t_stop > 0\<close> pos_via_half_dist
by force
define t' where "t' = ego.t_stop - min (ego.t_stop / 2) (e / 2)"
have "0 < ego.t_stop" using 1 by auto
have "other.p t' - ego.p t' < other.p ego.t_stop - ego.p ego.t_stop"
apply (rule DERIV_pos_imp_increasing[of t'])
apply (force simp: t'_def e min_def \<open>0 < ego.t_stop\<close>)
apply (auto intro!: exI[where x = "?d x" for x] intro!: derivative_intros e)
using \<open>e > 0\<close>
apply (auto simp: t'_def dist_real_def algebra_simps)
done
also have "\<dots> = 0" using 1 by (simp add: ego.p_t_stop)
finally have less: "other.p t' < ego.p t'" by simp
have "t' > 0"
using 1 by (auto simp: t'_def algebra_simps min_def)
have "t' < ego.t_stop" by (auto simp: t'_def \<open>e > 0\<close> \<open>ego.t_stop > 0\<close>)
from less_le_trans[OF \<open>t' < ego.t_stop\<close> \<open>ego.t_stop \<le> other.t_stop\<close>]
have "t' < other.t_stop" .
from ego_other_strict_ivt[of t'] less
have "collision {0..<t'}"
using \<open>t' > 0\<close> \<open>t' < ego.t_stop\<close> \<open>t' < other.t_stop\<close>
by (auto simp: other.s_def ego.s_def split: if_splits)
thus ?case
using \<open>t' > 0\<close> \<open>t' < ego.t_stop\<close> \<open>t' < other.t_stop\<close>
apply (auto simp: collision_def ego.s_def other.s_def movement.p_def
split: if_splits)
apply (rule_tac x = t in exI)
apply (auto simp: movement.p_def)[]
done
qed
qed
qed
lemma collision_within_eq:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "collision {0..} \<longleftrightarrow> collision {0 ..< min ego.t_stop other.t_stop}"
unfolding collision_within_p[OF assms]
unfolding collision_def
by (safe; force
simp: ego.s_def other.s_def movement.p_def ego.t_stop_def other.t_stop_def
split: if_splits)
lemma collision_excluded: "(\<And>t. t \<in> T \<Longrightarrow> ego.s t \<noteq> other.s t) \<Longrightarrow> collision S \<longleftrightarrow> collision (S - T)"
by (auto simp: collision_def)
lemma collision_within_less:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "collision {0..} \<longleftrightarrow> collision {0 <..< min ego.t_stop other.t_stop}"
proof -
note collision_within_eq[OF assms]
also have "collision {0 ..< min ego.t_stop other.t_stop} \<longleftrightarrow>
collision ({0 ..< min ego.t_stop other.t_stop} - {0})"
using hyps assms
by (intro collision_excluded) (auto simp: ego.s_def other.s_def)
also have "{0 ..< min ego.t_stop other.t_stop} - {0} = {0 <..< min ego.t_stop other.t_stop}"
by auto
finally show ?thesis
unfolding collision_def
by (safe;
force
simp: ego.s_def other.s_def movement.p_def ego.t_stop_def other.t_stop_def
split: if_splits)
qed
theorem cond_3:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "collision {0..} \<longleftrightarrow> (a\<^sub>o > a\<^sub>e \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)"
proof -
have "v\<^sub>o \<noteq> 0"
using assms(1) assms(2) movement.s_def movement.t_stop_def by auto
with hyps have "v\<^sub>o > 0" by auto
note hyps = hyps this
define t1 where "t1 = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)"
define t2 where "t2 = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)"
define bounded where "bounded \<equiv> \<lambda>t. (0 \<le> t \<and> t \<le> ego.t_stop \<and> t \<le> other.t_stop)"
have ego_other_conv:
"\<And>t. bounded t \<Longrightarrow> ego.p t = other.p t \<longleftrightarrow> ego_other.p t = 0"
by (auto simp: movement.p_def field_split_simps)
let ?r = "{0 <..< min ego.t_stop other.t_stop}"
have D2: "D2 = (v\<^sub>o - v\<^sub>e)\<^sup>2 - 4 * ((a\<^sub>o - a\<^sub>e) / 2) * (s\<^sub>o - s\<^sub>e)" by simp
define D where "D = D2"
note D = D_def[symmetric]
define x1 where "x1 \<equiv> (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (2 * ((a\<^sub>o - a\<^sub>e) / 2))"
define x2 where "x2 \<equiv> (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (2 * ((a\<^sub>o - a\<^sub>e) / 2))"
have x2: "x2 =(- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)"
by (simp add: x2_def field_split_simps)
have x1: "x1 =(- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)"
by (simp add: x1_def field_split_simps)
from collision_within_less[OF assms]
have coll_eq: "collision {0..} = collision ?r"
by (auto simp add: bounded_def)
also have "\<dots> \<longleftrightarrow> (a\<^sub>o > a\<^sub>e \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)"
proof safe
assume H: "a\<^sub>e < a\<^sub>o" "v\<^sub>o < v\<^sub>e" "0 \<le> D2"
assume sqrt: "sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o"
have nz: "(a\<^sub>o - a\<^sub>e) / 2 \<noteq> 0" using \<open>a\<^sub>e < a\<^sub>o\<close> by simp
note sol = quadratic_eq_zeroes_iff[OF D2 x1_def[THEN meta_eq_to_obj_eq] x2_def[THEN meta_eq_to_obj_eq] nz]
from sol[of x2] \<open>0 \<le> D2\<close>
have "other.p x2 = ego.p x2"
by (auto simp: ego.p_def other.p_def field_split_simps)
moreover
have "x2 > 0"
proof (rule ccontr)
assume "\<not> 0 < x2"
then have "ego_other.p x2 \<ge> ego_other.p 0"
using H hyps
by (intro DERIV_nonpos_imp_nonincreasing[of x2])
(auto intro!: exI[where x="ego_other.p' x" for x] derivative_eq_intros
simp: ego_other.p'_def add_nonpos_nonpos mult_nonneg_nonpos)
also have "ego_other.p 0 > 0" using hyps by (simp add: ego_other.p_def)
finally (xtrans) show False using \<open>other.p x2 = ego.p x2\<close>
by (simp add: movement.p_def field_split_simps power2_eq_square)
qed
moreover
have "x2 < other.t_stop"
using sqrt H hyps
by (auto simp: x2 other.t_stop_def field_split_simps power2_eq_square)
ultimately
show "collision {0<..<min ego.t_stop other.t_stop}"
proof (cases "x2 < ego.t_stop", goal_cases)
case 2
then have "other.s x2 = other.p x2"
by (auto simp: other.s_def)
also from 2 have "\<dots> \<le> ego.p ego.t_stop"
by (auto intro!: ego.p_antimono)
also have "\<dots> = ego.s x2"
using 2 by (auto simp: ego.s_def ego.p_t_stop)
finally have "other.s x2 \<le> ego.s x2" .
from ego_other_ivt[OF this]
show ?thesis
unfolding coll_eq[symmetric]
by (rule collision_subset) auto
qed (auto simp: collision_def ego.s_def other.s_def not_le intro!: bexI[where x=x2])
next
let ?max = "max ego.t_stop other.t_stop"
let ?min = "min ego.t_stop other.t_stop"
assume "collision ?r"
then obtain t where t: "ego.p t = other.p t" "0 < t" "t < ?min"
by (auto simp: collision_def ego.s_def other.s_def)
then have "t < - (v\<^sub>e / a\<^sub>e)" "t < - (v\<^sub>o / a\<^sub>o)" "t < other.t_stop"
by (simp_all add: ego.t_stop_def other.t_stop_def)
from t have "ego_other.p t = 0"
by (auto simp: movement.p_def field_split_simps)
from t have "t < ?max" by auto
from hyps assms have "0 < ego_other.p 0"
by simp
from ego_other.p_def[abs_def, THEN meta_eq_to_obj_eq]
have eop_eq: "ego_other.p = (\<lambda>t. 1 / 2 * (a\<^sub>o - a\<^sub>e) * t\<^sup>2 + (v\<^sub>o - v\<^sub>e) * t + (s\<^sub>o - s\<^sub>e))"
by (simp add: algebra_simps)
show "a\<^sub>o > a\<^sub>e"
proof -
have "ego.p other.t_stop \<le> ego.p_max"
by (rule ego.p_max)
also have "... \<le> other.p other.t_stop" using hyps assms
by (auto simp:other.s_def ego.s_def ego.p_t_stop split:if_splits)
finally have "0 \<le> ego_other.p other.t_stop"
by (auto simp add:movement.p_def field_simps)
from p_convex[OF eop_eq, of 0 t other.t_stop, simplified \<open>ego_other.p t = 0\<close>,
OF \<open>0 < t\<close> \<open>t < other.t_stop\<close> \<open>0 < ego_other.p 0\<close> \<open>0 \<le> ego_other.p other.t_stop\<close>]
show "a\<^sub>o > a\<^sub>e" by (simp add: algebra_simps)
qed
have rewr: "4 * ((a\<^sub>o - a\<^sub>e) / 2) = 2 * (a\<^sub>o - a\<^sub>e)" by simp
from \<open>a\<^sub>o > a\<^sub>e\<close> \<open>ego_other.p t = 0\<close> ego_other.p_all_zeroes[OF D2[symmetric], of t]
have "0 \<le> D2" and disj: "(t = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e) \<or> t = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e))"
using hyps assms
unfolding rewr by simp_all
show "0 \<le> D2" by fact
- from add_strict_mono[OF \<open>t < - (v\<^sub>e / a\<^sub>e)\<close> \<open>t < - (v\<^sub>o / a\<^sub>o)\<close>] `0 < t` \<open>a\<^sub>o > a\<^sub>e\<close>
+ from add_strict_mono[OF \<open>t < - (v\<^sub>e / a\<^sub>e)\<close> \<open>t < - (v\<^sub>o / a\<^sub>o)\<close>] \<open>0 < t\<close> \<open>a\<^sub>o > a\<^sub>e\<close>
have "0 < - (v\<^sub>e / a\<^sub>e) + - (v\<^sub>o / a\<^sub>o)" by (simp add: divide_simps)
then have "0 > v\<^sub>e * a\<^sub>o + a\<^sub>e * v\<^sub>o" using hyps
by (simp add: field_split_simps split: if_splits)
show "v\<^sub>o < v\<^sub>e"
- using `a\<^sub>e < a\<^sub>o` `movement.p (a\<^sub>o - a\<^sub>e) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e) t = 0` in_front t(2)
+ using \<open>a\<^sub>e < a\<^sub>o\<close> \<open>movement.p (a\<^sub>o - a\<^sub>e) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e) t = 0\<close> in_front t(2)
apply (auto simp: movement.p_def divide_less_0_iff algebra_simps power2_eq_square)
by (smt divide_less_0_iff mult_le_cancel_right mult_mono mult_nonneg_nonneg nonneg_vel_ego)
from disj have "x2 < ?min"
proof rule
assume "t = (- (v\<^sub>o - v\<^sub>e) - sqrt D2) / (a\<^sub>o - a\<^sub>e)"
then show ?thesis
using \<open>t < ?min\<close>
by (simp add: x2)
next
assume "t = (- (v\<^sub>o - v\<^sub>e) + sqrt D2) / (a\<^sub>o - a\<^sub>e)"
also have "\<dots> \<ge> x2"
unfolding x2
apply (rule divide_right_mono)
apply (subst (2) diff_conv_add_uminus)
apply (rule add_left_mono)
using \<open>a\<^sub>o > a\<^sub>e\<close> \<open>D2 \<ge> 0\<close>
by auto
also (xtrans) note \<open>t < ?min\<close>
finally show ?thesis .
qed
then show "sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o"
using hyps \<open>a\<^sub>o > a\<^sub>e\<close>
by (auto simp: x2 field_split_simps other.t_stop_def)
qed
finally show ?thesis .
qed
subsubsection \<open>Formalising Safe Distance\<close>
text \<open>First definition for Safe Distance based on \<open>cond_1\<close>.\<close>
definition absolute_safe_distance :: real where
"absolute_safe_distance = - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)"
lemma absolute_safe_distance:
assumes "s\<^sub>o - s\<^sub>e > absolute_safe_distance"
shows "no_collision {0..}"
proof -
from assms hyps absolute_safe_distance_def have "ego.s_stop < s\<^sub>o"
by (auto simp add:ego.s_def ego.p_def ego.t_stop_def power_def)
thus ?thesis by (rule cond_1)
qed
text \<open>First Fallback for Safe Distance.\<close>
definition fst_safe_distance :: real where
"fst_safe_distance = v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)"
definition distance_leq_d2 :: real where
"distance_leq_d2 = (a\<^sub>e + a\<^sub>o) / (2 * a\<^sub>o\<^sup>2) * v\<^sub>o\<^sup>2 - v\<^sub>o * v\<^sub>e / a\<^sub>o"
lemma snd_leq_fst_exp: "distance_leq_d2 \<le> fst_safe_distance"
proof -
have "0 \<le> (other.t_stop - ego.t_stop)\<^sup>2" by auto
hence "- ego.t_stop\<^sup>2 \<le> other.t_stop\<^sup>2 - 2 * other.t_stop * ego.t_stop" by (simp add:power_def algebra_simps)
with hyps(3) have "- ego.t_stop\<^sup>2 * (a\<^sub>e / 2) \<ge> (other.t_stop\<^sup>2 - 2 * other.t_stop * ego.t_stop) * (a\<^sub>e / 2)"
by (smt half_gt_zero_iff mult_le_cancel_right)
with ego.t_stop_def other.t_stop_def hyps
have "- v\<^sub>e\<^sup>2 / (2 * a\<^sub>e) \<ge> a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o" by (simp add:power_def algebra_simps)
with fst_safe_distance_def distance_leq_d2_def
have 1: "fst_safe_distance \<ge> a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + v\<^sub>o\<^sup>2 / (2 * a\<^sub>o)" by (auto simp add:algebra_simps)
have "a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) = distance_leq_d2" (is "?LHS = _")
proof -
have "?LHS = a\<^sub>e * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2) - v\<^sub>o * v\<^sub>e / a\<^sub>o + a\<^sub>o * v\<^sub>o\<^sup>2 / (2 * a\<^sub>o\<^sup>2)"
by (auto simp add: algebra_simps power_def)
also have "... = distance_leq_d2"
by (auto simp add: power_def field_split_simps distance_leq_d2_def)
finally show ?thesis by auto
qed
with 1 show ?thesis by auto
qed
lemma sqrt_D2_leq_stop_time_diff:
assumes "a\<^sub>e < a\<^sub>o"
assumes "0 \<le> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o "
assumes "s\<^sub>o - s\<^sub>e \<ge> distance_leq_d2"
shows "sqrt D2 \<le> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o"
proof -
from assms have "- 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e) \<le> - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2" (is "?L \<le> ?R")
by simp
hence "D2 \<le> (v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2" by (simp add: algebra_simps)
also have "... = (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2"
proof -
from distance_leq_d2_def
have 1: "(v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * distance_leq_d2 =
(v\<^sub>o - v\<^sub>e)\<^sup>2 - (a\<^sub>o - a\<^sub>e) * (a\<^sub>e + a\<^sub>o) / a\<^sub>o\<^sup>2 * v\<^sub>o\<^sup>2 + 2 * (a\<^sub>o - a\<^sub>e) * v\<^sub>o * v\<^sub>e / a\<^sub>o"
by (auto simp add: field_split_simps)
with hyps(4) have "... = (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2"
by (auto simp add: power_def field_split_simps)
with 1 show ?thesis by auto
qed
finally show ?thesis by (smt assms(2) real_le_lsqrt real_sqrt_le_0_iff)
qed
lemma cond2_imp_pos_vo:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "v\<^sub>o \<noteq> 0"
proof (rule ccontr)
assume "\<not> v\<^sub>o \<noteq> 0"
with other.s_def other.t_stop_def have "other.s_stop = s\<^sub>o" by auto
with assms(2) have "ego.s_stop < s\<^sub>o" by auto
with assms(1) show "False" by auto
qed
lemma cond2_imp_gt_fst_sd:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "fst_safe_distance < s\<^sub>o - s\<^sub>e"
proof (cases "v\<^sub>e \<noteq> 0")
case True
from fst_safe_distance_def assms ego.s_def ego.t_stop_pos[OF \<open>v\<^sub>e \<noteq> 0\<close>] ego.p_def ego.t_stop_def
other.s_def other.t_stop_pos[OF cond2_imp_pos_vo[OF assms]] other.p_def other.t_stop_def hyps
show ?thesis by (simp add: power_def algebra_simps)
next
case False
with fst_safe_distance_def have "fst_safe_distance = v\<^sub>o\<^sup>2 / (2 * a\<^sub>o)" by auto
also have "... \<le> 0" by (simp add: divide_nonneg_neg hyps)
also have "... < s\<^sub>o - s\<^sub>e" by (simp add: algebra_simps hyps)
finally show ?thesis by auto
qed
text \<open>Second Fallback for Safe Distance.\<close>
definition snd_safe_distance :: real where
"snd_safe_distance = (v\<^sub>o - v\<^sub>e)\<^sup>2 / (2 * (a\<^sub>o - a\<^sub>e))"
lemma fst_leq_snd_safe_distance:
assumes "a\<^sub>e < a\<^sub>o"
shows"fst_safe_distance \<le> snd_safe_distance"
proof -
have "0 \<le> (v\<^sub>o / a\<^sub>o - v\<^sub>e / a\<^sub>e)\<^sup>2" by auto
hence 1: "0 \<le> (v\<^sub>o / a\<^sub>o)\<^sup>2 - 2 * v\<^sub>o * v\<^sub>e / (a\<^sub>o * a\<^sub>e) + (v\<^sub>e / a\<^sub>e)\<^sup>2" by (auto simp add: power_def algebra_simps)
from hyps have "0 \<le> a\<^sub>o * a\<^sub>e" by (simp add: mult_nonpos_nonpos)
from mult_right_mono[OF 1 this] hyps
have "0 \<le> v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o - 2 * v\<^sub>o * v\<^sub>e + v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e" by (auto simp add: power_def algebra_simps)
with hyps have 2: "(v\<^sub>o\<^sup>2 / (2 * a\<^sub>o) - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)) * (2 * (a\<^sub>o - a\<^sub>e)) \<le> (v\<^sub>o - v\<^sub>e)\<^sup>2"
by (auto simp add: power_def field_split_simps)
from assms have "0 \<le> 2 * (a\<^sub>o - a\<^sub>e)" by auto
from divide_right_mono[OF 2 this] assms fst_safe_distance_def snd_safe_distance_def
show ?thesis by auto
qed
lemma snd_safe_distance_iff_nonneg_D2:
assumes "a\<^sub>e < a\<^sub>o"
shows "s\<^sub>o - s\<^sub>e \<le> snd_safe_distance \<longleftrightarrow> 0 \<le> D2"
proof -
from snd_safe_distance_def assms pos_le_divide_eq[of "2 * (a\<^sub>o - a\<^sub>e)"]
have "s\<^sub>o - s\<^sub>e \<le> snd_safe_distance \<longleftrightarrow> (s\<^sub>o - s\<^sub>e) * (2 * (a\<^sub>o - a\<^sub>e)) \<le> (v\<^sub>o - v\<^sub>e)\<^sup>2" by auto
also have "... \<longleftrightarrow> 0 \<le> D2" by (auto simp add: algebra_simps)
finally show ?thesis by auto
qed
lemma t_stop_diff_neg_means_leq_D2:
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop" "a\<^sub>e < a\<^sub>o" "0 \<le> D2"
shows "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 \<longleftrightarrow> sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o"
proof
assume only_if: "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0"
from assms have "... \<le> sqrt D2" by auto
with only_if show "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2" by linarith
next
assume if_part: "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2"
from cond2_imp_gt_fst_sd[OF assms(1) assms(2)] snd_leq_fst_exp have "distance_leq_d2 \<le> s\<^sub>o - s\<^sub>e" by auto
from if_part and sqrt_D2_leq_stop_time_diff [OF \<open>a\<^sub>e < a\<^sub>o\<close> _ \<open>distance_leq_d2 \<le> s\<^sub>o - s\<^sub>e\<close>]
show " v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0" by linarith
qed
theorem cond_3':
assumes "s\<^sub>o \<le> ego.s_stop" "ego.s_stop < other.s_stop"
shows "collision {0..} \<longleftrightarrow> (a\<^sub>o > a\<^sub>e \<and> v\<^sub>o < v\<^sub>e \<and> s\<^sub>o - s\<^sub>e \<le> snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0)"
proof (cases "a\<^sub>o \<le> a\<^sub>e \<or> v\<^sub>o \<ge> v\<^sub>e")
case True
with cond_3[OF assms] show ?thesis by auto
next
case False
from \<open>\<not> (a\<^sub>o \<le> a\<^sub>e \<or> v\<^sub>e \<le> v\<^sub>o)\<close> have "a\<^sub>o > a\<^sub>e" by auto
from \<open>\<not> (a\<^sub>o \<le> a\<^sub>e \<or> v\<^sub>e \<le> v\<^sub>o)\<close> have "v\<^sub>o < v\<^sub>e" by auto
show ?thesis
proof -
from snd_safe_distance_iff_nonneg_D2 [OF \<open>a\<^sub>o > a\<^sub>e\<close>]
have 1: "(a\<^sub>e < a\<^sub>o \<and> v\<^sub>o < v\<^sub>e \<and> s\<^sub>o - s\<^sub>e \<le> snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0) \<longleftrightarrow>
(a\<^sub>e < a\<^sub>o \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0)" by auto
from t_stop_diff_neg_means_leq_D2[OF assms \<open>a\<^sub>e < a\<^sub>o\<close>]
have "... = (a\<^sub>e < a\<^sub>o \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> sqrt D2 > v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)" by auto
with 1 cond_3[OF assms] show ?thesis by blast
qed
qed
definition d :: "real \<Rightarrow> real" where
"d t = (
if t \<le> 0 then s\<^sub>o - s\<^sub>e
else if t \<le> ego.t_stop \<and> t \<le> other.t_stop then ego_other.p t
else if ego.t_stop \<le> t \<and> t \<le> other.t_stop then other.p t - ego.s_stop
else if other.t_stop \<le> t \<and> t \<le> ego.t_stop then other.s_stop - ego.p t
else other.s_stop - ego.s_stop
)"
lemma d_diff: "d t = other.s t - ego.s t"
by (auto simp: d_def ego.s_eq_s_stop other.s_eq_s_stop ego.s_cond_def other.s_cond_def
movement.p_def field_split_simps)
lemma collision_d: "collision S \<longleftrightarrow> (\<exists>t\<in>S. d t = 0)"
by (force simp: d_diff collision_def )
lemma collision_restrict: "collision {0..} \<longleftrightarrow> collision {0..max ego.t_stop other.t_stop}"
by (auto simp: max.coboundedI1 ego.t_stop_nonneg min_def
ego.s_eq_s_stop other.s_eq_s_stop collision_def
intro!: bexI[where x = "min t (max (movement.t_stop a\<^sub>e v\<^sub>e) (movement.t_stop a\<^sub>o v\<^sub>o))" for t])
lemma collision_union: "collision (A \<union> B) \<longleftrightarrow> collision A \<or> collision B"
by (auto simp: collision_def)
lemma symbolic_checker:
"collision {0..} \<longleftrightarrow>
(quadroot_in 0 (min ego.t_stop other.t_stop) (1/2 * (a\<^sub>o - a\<^sub>e)) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e)) \<or>
(quadroot_in ego.t_stop other.t_stop (1/2 * a\<^sub>o) v\<^sub>o (s\<^sub>o - ego.s_stop)) \<or>
(quadroot_in other.t_stop ego.t_stop (1/2 * a\<^sub>e) v\<^sub>e (s\<^sub>e - other.s_stop))"
(is "_ \<longleftrightarrow> ?q1 \<or> ?q2 \<or> ?q3")
proof -
have *: "{0..max ego.t_stop other.t_stop} =
{0 .. min ego.t_stop other.t_stop} \<union> {ego.t_stop .. other.t_stop} \<union> {other.t_stop .. ego.t_stop}"
using ego.t_stop_nonneg other.t_stop_nonneg
by auto
have "collision {0..min (movement.t_stop a\<^sub>e v\<^sub>e) (movement.t_stop a\<^sub>o v\<^sub>o)} = ?q1"
by (force simp: collision_def quadroot_in_def root_in_def d_def
power2_eq_square field_split_simps movement.p_def movement.s_cond_def)
moreover
have "collision {ego.t_stop .. other.t_stop} = ?q2"
using ego.t_stop_nonneg
by (force simp: collision_def quadroot_in_def root_in_def d_def
ego.s_eq_s_stop movement.s_cond_def movement.p_def)
moreover
have "collision {other.t_stop .. ego.t_stop} = ?q3"
using other.t_stop_nonneg
by (force simp: collision_def quadroot_in_def root_in_def d_def
other.s_eq_s_stop movement.s_cond_def movement.p_def)
ultimately
show ?thesis
unfolding collision_restrict * collision_union
by auto
qed
end
subsection \<open>Checker Design\<close>
definition rel_dist_to_stop :: "real \<Rightarrow> real \<Rightarrow> real" where
"rel_dist_to_stop v a \<equiv> - v\<^sup>2 / (2 * a)"
context includes floatarith_notation begin
definition rel_dist_to_stop_expr :: "nat \<Rightarrow> nat \<Rightarrow> floatarith" where
"rel_dist_to_stop_expr v a = Mult (Minus (Power (Var v) 2)) (Inverse (Mult (Num 2) (Var a)))"
definition rel_dist_to_stop' :: "nat \<Rightarrow> float interval option \<Rightarrow> float interval option \<Rightarrow> float interval option" where
"rel_dist_to_stop' p v a = approx p (rel_dist_to_stop_expr 0 1) [v, a]"
lemma rel_dist_to_stop': "interpret_floatarith (rel_dist_to_stop_expr 0 1) [v, a] = rel_dist_to_stop v a"
by (simp add: rel_dist_to_stop_def rel_dist_to_stop_expr_def inverse_eq_divide)
definition first_safe_dist :: "real \<Rightarrow> real \<Rightarrow> real" where
"first_safe_dist v\<^sub>e a\<^sub>e \<equiv> rel_dist_to_stop v\<^sub>e a\<^sub>e"
definition second_safe_dist :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o \<equiv> rel_dist_to_stop v\<^sub>e a\<^sub>e - rel_dist_to_stop v\<^sub>o a\<^sub>o"
definition second_safe_dist_expr :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> floatarith" where
"second_safe_dist_expr ve ae vo ao = Add (rel_dist_to_stop_expr ve ae) (Minus (rel_dist_to_stop_expr vo ao))"
definition second_safe_dist' :: "nat \<Rightarrow> float interval option \<Rightarrow> float interval option
\<Rightarrow> float interval option \<Rightarrow> float interval option \<Rightarrow> float interval option" where
"second_safe_dist' p v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o = approx p (second_safe_dist_expr 0 1 2 3) [v\<^sub>e, a\<^sub>e, v\<^sub>o, a\<^sub>o]"
lemma second_safe_dist':
"interpret_floatarith (second_safe_dist_expr 0 1 2 3) [v, a, v', a'] = second_safe_dist v a v' a'"
by (simp add: second_safe_dist_def second_safe_dist_expr_def rel_dist_to_stop_def rel_dist_to_stop_expr_def inverse_eq_divide)
definition t_stop :: "real \<Rightarrow> real \<Rightarrow> real" where
"t_stop v a \<equiv> - v / a"
definition t_stop_expr :: "nat \<Rightarrow> nat \<Rightarrow> floatarith" where
"t_stop_expr v a = Minus (Mult (Var v) (Inverse (Var a)))"
end
definition s_stop :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"s_stop s v a \<equiv> s + rel_dist_to_stop v a"
definition discriminant :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv> (v\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * (a\<^sub>o - a\<^sub>e) * (s\<^sub>o - s\<^sub>e)"
definition suff_cond_safe_dist2 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv>
let D2 = discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o
in \<not> (a\<^sub>e < a\<^sub>o \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2
)"
lemma less_sqrt_iff: "y \<ge> 0 \<Longrightarrow> x < sqrt y \<longleftrightarrow> (x \<ge> 0 \<longrightarrow> x\<^sup>2 < y)"
by (smt real_le_lsqrt real_less_rsqrt real_sqrt_ge_zero)
lemma suff_cond_safe_dist2_code[code]:
"suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o =
(let D2 = discriminant s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o in
(a\<^sub>e < a\<^sub>o \<longrightarrow> v\<^sub>o < v\<^sub>e \<longrightarrow> 0 \<le> D2 \<longrightarrow> (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o \<ge> 0 \<and> (v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o)\<^sup>2 \<ge> D2)))"
using real_sqrt_ge_zero real_less_rsqrt less_sqrt_iff
by (auto simp: suff_cond_safe_dist2_def Let_def)
text \<open>
There are two expressions for safe distance. The first safe distance \<open>first_safe_dist\<close> is always valid.
Whenever the distance is bigger than \<open>first_safe_dist\<close>, it is guarantee to be collision free.
The second one is \<open>second_safe_dist\<close>. If the sufficient condition \<open>suff_cond_safe_dist2\<close> is satisfied and
the distance is bigger than \<open>second_safe_dist\<close>, it is guaranteed to be collision free.
\<close>
definition check_precond :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<longleftrightarrow> s\<^sub>o > s\<^sub>e \<and> 0 \<le> v\<^sub>e \<and> 0 \<le> v\<^sub>o \<and> a\<^sub>e < 0 \<and> a\<^sub>o < 0 "
lemma check_precond_safe_distance:
"check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o"
proof
assume "safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o"
then interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o .
show "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
by (auto simp: check_precond_def in_front nonneg_vel_ego other.nonneg_vel ego.decel other.decel)
qed (unfold_locales; auto simp: check_precond_def)
subsubsection \<open>Prescriptive Checker\<close>
definition checker :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o;
safe_dist1 = first_safe_dist v\<^sub>e a\<^sub>e;
safe_dist2 = second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o;
cond2 = suff_cond_safe_dist2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o
in precond \<and> (safe_dist1 < distance \<or> (safe_dist2 < distance \<and> distance \<le> safe_dist1 \<and> cond2))"
lemma aux_logic:
assumes "a \<Longrightarrow> b"
assumes "b \<Longrightarrow> a \<longleftrightarrow> c"
shows "a \<longleftrightarrow> b \<and> c"
using assms by blast
theorem soundness_correctness:
"checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<longleftrightarrow> check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<and> safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}"
proof (rule aux_logic, simp add: checker_def Let_def)
assume cp: "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
then have in_front': "s\<^sub>o > s\<^sub>e"
and nonneg_vel_ego: "0 \<le> v\<^sub>e"
and nonneg_vel_other: "0 \<le> v\<^sub>o"
and decelerate_ego: "a\<^sub>e < 0"
and decelerate_other: "a\<^sub>o < 0"
by (auto simp: check_precond_def)
from in_front' have in_front: "0 < s\<^sub>o - s\<^sub>e" by arith
interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact)
interpret ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; fact)
interpret other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact)
have "ego.p_max < s\<^sub>o \<or> other.p_max \<le> ego.p_max \<or> s\<^sub>o \<le> ego.p_max \<and> ego.p_max < other.p_max"
by arith
then show "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}"
proof (elim disjE)
assume "ego.p_max < s\<^sub>o"
then have "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
using \<open>a\<^sub>e < 0\<close> cp
by (simp add: checker_def Let_def first_safe_dist_def rel_dist_to_stop_def ego.p_max_def
ego.p_def ego.t_stop_def algebra_simps power2_eq_square)
moreover
have "no_collision {0..}"
using \<open>ego.p_max < s\<^sub>o\<close>
by (intro cond_1) (auto simp: ego.s_t_stop)
ultimately show ?thesis by auto
next
assume "other.p_max \<le> ego.p_max"
then have "\<not> checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
using \<open>a\<^sub>e < 0\<close> \<open>a\<^sub>o < 0\<close> other.nonneg_vel
by (auto simp add: checker_def Let_def first_safe_dist_def second_safe_dist_def
rel_dist_to_stop_def movement.p_max_def
movement.p_def movement.t_stop_def algebra_simps power2_eq_square)
(smt divide_nonneg_neg mult_nonneg_nonneg)
moreover have "collision {0..}"
using \<open>other.p_max \<le> ego.p_max\<close>
by (intro cond_2) (auto simp: other.s_t_stop ego.s_t_stop)
ultimately show ?thesis by auto
next
assume H: "s\<^sub>o \<le> ego.p_max \<and> ego.p_max < other.p_max"
then have "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = (\<not> (a\<^sub>e < a\<^sub>o \<and> v\<^sub>o < v\<^sub>e \<and> 0 \<le> D2 \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < sqrt D2))"
using \<open>a\<^sub>e < 0\<close> \<open>a\<^sub>o < 0\<close> cp
by (simp add: checker_def Let_def first_safe_dist_def rel_dist_to_stop_def ego.p_max_def
ego.p_def ego.t_stop_def algebra_simps power2_eq_square second_safe_dist_def
suff_cond_safe_dist2_def discriminant_def not_less other.p_max_def other.p_def other.t_stop_def)
also have "\<dots> = no_collision {0..}"
using H
unfolding Not_eq_iff
by (intro cond_3[symmetric]) (auto simp: ego.s_t_stop other.s_t_stop)
finally show ?thesis by auto
qed
qed
definition checker2 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o;
safe_dist1 = first_safe_dist v\<^sub>e a\<^sub>e;
safe_dist2 = second_safe_dist v\<^sub>e a\<^sub>e v\<^sub>o a\<^sub>o;
safe_dist3 = - rel_dist_to_stop (v\<^sub>o - v\<^sub>e) (a\<^sub>o - a\<^sub>e)
in
if \<not> precond then False
else if distance > safe_dist1 then True
else if a\<^sub>o > a\<^sub>e \<and> v\<^sub>o < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 then distance > safe_dist3
else distance > safe_dist2"
theorem checker_eq_checker2: "checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<longleftrightarrow> checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
proof (cases "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o")
case False
with checker_def checker2_def
show ?thesis by auto
next
case True
with check_precond_def safe_distance_def
have "safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o" by (simp add: check_precond_safe_distance)
from this interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o by auto
interpret ego: braking_movement a\<^sub>e v\<^sub>e s\<^sub>e by (unfold_locales; fact)
interpret other: braking_movement a\<^sub>o v\<^sub>o s\<^sub>o by (unfold_locales; fact)
from \<open>check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o\<close> cond_3 cond_3'[symmetric] fst_leq_snd_safe_distance
ego.s_t_stop ego.p_max_def ego.p_def ego.t_stop_def hyps other.s_t_stop other.p_max_def other.p_def
other.t_stop_def checker2_def checker_def suff_cond_safe_dist2_def fst_safe_distance_def
first_safe_dist_def snd_safe_distance_def second_safe_dist_def rel_dist_to_stop_def discriminant_def
show ?thesis
by (auto simp add:power_def Let_def split: if_splits)
qed
definition checker3 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker3 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o;
s_stop_e = s\<^sub>e + rel_dist_to_stop v\<^sub>e a\<^sub>e;
s_stop_o = s\<^sub>o + rel_dist_to_stop v\<^sub>o a\<^sub>o
in precond \<and> (s_stop_e < s\<^sub>o
\<or> (s\<^sub>o \<le> s_stop_e \<and> s_stop_e < s_stop_o \<and>
(\<not>(a\<^sub>o > a\<^sub>e \<and> v\<^sub>o < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o < 0 \<and> distance * (a\<^sub>o - a\<^sub>e) \<le> (v\<^sub>o - v\<^sub>e)\<^sup>2 / 2))))"
theorem checker2_eq_checker3:
"checker2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<longleftrightarrow> checker3 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
apply (auto simp: checker2_def checker3_def Let_def first_safe_dist_def not_less
suff_cond_safe_dist2_def second_safe_dist_def rel_dist_to_stop_def check_precond_def)
proof goal_cases
case 1
then interpret safe_distance
by unfold_locales auto
from fst_leq_snd_safe_distance 1
show ?case
by (auto simp: fst_safe_distance_def snd_safe_distance_def)
next
case 2
then interpret safe_distance
by unfold_locales auto
from fst_leq_snd_safe_distance 2
show ?case
by (auto simp: fst_safe_distance_def snd_safe_distance_def field_split_simps)
next
case 3
then interpret safe_distance
by unfold_locales auto
from fst_leq_snd_safe_distance 3
show ?case
by (auto simp: fst_safe_distance_def snd_safe_distance_def field_split_simps)
qed
subsubsection \<open>Approximate Checker\<close>
lemma checker2_def': "checker2 a b c d e f = (
let distance = d - a;
precond = check_precond a b c d e f;
safe_dist1 = first_safe_dist b c;
safe_dist2 = second_safe_dist b c e f;
C = c < f \<and> e < b \<and> b * f > c * e;
P1 = (e - b)\<^sup>2 < 2 * distance * (f - c);
P2 = - b\<^sup>2 / c + e\<^sup>2 / f < 2 * distance
in precond \<and> (safe_dist1 < distance \<or>
safe_dist1 \<ge> distance \<and> (C \<and> P1 \<or> \<not>C \<and> P2)))"
unfolding checker2_def
by (auto simp: Let_def field_split_simps check_precond_def second_safe_dist_def
rel_dist_to_stop_def)
lemma power2_less_sqrt_iff: "(x::real)\<^sup>2 < y \<longleftrightarrow> (y \<ge> 0 \<and> abs x < sqrt y)"
apply (auto simp: real_less_rsqrt abs_real_def less_sqrt_iff)
apply (meson le_less le_less_trans not_less power2_less_0)+
done
schematic_goal checker_form: "interpret_form ?x ?y \<Longrightarrow> checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
unfolding checker_eq_checker2 checker2_eq_checker3 checker3_def check_precond_def first_safe_dist_def second_safe_dist_def
suff_cond_safe_dist2_def Let_def t_stop_def s_stop_def
rel_dist_to_stop_def
discriminant_def
not_le not_less
de_Morgan_conj
de_Morgan_disj
power2_less_sqrt_iff
apply (tactic \<open>(Reification.tac @{context} @{thms interpret_form.simps interpret_floatarith.simps interpret_floatarith_divide interpret_floatarith_diff}) NONE 1\<close>)
apply assumption
done
context includes floatarith_notation begin
definition "checker' p s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = approx_form p
(Conj (Conj (Less (Var (Suc (Suc 0))) (Var (Suc (Suc (Suc 0)))))
(Conj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Var (Suc (Suc (Suc (Suc (Suc 0)))))))
(Conj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))
(Conj (Less (Var 0) (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))))
(Less (Var (Suc 0)) (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))))))))
(Disj (Less (Add (Var (Suc (Suc 0)))
(Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0)))))
(Var (Suc (Suc (Suc 0)))))
(Conj (LessEqual (Var (Suc (Suc (Suc 0))))
(Add (Var (Suc (Suc 0)))
(Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0))))))
(Conj (Less (Add (Var (Suc (Suc 0)))
(Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc 0)))))) 2)) (Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var 0)))))
(Add (Var (Suc (Suc (Suc 0))))
(Mult (Minus (Power (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) 2))
(Inverse (Mult (Var (Suc (Suc (Suc (Suc 0))))) (Var (Suc 0)))))))
(Disj (LessEqual (Var (Suc 0)) (Var 0))
(Disj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc 0)))))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))
(Disj (LessEqual (Var (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0))))))))
(Add (Var (Suc (Suc (Suc (Suc (Suc 0))))))
(Minus (Mult (Mult (Var 0) (Inverse (Var (Suc 0)))) (Var (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))))))
(Less (Mult (Power (Add (Var (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) (Minus (Var (Suc (Suc (Suc (Suc (Suc 0)))))))) 2)
(Inverse (Var (Suc (Suc (Suc (Suc 0)))))))
(Mult (Add (Var (Suc (Suc (Suc 0)))) (Minus (Var (Suc (Suc 0))))) (Add (Var (Suc 0)) (Minus (Var 0))))))))))))
([a\<^sub>e, a\<^sub>o, s\<^sub>e, s\<^sub>o, Interval' (Float 2 0) (Float 2 0), v\<^sub>e, v\<^sub>o, Interval' (Float 0 1) (Float 0 1)]) (replicate 8 0)"
lemma less_Suc_iff_disj: "i < Suc x \<longleftrightarrow> i = x \<or> i < x"
by auto
lemma checker'_soundness_correctness:
assumes "a \<in> {real_of_float al .. real_of_float au}"
assumes "b \<in> {real_of_float bl .. real_of_float bu}"
assumes "c \<in> {real_of_float cl .. real_of_float cu}"
assumes "d \<in> {real_of_float dl .. real_of_float du}"
assumes "e \<in> {real_of_float el .. real_of_float eu}"
assumes "f \<in> {real_of_float fl .. real_of_float fu}"
assumes chk: "checker' p (Interval' al au) (Interval' bl bu) (Interval' cl cu) (Interval' dl du) (Interval' el eu) (Interval' fl fu)"
shows "checker a b c d e f"
apply (rule checker_form)
apply (rule approx_form_aux)
apply (rule chk[unfolded checker'_def])
using assms(1-6)
unfolding bounded_by_def
proof (auto split: option.splits)
fix i x2
assume *: "[Interval' cl cu, Interval' fl fu, Interval' al au, Interval' dl du,
Interval' (Float 2 0) (Float 2 0), Interval' bl bu, Interval' el eu, Interval' 0 0] ! i = Some x2"
assume " i < Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))"
then consider "i = 0" | "i = 1" | "i = 2" | "i = 3" | "i = 4" | "i = 5" | "i = 6" | "i = 7"
by linarith
thus " [c, f, a, d, 2, b, e, 0] ! i \<in>\<^sub>r x2"
apply cases using assms(1-6) *
by (auto intro!: in_real_intervalI dest!: Interval'_eq_Some)
qed
lemma approximate_soundness_correctness:
assumes "a \<in> {real_of_float al .. real_of_float au}"
assumes "b \<in> {real_of_float bl .. real_of_float bu}"
assumes "c \<in> {real_of_float cl .. real_of_float cu}"
assumes "d \<in> {real_of_float dl .. real_of_float du}"
assumes "e \<in> {real_of_float el .. real_of_float eu}"
assumes "f \<in> {real_of_float fl .. real_of_float fu}"
assumes chk: "checker' p (Interval' al au) (Interval' bl bu) (Interval' cl cu) (Interval' dl du) (Interval' el eu) (Interval' fl fu)"
shows checker'_precond: "check_precond a b c d e f"
and checker'_no_collision: "safe_distance.no_collision c b a f e d {0..}"
unfolding atomize_conj
apply (subst soundness_correctness[symmetric])
using checker'_soundness_correctness[OF assms]
by (auto simp: checker_def Let_def)
subsubsection \<open>Symbolic Checker\<close>
definition symbolic_checker :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<equiv>
let e_stop = - v\<^sub>e / a\<^sub>e;
o_stop = - v\<^sub>o / a\<^sub>o
in check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<and>
(\<not>quadroot_in 0 (min e_stop o_stop) (1/2 * (a\<^sub>o - a\<^sub>e)) (v\<^sub>o - v\<^sub>e) (s\<^sub>o - s\<^sub>e) \<and>
\<not>quadroot_in e_stop o_stop (1/2 * a\<^sub>o) v\<^sub>o (s\<^sub>o - movement.p a\<^sub>e v\<^sub>e s\<^sub>e e_stop) \<and>
\<not>quadroot_in o_stop e_stop (1/2 * a\<^sub>e) v\<^sub>e (s\<^sub>e - movement.p a\<^sub>o v\<^sub>o s\<^sub>o o_stop))"
theorem symbolic_soundness_correctness:
"symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<longleftrightarrow> check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<and> safe_distance.no_collision a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o {0..}"
proof -
{
assume c: "check_precond s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o"
then interpret safe_distance a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o
by (simp add: check_precond_safe_distance)
have "symbolic_checker s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o = no_collision {0..}"
using c
unfolding symbolic_checker symbolic_checker_def ego.s_t_stop other.s_t_stop ego.p_max_def other.p_max_def
by (auto simp: Let_def movement.t_stop_def)
}
then show ?thesis
by (auto simp: symbolic_checker_def Let_def)
qed
end
end
\ No newline at end of file
diff --git a/thys/Safe_Distance/Safe_Distance_Reaction.thy b/thys/Safe_Distance/Safe_Distance_Reaction.thy
--- a/thys/Safe_Distance/Safe_Distance_Reaction.thy
+++ b/thys/Safe_Distance/Safe_Distance_Reaction.thy
@@ -1,1757 +1,1757 @@
\<^marker>\<open>creator "Albert Rizaldi" "Fabian Immler\<close>
section \<open>Safe Distance with Reaction Time\<close>
theory Safe_Distance_Reaction
imports
Safe_Distance
begin
subsection \<open>Normal Safe Distance\<close>
locale safe_distance_normal = safe_distance +
fixes \<delta> :: real
assumes pos_react: "0 < \<delta>"
begin
sublocale ego2: braking_movement a\<^sub>e v\<^sub>e "(ego.q \<delta>)" ..
lemma ego2_s_init: "ego2.s 0 = ego.q \<delta>" unfolding ego2.s_def by auto
definition \<tau> :: "real \<Rightarrow> real" where
"\<tau> t = t - \<delta>"
definition \<tau>' :: "real \<Rightarrow> real" where
"\<tau>' t = 1"
lemma \<tau>_continuous[continuous_intros]: "continuous_on T \<tau>"
unfolding \<tau>_def by (auto intro: continuous_intros)
lemma isCont_\<tau>[continuous_intros]: "isCont \<tau> x"
using \<tau>_continuous[of UNIV] by (auto simp: continuous_on_eq_continuous_at)
lemma del_has_vector_derivative[derivative_intros]: "(\<tau> has_vector_derivative \<tau>' t) (at t within u)"
by (auto simp: \<tau>_def[abs_def] \<tau>'_def has_vector_derivative_def algebra_simps
intro!: derivative_eq_intros)
lemma del_has_real_derivative[derivative_intros]: "(\<tau> has_real_derivative \<tau>' t) (at t within u)"
using del_has_vector_derivative
by (simp add:has_field_derivative_iff_has_vector_derivative)
lemma delay_image: "\<tau> ` {\<delta>..} = {0..}"
proof (rule subset_antisym, unfold image_def, unfold \<tau>_def)
show "{y. \<exists>x\<in>{\<delta>..}. y = x - \<delta>} \<subseteq> {0..}" by auto
next
show "{0..} \<subseteq> {y. \<exists>x\<in>{\<delta>..}. y = x - \<delta>}"
proof (rule subsetI)
fix a
assume "(a::real) \<in> {0..}"
hence "0 \<le> a" by simp
hence "\<exists>x\<in>{\<delta>..}. a = x - \<delta>" using bexI[where x = "a + \<delta>"] by auto
thus "a \<in> {y. \<exists>x\<in>{\<delta>..}. y = x - \<delta>}" by auto
qed
qed
lemma s_delayed_has_real_derivative[derivative_intros]:
assumes "\<delta> \<le> t"
shows "((ego2.s \<circ> \<tau>) has_field_derivative ego2.s' (t - \<delta>) * \<tau>' t) (at t within {\<delta>..})"
proof (rule DERIV_image_chain)
from assms have 0: "0 \<le> t - \<delta>" by simp
from ego2.t_stop_nonneg have 1: "v\<^sub>e / a\<^sub>e \<le> 0" unfolding ego2.t_stop_def by simp
from ego2.decel have 2: "a\<^sub>e \<noteq> 0" by simp
show "(ego2.s has_real_derivative ego2.s' (t - \<delta>)) (at (\<tau> t) within \<tau> ` {\<delta>..})"
using ego2.s_has_real_derivative[OF 0 1 2] sym[OF delay_image]
unfolding \<tau>_def by simp
next
from del_has_real_derivative show "(\<tau> has_real_derivative \<tau>' t) (at t within {\<delta>..})"
by auto
qed
lemma s_delayed_has_real_derivative' [derivative_intros]:
assumes "\<delta> \<le> t"
shows "((ego2.s \<circ> \<tau>) has_field_derivative (ego2.s' \<circ> \<tau>) t) (at t within {\<delta>..})"
proof -
from s_delayed_has_real_derivative[OF assms] have
"((ego2.s \<circ> \<tau>) has_field_derivative ego2.s' (t - \<delta>) * \<tau>' t) (at t within {\<delta>..})"
by auto
hence "((ego2.s \<circ> \<tau>) has_field_derivative ego2.s' (t - \<delta>) * 1) (at t within {\<delta>..})"
using \<tau>'_def[of t] by metis
hence "((ego2.s \<circ> \<tau>) has_field_derivative ego2.s' (t - \<delta>)) (at t within {\<delta>..})"
by (simp add:algebra_simps)
thus ?thesis unfolding comp_def \<tau>_def by auto
qed
lemma s_delayed_has_vector_derivative' [derivative_intros]:
assumes "\<delta> \<le> t"
shows "((ego2.s \<circ> \<tau>) has_vector_derivative (ego2.s' \<circ> \<tau>) t) (at t within {\<delta>..})"
using s_delayed_has_real_derivative'[OF assms]
by (simp add:has_field_derivative_iff_has_vector_derivative)
definition u :: "real \<Rightarrow> real" where
"u t = ( if t \<le> 0 then s\<^sub>e
else if t \<le> \<delta> then ego.q t
else (ego2.s \<circ> \<tau>) t)"
lemma init_u: "t \<le> 0 \<Longrightarrow> u t = s\<^sub>e" unfolding u_def by auto
lemma u_delta: "u \<delta> = ego2.s 0"
proof -
have "u \<delta> = ego.q \<delta>" using pos_react unfolding u_def by auto
also have "... = ego2.s 0" unfolding ego2.s_def by auto
finally show "u \<delta> = ego2.s 0" .
qed
lemma q_delta: "ego.q \<delta> = ego2.s 0" using u_delta pos_react unfolding u_def by auto
definition u' :: "real \<Rightarrow> real" where
"u' t = (if t \<le> \<delta> then ego.q' t else ego2.s' (t - \<delta>))"
lemma u'_delta: "u' \<delta> = ego2.s' 0"
proof -
have "u' \<delta> = ego.q' \<delta>" unfolding u'_def by auto
also have "... = v\<^sub>e" unfolding ego2.q'_def by simp
also have "... = ego2.p' 0" unfolding ego2.p'_def by simp
also have "... = ego2.s' 0" using ego2.t_stop_nonneg unfolding ego2.s'_def by auto
finally show "u' \<delta> = ego.s' 0" .
qed
lemma q'_delta: "ego.q' \<delta> = ego2.s' 0" using u'_delta unfolding u'_def by auto
lemma u_has_real_derivative[derivative_intros]:
assumes nonneg_t: "t \<ge> 0"
shows "(u has_real_derivative u' t) (at t within {0..})"
proof -
from pos_react have "0 \<le> \<delta>" by simp
have temp: "((\<lambda>t. if t \<in> {0 .. \<delta>} then ego.q t else (ego2.s \<circ> \<tau>) t) has_real_derivative
(if t \<in> {0..\<delta>} then ego.q' t else (ego2.s' \<circ> \<tau>) t)) (at t within {0..})" (is "(?f1 has_real_derivative ?f2) (?net)")
unfolding u_def[abs_def] u'_def
has_field_derivative_iff_has_vector_derivative
apply (rule has_vector_derivative_If_within_closures[where T = "{\<delta>..}"])
using \<open>0 \<le> \<delta>\<close> q_delta q'_delta ego.s_has_vector_derivative[OF assms] ego.decel ego.t_stop_nonneg
s_delayed_has_vector_derivative'[of "t"] \<tau>_def
unfolding comp_def
by (auto simp: assms max_def insert_absorb
intro!: ego.q_has_vector_derivative)
show ?thesis
unfolding has_vector_derivative_def has_field_derivative_iff_has_vector_derivative
u'_def u_def[abs_def]
proof (rule has_derivative_transform[where f="(\<lambda>t. if t \<in> {0..\<delta>} then ego.q t else (ego2.s \<circ> \<tau>) t)"])
from nonneg_t show " t \<in> {0..}" by auto
next
fix x
assume "(x::real) \<in> {0..}"
hence "x \<le> \<delta> \<longleftrightarrow> x \<in> {0 .. \<delta>}" by simp
thus " (if x \<le> 0 then s\<^sub>e else if x \<le> \<delta> then ego.q x else (ego2.s \<circ> \<tau>) x) =
(if x \<in> {0..\<delta>} then ego.q x else (ego2.s \<circ> \<tau>) x)" using pos_react unfolding ego.q_def by auto
next
from temp have "(?f1 has_vector_derivative ?f2) ?net"
using has_field_derivative_iff_has_vector_derivative by auto
moreover with assms have "t \<in> {0 .. \<delta>} \<longleftrightarrow> t \<le> \<delta>" by auto
ultimately show " ((\<lambda>t. if t \<in> {0..\<delta>} then ego.q t else (ego2.s \<circ> \<tau>) t) has_derivative
(\<lambda>x. x *\<^sub>R (if t \<le> \<delta> then ego2.q' t else ego2.s' (t - \<delta>)))) (at t within {0..})"
unfolding comp_def \<tau>_def has_vector_derivative_def by auto
qed
qed
definition t_stop :: real where
"t_stop = ego2.t_stop + \<delta>"
lemma t_stop_nonneg: "0 \<le> t_stop"
unfolding t_stop_def
using ego2.t_stop_nonneg pos_react
by auto
lemma t_stop_pos: "0 < t_stop"
unfolding t_stop_def
using ego2.t_stop_nonneg pos_react
by auto
lemma t_stop_zero:
assumes "t_stop \<le> x"
assumes "x \<le> \<delta>"
shows "v\<^sub>e = 0"
using assms unfolding t_stop_def using ego2.t_stop_nonneg pos_react ego2.t_stop_zero by auto
lemma u'_stop_zero: "u' t_stop = 0"
unfolding u'_def t_stop_def ego2.q'_def ego2.s'_def
using ego2.t_stop_nonneg ego2.p'_stop_zero decelerate_ego ego2.t_stop_zero by auto
definition u_max :: real where
"u_max = u (ego2.t_stop + \<delta>)"
lemma u_max_eq: "u_max = ego.q \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2"
proof (cases "ego2.t_stop = 0")
assume "ego2.t_stop = 0"
hence "v\<^sub>e = 0" using ego2.t_stop_zero by simp
with \<open>ego2.t_stop = 0\<close> show "u_max = ego.q \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" unfolding u_max_def u_def using pos_react by auto
next
assume "ego2.t_stop \<noteq> 0"
hence "u_max = (ego2.s \<circ> \<tau>) (ego2.t_stop + \<delta>)"
unfolding u_max_def u_def using ego2.t_stop_nonneg pos_react by auto
moreover have "... = ego2.s ego2.t_stop" unfolding comp_def \<tau>_def by auto
moreover have "... = ego2.p_max"
unfolding ego2.s_def ego2.p_max_def using \<open>ego2.t_stop \<noteq> 0\<close> ego2.t_stop_nonneg by auto
moreover have "... = ego.q \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2" using ego2.p_max_eq .
ultimately show ?thesis by auto
qed
lemma u_mono:
assumes "x \<le> y" "y \<le> t_stop"
shows "u x \<le> u y"
proof -
have "y \<le> 0 \<or> (0 < y \<and> y \<le> \<delta>) \<or> \<delta> < y" by auto
moreover
{ assume "y \<le> 0"
with assms have "x \<le> 0" by auto
with \<open>y \<le> 0\<close> have "u x \<le> u y" unfolding u_def by auto }
moreover
{ assume "0 < y \<and> y \<le> \<delta>"
with assms have "x \<le> \<delta>" by auto
hence "u x \<le> u y"
proof (cases "x \<le> 0")
assume "x \<le> 0"
with \<open>x \<le> \<delta>\<close> and \<open>0 < y \<and> y \<le> \<delta>\<close> show "u x \<le> u y" unfolding u_def using ego.q_min by auto
next
assume "\<not> x \<le> 0"
with \<open>0 < y \<and> y \<le> \<delta>\<close> and assms show "u x \<le> u y"
unfolding u_def using ego.q_mono by auto
qed }
moreover
{ assume "\<delta> < y"
have "u x \<le> u y"
proof (cases "\<delta> < x")
assume "\<delta> < x"
with pos_react have "\<not> x \<le> 0" by auto
moreover from \<open>\<delta> < y\<close> and pos_react have "\<not> y \<le> 0" by auto
ultimately show "u x \<le> u y" unfolding u_def comp_def
using assms ego2.s_mono[of "x - \<delta>" "y - \<delta>"] \<open>\<delta> < y\<close> \<open>\<delta> < x\<close> by (auto simp:\<tau>_def)
next
assume "\<not> \<delta> < x"
hence "x \<le> \<delta>" by simp
hence "u x \<le> ego.q \<delta>" unfolding u_def using pos_react nonneg_vel_ego
by (auto simp add:ego.q_def mult_left_mono)
also have "... = ego2.s (\<tau> \<delta>)" unfolding ego2.s_def unfolding \<tau>_def by auto
also have "... \<le> ego2.s (\<tau> y)" unfolding \<tau>_def using \<open>\<delta> < y\<close> by (auto simp add:ego2.s_mono)
also have "... = u y" unfolding u_def using \<open>\<delta> < y\<close> pos_react by auto
ultimately show "u x \<le> u y" by auto
qed }
ultimately show "u x \<le> u y" by auto
qed
lemma u_antimono: "x \<le> y \<Longrightarrow> t_stop \<le> x \<Longrightarrow> u y \<le> u x"
proof -
assume 1: "x \<le> y"
assume 2: "t_stop \<le> x"
hence "\<delta> \<le> x" unfolding \<tau>_def t_stop_def using pos_react ego2.t_stop_nonneg by auto
with 1 have "\<delta> \<le> y" by auto
from 1 and 2 have 3: "t_stop \<le> y" by auto
show "u y \<le> u x"
proof (cases "x \<noteq> \<delta> \<and> y \<noteq> \<delta>")
assume "x \<noteq> \<delta> \<and> y \<noteq> \<delta>"
hence "x \<noteq> \<delta>" and "y \<noteq> \<delta>" by auto
have "u y \<le> (ego2.s \<circ> \<tau>) y" unfolding u_def using \<open>\<delta> \<le> y\<close> \<open>y \<noteq> \<delta>\<close> pos_react by auto
also have "... \<le> (ego2.s \<circ> \<tau>) x" unfolding comp_def
proof (intro ego2.s_antimono)
show "\<tau> x \<le> \<tau> y" unfolding \<tau>_def using \<open>x \<le> y\<close> by auto
next
show "ego2.t_stop \<le> \<tau> x" unfolding \<tau>_def using \<open>t_stop \<le> x\<close> by (auto simp: t_stop_def)
qed
also have "... \<le> u x" unfolding u_def using \<open>\<delta> \<le> x\<close>\<open>x \<noteq> \<delta>\<close> pos_react by auto
ultimately show "u y \<le> u x" by auto
next
assume "\<not> (x \<noteq> \<delta> \<and> y \<noteq> \<delta>)"
have "x \<noteq> \<delta> \<longrightarrow> y \<noteq> \<delta>"
proof (rule impI; erule contrapos_pp[where Q="\<not> x = \<delta>"])
assume "\<not> y \<noteq> \<delta>"
hence "y = \<delta>" by simp
with \<open>t_stop \<le> y\<close> have "ego2.t_stop = 0" unfolding t_stop_def
using ego2.t_stop_nonneg by auto
with \<open>t_stop \<le> x\<close> have "x = \<delta>" unfolding t_stop_def using \<open>x \<le> y\<close> \<open>y = \<delta>\<close> by auto
thus "\<not> x \<noteq> \<delta>" by auto
qed
with \<open>\<not> (x \<noteq> \<delta> \<and> y \<noteq> \<delta>)\<close> have "(x = \<delta> \<and> y = \<delta>) \<or> (x = \<delta>)" by auto
moreover
{ assume "x = \<delta> \<and> y = \<delta>"
hence "x = \<delta>" and "y = \<delta>" by auto
hence "u y \<le> ego.q \<delta>" unfolding u_def using pos_react by auto
also have "... \<le> u x" unfolding u_def using \<open>x = \<delta>\<close> pos_react by auto
ultimately have "u y \<le> u x" by auto }
moreover
{ assume "x = \<delta>"
hence "ego2.t_stop = 0" using \<open>t_stop \<le> x\<close> ego2.t_stop_nonneg by (auto simp:t_stop_def)
hence "v\<^sub>e = 0" by (rule ego2.t_stop_zero)
hence "u y \<le> ego.q \<delta>"
using pos_react \<open>x = \<delta>\<close> \<open>x \<le> y\<close> \<open>v\<^sub>e = 0\<close>
unfolding u_def comp_def \<tau>_def ego2.s_def ego2.p_def ego2.p_max_def ego2.t_stop_def
by auto
also have "... \<le> u x" using \<open>x = \<delta>\<close> pos_react unfolding u_def by auto
ultimately have "u y \<le> u x" by auto }
ultimately show ?thesis by auto
qed
qed
lemma u_max: "u x \<le> u_max"
unfolding u_max_def using t_stop_def
by (cases "x \<le> t_stop") (auto intro: u_mono u_antimono)
lemma u_eq_u_stop: "NO_MATCH t_stop x \<Longrightarrow> x \<ge> t_stop \<Longrightarrow> u x = u_max"
proof -
assume "t_stop \<le> x"
with t_stop_pos have "0 < x" by auto
from \<open>t_stop \<le> x\<close> have "\<delta> \<le> x" unfolding t_stop_def using ego2.t_stop_nonneg by auto
show "u x = u_max"
proof (cases "x \<le> \<delta>")
assume "x \<le> \<delta>"
with \<open>t_stop \<le> x\<close> have "v\<^sub>e = 0" by (rule t_stop_zero)
also have "x = \<delta>" using \<open>x \<le> \<delta>\<close> and \<open>\<delta> \<le> x\<close> by auto
ultimately have "u x = ego.q \<delta>" unfolding u_def using pos_react by auto
also have "... = u_max" unfolding u_max_eq using \<open>v\<^sub>e = 0\<close> by auto
ultimately show "u x = u_max" by simp
next
assume "\<not> x \<le> \<delta>"
hence "\<delta> < x" by auto
hence "u x = (ego2.s \<circ> \<tau>) x" unfolding u_def using pos_react by auto
also have "... = ego2.s ego2.t_stop"
proof (unfold comp_def; unfold \<tau>_def; intro order.antisym)
have "x - \<delta> \<ge> ego2.t_stop" using \<open>t_stop \<le> x\<close> unfolding t_stop_def by auto
thus "ego2.s (x - \<delta>) \<le> ego2.s ego2.t_stop" by (rule ego2.s_antimono) simp
next
have "x - \<delta> \<ge> ego2.t_stop" using \<open>t_stop \<le> x\<close> unfolding t_stop_def by auto
thus "ego2.s ego2.t_stop \<le> ego2.s (x - \<delta>)" using ego2.t_stop_nonneg by (rule ego2.s_mono)
qed
also have "... = u_max" unfolding u_max_eq ego2.s_t_stop ego2.p_max_eq by auto
ultimately show "u x = u_max" by auto
qed
qed
lemma at_least_delta:
assumes "x \<le> \<delta>"
assumes "t_stop \<le> x"
shows "ego.q x = ego2.s (x - \<delta>)"
using assms ego2.t_stop_nonneg
unfolding t_stop_def ego2.s_def less_eq_real_def by auto
lemma continuous_on_u[continuous_intros]: "continuous_on T u"
unfolding u_def[abs_def]
using t_stop_nonneg pos_react at_least_delta
proof (intro continuous_on_subset[where t=T and s = "{..0} \<union> ({0..\<delta>} \<union> ({\<delta> .. t_stop} \<union> {t_stop ..}))"] continuous_on_If continuous_intros)
fix x
assume " \<not> x \<le> \<delta>"
assume "x \<in> {0..\<delta>}"
hence "0 \<le> x" and "x \<le> \<delta>" by auto
thus "ego.q x = (ego2.s \<circ> \<tau>) x"
unfolding comp_def \<tau>_def ego2.s_def
using \<open>\<not> x \<le> \<delta>\<close> by auto
next
fix x
assume "x \<in> {\<delta>..t_stop} \<union> {t_stop..}"
hence "\<delta> \<le> x" unfolding t_stop_def using pos_react ego.t_stop_nonneg by auto
also assume "x \<le> \<delta>"
ultimately have "x = \<delta>" by auto
thus "ego.q x = (ego2.s \<circ> \<tau>) x" unfolding comp_def \<tau>_def ego2.s_def by auto
next
fix t::real
assume "t \<in> {.. 0}"
hence "t \<le> 0" by auto
also assume "\<not> t \<le> 0"
ultimately have "t = 0" by auto
hence "s\<^sub>e = ego.q t" unfolding ego.q_def by auto
with pos_react \<open>t = 0\<close> show "s\<^sub>e = (if t \<le> \<delta> then ego.q t else (ego2.s \<circ> \<tau>) t)" by auto
next
fix t::real
assume "t \<in> {0..\<delta>} \<union> ({\<delta>..t_stop} \<union> {t_stop..})"
hence "0 \<le> t" using pos_react ego2.t_stop_nonneg by (auto simp: t_stop_def)
also assume "t \<le> 0"
ultimately have "t = 0" by auto
hence " s\<^sub>e = (if t \<le> \<delta> then ego.q t else (ego2.s \<circ> \<tau>) t)" using pos_react ego.init_q by auto
thus "s\<^sub>e = (if t \<le> \<delta> then ego.q t else (ego2.s \<circ> \<tau>) t)" by auto
next
show "T \<subseteq> {..0} \<union> ({0..\<delta>} \<union> ({\<delta>..t_stop} \<union> {t_stop..}))" by auto
qed
lemma isCont_u[continuous_intros]: "isCont u x"
using continuous_on_u[of UNIV]
by (auto simp:continuous_on_eq_continuous_at)
definition collision_react :: "real set \<Rightarrow> bool" where
"collision_react time_set \<equiv> (\<exists>t\<in>time_set. u t = other.s t )"
abbreviation no_collision_react :: "real set \<Rightarrow> bool" where
"no_collision_react time_set \<equiv> \<not> collision_react time_set"
lemma no_collision_reactI:
assumes "\<And>t. t \<in> S \<Longrightarrow> u t \<noteq> other.s t"
shows "no_collision_react S"
using assms
unfolding collision_react_def
by blast
lemma no_collision_union:
assumes "no_collision_react S"
assumes "no_collision_react T"
shows "no_collision_react (S \<union> T)"
using assms
unfolding collision_react_def
by auto
lemma collision_trim_subset:
assumes "collision_react S"
assumes "no_collision_react T"
assumes "T \<subseteq> S"
shows "collision_react (S - T)"
using assms
unfolding collision_react_def by auto
theorem cond_1r : "u_max < s\<^sub>o \<Longrightarrow> no_collision_react {0..}"
proof (rule no_collision_reactI, simp)
fix t :: real
assume "0 \<le> t"
have "u t \<le> u_max" by (rule u_max)
also assume "... < s\<^sub>o"
also have "... = other.s 0"
by (simp add: other.init_s)
also have "... \<le> other.s t"
using \<open>0 \<le> t\<close> hyps
by (intro other.s_mono) auto
finally show "u t \<noteq> other.s t"
by simp
qed
definition safe_distance_1r :: real where
"safe_distance_1r = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2"
lemma sd_1r_eq: "(s\<^sub>o - s\<^sub>e > safe_distance_1r) = (u_max < s\<^sub>o)"
proof -
have "(s\<^sub>o - s\<^sub>e > safe_distance_1r) = (s\<^sub>o - s\<^sub>e > v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2)" unfolding safe_distance_1r_def by auto
moreover have "... = (s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 < s\<^sub>o)" by auto
ultimately show ?thesis using u_max_eq ego.q_def by auto
qed
lemma sd_1r_correct:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_1r"
shows "no_collision_react {0..}"
proof -
from assms have "u_max < s\<^sub>o" using sd_1r_eq by auto
thus ?thesis by (rule cond_1r)
qed
lemma u_other_strict_ivt:
assumes "u t > other.s t"
shows "collision_react {0..<t}"
proof cases
assume "0 \<le> t"
with assms in_front
have "\<exists>x\<ge>0. x \<le> t \<and> other.s x - u x = 0"
by (intro IVT2) (auto simp: init_u other.init_s continuous_diff isCont_u other.isCont_s)
then show ?thesis
using assms
by (auto simp add: algebra_simps collision_react_def Bex_def order.order_iff_strict)
qed(insert assms hyps, auto simp: collision_react_def init_u other.init_s)
lemma collision_react_subset: "collision_react s \<Longrightarrow> s \<subseteq> t \<Longrightarrow> collision_react t"
by (auto simp:collision_react_def)
lemma u_other_ivt:
assumes "u t \<ge> other.s t"
shows "collision_react {0 .. t}"
proof cases
assume "u t > other.s t"
from u_other_strict_ivt[OF this]
show ?thesis
by (rule collision_react_subset) auto
qed (insert hyps assms; cases "t \<ge> 0"; force simp: collision_react_def init_u other.init_s)
theorem cond_2r:
assumes "u_max \<ge> other.s_stop"
shows "collision_react {0 ..}"
using assms
apply(intro collision_react_subset[where t="{0..}" and s ="{0 .. max t_stop other.t_stop}"])
apply(intro u_other_ivt[where t ="max t_stop other.t_stop"])
apply(auto simp: u_eq_u_stop other.s_eq_s_stop)
done
definition ego_other2 :: "real \<Rightarrow> real" where
"ego_other2 t = other.s t - u t"
lemma continuous_on_ego_other2[continuous_intros]: "continuous_on T ego_other2"
unfolding ego_other2_def[abs_def]
by (intro continuous_intros)
lemma isCont_ego_other2[continuous_intros]: "isCont ego_other2 x"
using continuous_on_ego_other2[of UNIV]
by (auto simp: continuous_on_eq_continuous_at)
definition ego_other2' :: "real \<Rightarrow> real" where
"ego_other2' t = other.s' t - u' t"
lemma ego_other2_has_real_derivative[derivative_intros]:
assumes "0 \<le> t"
shows "(ego_other2 has_real_derivative ego_other2' t) (at t within {0..})"
using assms other.t_stop_nonneg decelerate_other
unfolding other.t_stop_def
by (auto simp: ego_other2_def[abs_def] ego_other2'_def algebra_simps
intro!: derivative_eq_intros)
theorem cond_3r_1:
assumes "u \<delta> \<ge> other.s \<delta>"
shows "collision_react {0 .. \<delta>}"
proof (unfold collision_react_def)
have 1: "\<exists>t\<ge>0. t \<le> \<delta> \<and> ego_other2 t = 0"
proof (intro IVT2)
show "ego_other2 \<delta> \<le> 0" unfolding ego_other2_def using assms by auto
next
show "0 \<le> ego_other2 0" unfolding ego_other2_def
using other.init_s[of 0] init_u[of 0] in_front by auto
next
show "0 \<le> \<delta>" using pos_react by auto
next
show "\<forall>t. 0 \<le> t \<and> t \<le> \<delta> \<longrightarrow> isCont ego_other2 t"
using isCont_ego_other2 by auto
qed
then obtain t where "0 \<le> t \<and> t \<le> \<delta> \<and> ego_other2 t = 0" by auto
hence "t \<in> {0 .. \<delta>}" and "u t = other.s t" unfolding ego_other2_def by auto
thus "\<exists>t\<in>{0..\<delta>}. u t = other.s t" by (intro bexI)
qed
definition distance0 :: real where
"distance0 = v\<^sub>e * \<delta> - v\<^sub>o * \<delta> - a\<^sub>o * \<delta>\<^sup>2 / 2"
definition distance0_2 :: real where
"distance0_2 = v\<^sub>e * \<delta> + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o"
theorem cond_3r_1':
assumes "s\<^sub>o - s\<^sub>e \<le> distance0"
assumes "\<delta> \<le> other.t_stop"
shows "collision_react {0 .. \<delta>}"
proof -
from assms have "u \<delta> \<ge> other.s \<delta>" unfolding distance0_def other.s_def
other.p_def u_def ego.q_def using pos_react by auto
thus ?thesis using cond_3r_1 by auto
qed
theorem distance0_2_eq:
assumes "\<delta> > other.t_stop"
shows "(u \<delta> < other.s \<delta>) = (s\<^sub>o - s\<^sub>e > distance0_2)"
proof -
from assms have "(u \<delta> < other.s \<delta>) = (ego.q \<delta> < other.p_max)"
using u_def other.s_def pos_react by auto
also have "... = (s\<^sub>e + v\<^sub>e * \<delta> < s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2)"
using ego.q_def other.p_max_def other.p_def other.t_stop_def by auto
also have "... = (v\<^sub>e * \<delta> - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2 < s\<^sub>o - s\<^sub>e)" by linarith
also have "... = (v\<^sub>e * \<delta> + v\<^sub>o\<^sup>2 / a\<^sub>o - 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)"
using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto
also have "... = (v\<^sub>e * \<delta> + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" by linarith
thus ?thesis using distance0_2_def by (simp add: calculation)
qed
theorem cond_3r_1'_2:
assumes "s\<^sub>o - s\<^sub>e \<le> distance0_2"
assumes "\<delta> > other.t_stop"
shows "collision_react {0 .. \<delta>}"
proof -
from assms distance0_2_eq have "u \<delta> \<ge> other.s \<delta>" unfolding distance0_def other.s_def
other.p_def u_def ego.q_def using pos_react by auto
thus ?thesis using cond_3r_1 by auto
qed
definition safe_distance_3r :: real where
"safe_distance_3r = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2"
lemma distance0_at_most_sd3r:
"distance0 \<le> safe_distance_3r"
unfolding distance0_def safe_distance_3r_def using nonneg_vel_ego decelerate_ego
by (auto simp add:field_simps)
definition safe_distance_4r :: real where
"safe_distance_4r = (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>"
lemma distance0_at_most_sd4r:
assumes "a\<^sub>o > a\<^sub>e"
shows "distance0 \<le> safe_distance_4r"
proof -
from assms have "a\<^sub>o \<ge> a\<^sub>e" by auto
have "0 \<le> (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / (2 * a\<^sub>o - 2 * a\<^sub>e)"
- by (rule divide_nonneg_nonneg) (auto simp add:assms `a\<^sub>e \<le> a\<^sub>o`)
+ by (rule divide_nonneg_nonneg) (auto simp add:assms \<open>a\<^sub>e \<le> a\<^sub>o\<close>)
thus ?thesis unfolding distance0_def safe_distance_4r_def
by auto
qed
definition safe_distance_2r :: real where
"safe_distance_2r = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o"
lemma vo_start_geq_ve:
assumes "\<delta> \<le> other.t_stop"
assumes "other.s' \<delta> \<ge> v\<^sub>e"
shows "u \<delta> < other.s \<delta>"
proof -
from assms have "v\<^sub>e \<le> v\<^sub>o + a\<^sub>o * \<delta>" unfolding other.s'_def other.p'_def by auto
with mult_right_mono[OF this, of "\<delta>"] have "v\<^sub>e * \<delta> \<le> v\<^sub>o * \<delta> + a\<^sub>o * \<delta>\<^sup>2" (is "?l0 \<le> ?r0")
using pos_react by (auto simp add:field_simps power_def)
hence "s\<^sub>e + ?l0 \<le> s\<^sub>e + ?r0" by auto
also have "... < s\<^sub>o + ?r0" using in_front by auto
also have "... < s\<^sub>o + v\<^sub>o * \<delta> + a\<^sub>o * \<delta>\<^sup>2 / 2" using decelerate_other pos_react by auto
finally show ?thesis using pos_react assms(1)
unfolding u_def ego.q_def other.s_def other.t_stop_def other.p_def by auto
qed
theorem so_star_stop_leq_se_stop:
assumes "\<delta> \<le> other.t_stop"
assumes "other.s' \<delta> < v\<^sub>e"
assumes "\<not> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
shows "0 \<le> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 + (v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 / a\<^sub>o / 2"
proof -
consider "v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> \<ge> 0" | "\<not> (v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> \<ge> 0)" by auto
thus ?thesis
proof (cases)
case 1
hence "v\<^sub>e - a\<^sub>e / a\<^sub>o * (v\<^sub>o + a\<^sub>o * \<delta>) \<ge> 0" unfolding other.s'_def other.p'_def
by (auto simp add:assms(1))
hence "v\<^sub>e - a\<^sub>e / a\<^sub>o * v\<^sub>o - a\<^sub>e * \<delta> \<ge> 0" (is "?l0 \<ge> 0") using decelerate_other
by (auto simp add:field_simps)
- hence "?l0 / a\<^sub>e \<le> 0" using divide_right_mono_neg[OF `?l0 \<ge> 0`] decelerate_ego by auto
+ hence "?l0 / a\<^sub>e \<le> 0" using divide_right_mono_neg[OF \<open>?l0 \<ge> 0\<close>] decelerate_ego by auto
hence "0 \<ge> v\<^sub>e / a\<^sub>e - v\<^sub>o / a\<^sub>o - \<delta>" using decelerate_ego by (auto simp add:field_simps)
hence *: "- v\<^sub>e / a\<^sub>e \<ge> - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o" using decelerate_other by (auto simp add:field_simps)
from assms have **: "v\<^sub>o + a\<^sub>o * \<delta> \<le> v\<^sub>e" unfolding other.s'_def other.p'_def by auto
have vo_star_nneg: "v\<^sub>o + a\<^sub>o * \<delta> \<ge> 0"
proof -
from assms(1) have "- v\<^sub>o \<le> a\<^sub>o * \<delta>" unfolding other.t_stop_def using decelerate_other
by (auto simp add:field_simps)
thus ?thesis by auto
qed
- from mult_mono[OF * ** _ `0 \<le> v\<^sub>o + a\<^sub>o * \<delta>`]
+ from mult_mono[OF * ** _ \<open>0 \<le> v\<^sub>o + a\<^sub>o * \<delta>\<close>]
have "- (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o * (v\<^sub>o + a\<^sub>o * \<delta>) \<le> - v\<^sub>e / a\<^sub>e * v\<^sub>e" using nonneg_vel_ego decelerate_ego
by (auto simp add:field_simps)
hence "- (v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 / a\<^sub>o \<le> - v\<^sub>e\<^sup>2 / a\<^sub>e " by (auto simp add: field_simps power_def)
thus ?thesis by (auto simp add:field_simps)
next
case 2
with assms have "a\<^sub>o \<le> a\<^sub>e" by auto
from assms(2) have "(v\<^sub>o + a\<^sub>o * \<delta>) \<le> v\<^sub>e" unfolding other.s'_def using assms unfolding other.p'_def
by auto
have vo_star_nneg: "v\<^sub>o + a\<^sub>o * \<delta> \<ge> 0"
proof -
from assms(1) have "- v\<^sub>o \<le> a\<^sub>o * \<delta>" unfolding other.t_stop_def using decelerate_other
by (auto simp add:field_simps)
thus ?thesis by auto
qed
- with mult_mono[OF `v\<^sub>o + a\<^sub>o * \<delta> \<le> v\<^sub>e` `v\<^sub>o + a\<^sub>o * \<delta> \<le> v\<^sub>e`] have *: "(v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 \<le> v\<^sub>e\<^sup>2"
+ with mult_mono[OF \<open>v\<^sub>o + a\<^sub>o * \<delta> \<le> v\<^sub>e\<close> \<open>v\<^sub>o + a\<^sub>o * \<delta> \<le> v\<^sub>e\<close>] have *: "(v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 \<le> v\<^sub>e\<^sup>2"
using nonneg_vel_ego by (auto simp add:power_def)
- from `a\<^sub>o \<le> a\<^sub>e` have "- 1 /a\<^sub>o \<le> - 1 / a\<^sub>e" using decelerate_ego decelerate_other
+ from \<open>a\<^sub>o \<le> a\<^sub>e\<close> have "- 1 /a\<^sub>o \<le> - 1 / a\<^sub>e" using decelerate_ego decelerate_other
by (auto simp add:field_simps)
from mult_mono[OF * this] have "(v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 * (- 1 / a\<^sub>o) \<le> v\<^sub>e\<^sup>2 * (- 1 / a\<^sub>e)"
using nonneg_vel_ego decelerate_other by (auto simp add:field_simps)
then show ?thesis by auto
qed
qed
theorem distance0_at_most_distance2r:
assumes "\<delta> \<le> other.t_stop"
assumes "other.s' \<delta> < v\<^sub>e"
assumes "\<not> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
shows "distance0 \<le> safe_distance_2r"
proof -
from so_star_stop_leq_se_stop[OF assms] have " 0 \<le> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2 + (v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 / a\<^sub>o / 2 " (is "0 \<le> ?term")
by auto
have "safe_distance_2r = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto
also have "... = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + (v\<^sub>o + a\<^sub>o * \<delta>)\<^sup>2 / 2 / a\<^sub>o - v\<^sub>o * \<delta> - a\<^sub>o * \<delta>\<^sup>2 / 2"
using decelerate_other by (auto simp add:field_simps power_def)
also have "... = v\<^sub>e * \<delta> - v\<^sub>o * \<delta> - a\<^sub>o * \<delta>\<^sup>2 / 2 + ?term" (is "_ = ?left + ?term")
by (auto simp add:field_simps)
finally have "safe_distance_2r = distance0 + ?term" unfolding distance0_def by auto
- with `0 \<le> ?term` show "distance0 \<le> safe_distance_2r" by auto
+ with \<open>0 \<le> ?term\<close> show "distance0 \<le> safe_distance_2r" by auto
qed
theorem dist0_sd2r_1:
assumes "\<delta> \<le> other.t_stop"
assumes "\<not> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r"
shows "s\<^sub>o - s\<^sub>e > distance0"
proof (cases "other.s' \<delta> \<ge> v\<^sub>e")
assume "v\<^sub>e \<le> other.s' \<delta>"
from vo_start_geq_ve[OF assms(1) this] have "u \<delta> < other.s \<delta>" by auto
thus ?thesis unfolding distance0_def u_def using pos_react assms(1) unfolding ego.q_def
other.s_def other.p_def by auto
next
assume "\<not> v\<^sub>e \<le> other.s' \<delta>"
hence "v\<^sub>e > other.s' \<delta>" by auto
from distance0_at_most_distance2r[OF assms(1) this assms(2)] have "distance0 \<le> safe_distance_2r"
by auto
with assms(3) show ?thesis by auto
qed
theorem sd2r_eq:
assumes "\<delta> > other.t_stop"
shows "(u_max < other.s \<delta>) = (s\<^sub>o - s\<^sub>e > safe_distance_2r)"
proof -
from assms have "(u_max < other.s \<delta>) = (ego2.s (- v\<^sub>e / a\<^sub>e) < other.p_max)"
using u_max_def ego2.t_stop_def u_def other.s_def \<tau>_def pos_react ego2.p_max_eq ego2.s_t_stop u_max_eq by auto
also have "... = (s\<^sub>e + v\<^sub>e * \<delta> + v\<^sub>e * (- v\<^sub>e / a\<^sub>e) + 1 / 2 * a\<^sub>e * (- v\<^sub>e / a\<^sub>e)\<^sup>2 < s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2)"
using ego2.s_def ego2.p_def ego.q_def other.p_max_def other.p_def other.t_stop_def ego2.p_max_def ego2.s_t_stop ego2.t_stop_def by auto
also have "... = (v\<^sub>e * \<delta> + v\<^sub>e * (- v\<^sub>e / a\<^sub>e) + 1 / 2 * a\<^sub>e * (- v\<^sub>e / a\<^sub>e)\<^sup>2 - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1 / 2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2 < s\<^sub>o - s\<^sub>e)" by linarith
also have "... = (v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e + 1 / 2 * v\<^sub>e\<^sup>2 / a\<^sub>e + v\<^sub>o\<^sup>2 / a\<^sub>o - 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)"
using ego2.p_def ego2.p_max_def ego2.p_max_eq ego2.t_stop_def other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto
also have "... = (v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o < s\<^sub>o - s\<^sub>e)" by linarith
thus ?thesis using distance0_2_def by (simp add: calculation safe_distance_2r_def)
qed
theorem dist0_sd2r_2:
assumes "\<delta> > - v\<^sub>o / a\<^sub>o"
assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r"
shows "s\<^sub>o - s\<^sub>e > distance0_2"
proof -
have "- v\<^sub>e\<^sup>2 / 2 / a\<^sub>e \<ge> 0" using zero_le_power2 hyps(3) divide_nonneg_neg by (auto simp add:field_simps)
hence "v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o \<ge> v\<^sub>e * \<delta> + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" by simp
hence "safe_distance_2r \<ge> distance0_2" using safe_distance_2r_def distance0_2_def by auto
thus ?thesis using assms(2) by linarith
qed
end
subsection \<open>Safe Distance Delta\<close>
locale safe_distance_no_collsion_delta = safe_distance_normal +
assumes no_collision_delta: "u \<delta> < other.s \<delta>"
begin
sublocale delayed_safe_distance: safe_distance a\<^sub>e v\<^sub>e "ego.q \<delta>" a\<^sub>o "other.s' \<delta>" "other.s \<delta>"
proof (unfold_locales)
from nonneg_vel_ego show "0 \<le> v\<^sub>e" by auto
next
from nonneg_vel_other show "0 \<le> other.s' \<delta>" unfolding other.s'_def other.p'_def other.t_stop_def
using decelerate_other by (auto simp add: field_split_simps)
next
from decelerate_ego show "a\<^sub>e < 0" by auto
next
from decelerate_other show "a\<^sub>o < 0" by auto
next
from no_collision_delta show "ego.q \<delta> < other.s \<delta>" unfolding u_def using pos_react by auto
qed
lemma no_collision_react_initially_strict:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
shows "no_collision_react {0 <..< \<delta>}"
proof (rule no_collision_reactI)
fix t::real
assume "t \<in> {0 <..< \<delta>}"
show "u t \<noteq> other.s t"
proof (rule ccontr)
assume "\<not> u t \<noteq> other.s t"
hence "ego_other2 t = 0" unfolding ego_other2_def by auto
from \<open>t \<in> {0 <..< \<delta>}\<close> have "ego_other2 t = other.s t - ego.q t"
unfolding ego_other2_def u_def using ego.init_q by auto
have "\<delta> \<le> other.t_stop \<or> other.t_stop < \<delta>" by auto
moreover
{ assume le_t_stop: "\<delta> \<le> other.t_stop"
with \<open>ego_other2 t = other.s t - ego.q t\<close> have "ego_other2 t = other.p t - ego.q t"
unfolding other.s_def using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
with \<open>ego_other2 t = 0\<close> have "other.p t - ego.q t = 0" by auto
hence eq: "(s\<^sub>o- s\<^sub>e) + (v\<^sub>o - v\<^sub>e) * t + (1/2 * a\<^sub>o) * t\<^sup>2 = 0"
unfolding other.p_def ego.q_def by (auto simp: algebra_simps)
define p where "p \<equiv> \<lambda>x. (1/2 * a\<^sub>o) * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e)"
have "0 < 1/2 * a\<^sub>o"
proof (intro p_convex[where p=p and b="v\<^sub>o - v\<^sub>e" and c="s\<^sub>o - s\<^sub>e"])
show "0 < t" using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
next
show "t < \<delta>" using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
next
show "p t < p 0" unfolding p_def using eq in_front by (auto simp: algebra_simps)
next
from eq have "p t = 0" unfolding p_def by auto
also have "... < p \<delta>" using no_collision_delta pos_react le_t_stop
unfolding p_def u_def other.s_def ego.q_def other.p_def by (auto simp:algebra_simps)
finally have "p t < p \<delta>" by simp
thus "p t \<le> p \<delta>" by auto
next
show "p = (\<lambda>x. 1 / 2 * a\<^sub>o * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e))" unfolding p_def
by (rule refl)
qed
hence "0 < a\<^sub>o" by auto
with decelerate_other have False by simp }
moreover
{ assume gt_t_stop: "\<delta> > other.t_stop"
have t_lt_t_stop: "t < other.t_stop"
proof (rule ccontr)
assume "\<not> t < other.t_stop"
hence "other.t_stop \<le> t" by simp
from \<open>ego_other2 t = 0\<close> have "ego.q t = other.p_max"
unfolding ego_other2_def u_def other.s_def comp_def \<tau>_def other.p_max_def
using \<open>t \<in> {0 <..< \<delta>}\<close> \<open>other.t_stop \<le> t\<close> gt_t_stop by (auto split:if_splits)
have "ego.q t = u t" unfolding u_def using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
also have "... \<le> u_max" using u_max by auto
also have "... < other.p_max" using assms(2) other.s_t_stop by auto
finally have "ego.q t < other.p_max" by auto
with \<open>ego.q t = other.p_max\<close> show False by auto
qed
with \<open>ego_other2 t = other.s t - ego.q t\<close> have "ego_other2 t = other.p t - ego.q t"
unfolding other.s_def using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
with \<open>ego_other2 t = 0\<close> have "other.p t - ego.q t = 0" by auto
hence eq: "(s\<^sub>o- s\<^sub>e) + (v\<^sub>o - v\<^sub>e) * t + (1/2 * a\<^sub>o) * t\<^sup>2 = 0"
unfolding other.p_def ego.q_def by (auto simp: algebra_simps)
define p where "p \<equiv> \<lambda>x. (1/2 * a\<^sub>o) * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e)"
have "0 < 1/2 * a\<^sub>o"
proof (intro p_convex[where p=p and b="v\<^sub>o - v\<^sub>e" and c="s\<^sub>o - s\<^sub>e"])
show "0 < t" using \<open>t \<in> {0 <..< \<delta>}\<close> by auto
next
show "t < other.t_stop" using t_lt_t_stop by auto
next
show "p t < p 0" unfolding p_def using eq in_front by (auto simp: algebra_simps)
next
from eq have zero: "p t = 0" unfolding p_def by auto
have eq: "p other.t_stop = ego_other2 other.t_stop"
unfolding ego_other2_def other.s_t_stop u_def ego.q_def
other.s_def other.p_def p_def
using \<open>\<delta> > other.t_stop\<close> other.t_stop_nonneg other.t_stop_def
by (auto simp: diff_divide_distrib left_diff_distrib')
have "u other.t_stop \<le> u_max" using u_max by auto
also have "... < other.s_stop" using assms by auto
finally have "0 \<le> other.s_stop - u other.t_stop" by auto
hence "0 \<le> ego_other2 other.t_stop" unfolding ego_other2_def by auto
hence "0 \<le> p other.t_stop" using eq by auto
with zero show "p t \<le> p other.t_stop" by auto
next
show "p = (\<lambda>x. 1 / 2 * a\<^sub>o * x\<^sup>2 + (v\<^sub>o - v\<^sub>e) * x + (s\<^sub>o - s\<^sub>e))"
unfolding p_def by (rule refl)
qed
hence False using decelerate_other by auto }
ultimately show False by auto
qed
qed
lemma no_collision_react_initially:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
shows "no_collision_react {0 .. \<delta>}"
proof -
have "no_collision_react {0 <..< \<delta>}" by (rule no_collision_react_initially_strict[OF assms])
have "u 0 \<noteq> other.s 0" using init_u other.init_s in_front by auto
hence "no_collision_react {0}" unfolding collision_react_def by auto
with \<open>no_collision_react {0 <..< \<delta>}\<close> have "no_collision_react ({0} \<union> {0 <..< \<delta>})"
using no_collision_union[of "{0}" "{0 <..< \<delta>}"] by auto
moreover have "{0} \<union> {0 <..< \<delta>} = {0 ..< \<delta>}" using pos_react by auto
ultimately have "no_collision_react {0 ..< \<delta>}" by auto
have "u \<delta> \<noteq> other.s \<delta>" using no_collision_delta by auto
hence "no_collision_react {\<delta>}" unfolding collision_react_def by auto
with \<open>no_collision_react {0 ..< \<delta>}\<close> have "no_collision_react ({\<delta>} \<union> {0 ..< \<delta>})"
using no_collision_union[of "{\<delta>}" "{0 ..< \<delta>}"] by auto
moreover have "{\<delta>} \<union> {0 ..< \<delta>} = {0 .. \<delta>}" using pos_react by auto
ultimately show "no_collision_react {0 .. \<delta>}" by auto
qed
lemma collision_after_delta:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
shows "collision_react {0 ..} \<longleftrightarrow> collision_react {\<delta>..}"
proof
assume "collision_react {0 ..}"
have "no_collision_react {0 .. \<delta>}" by (rule no_collision_react_initially[OF assms])
with \<open>collision_react {0..}\<close> have "collision_react ({0..} - {0 .. \<delta>})"
using pos_react by (auto intro: collision_trim_subset)
moreover have "{0..} - {0 .. \<delta>} = {\<delta> <..}" using pos_react by auto
ultimately have "collision_react {\<delta> <..}" by auto
thus "collision_react {\<delta> ..}" by (auto intro:collision_react_subset)
next
assume "collision_react {\<delta>..}"
moreover have "{\<delta>..} \<subseteq> {0 ..}" using pos_react by auto
ultimately show "collision_react {0 ..}" by (rule collision_react_subset)
qed
lemma collision_react_strict:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
shows "collision_react {\<delta> ..} \<longleftrightarrow> collision_react {\<delta> <..}"
proof
assume asm: "collision_react {\<delta> ..}"
have "no_collision_react {\<delta>}" using no_collision_delta unfolding collision_react_def by auto
moreover have "{\<delta> <..} \<subseteq> {\<delta> ..}" by auto
ultimately have "collision_react ({\<delta> ..} - {\<delta>})" using asm collision_trim_subset by simp
moreover have "{\<delta> <..} = {\<delta> ..} - {\<delta>}" by auto
ultimately show "collision_react {\<delta> <..}" by auto
next
assume "collision_react {\<delta> <..}"
thus "collision_react {\<delta> ..}"
using collision_react_subset[where t="{\<delta> ..}" and s="{\<delta> <..}"] by fastforce
qed
lemma delayed_other_s_stop_eq: "delayed_safe_distance.other.s_stop = other.s_stop"
proof (unfold other.s_t_stop; unfold delayed_safe_distance.other.s_t_stop; unfold movement.p_max_eq)
have "\<delta> \<le> other.t_stop \<or> other.t_stop < \<delta>" by auto
moreover
{ assume "\<delta> \<le> other.t_stop"
hence "other.s \<delta> - (other.s' \<delta>)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2"
unfolding other.s_def other.s'_def
using pos_react decelerate_other
by (auto simp add: other.p_def other.p'_def power2_eq_square field_split_simps) }
moreover
{ assume "other.t_stop < \<delta>"
hence "other.s \<delta> - (other.s' \<delta>)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2"
unfolding other.s_def other.s'_def other.p_max_eq
using pos_react decelerate_other
by (auto) }
ultimately show "other.s \<delta> - (other.s' \<delta>)\<^sup>2 / a\<^sub>o / 2 = s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o / 2" by auto
qed
lemma delayed_cond3':
assumes "other.s \<delta> \<le> u_max"
assumes "u_max < other.s_stop"
shows "delayed_safe_distance.collision {0 ..} \<longleftrightarrow>
(a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
proof (rule delayed_safe_distance.cond_3')
have "other.s \<delta> \<le> u_max" using \<open>other.s \<delta> \<le> u_max\<close> .
also have "... = ego2.s_stop" unfolding u_max_eq ego2.s_t_stop ego2.p_max_eq by (rule refl)
finally show "other.s \<delta> \<le> ego2.s_stop" by auto
next
have "ego2.s_stop = u_max" unfolding ego2.s_t_stop ego2.p_max_eq u_max_eq by (rule refl)
also have "... < other.s_stop" using assms by auto
also have "... \<le> delayed_safe_distance.other.s_stop" using delayed_other_s_stop_eq by auto
finally show "ego2.s_stop < delayed_safe_distance.other.s_stop" by auto
qed
lemma delayed_other_t_stop_eq:
assumes "\<delta> \<le> other.t_stop"
shows "delayed_safe_distance.other.t_stop + \<delta> = other.t_stop"
using assms decelerate_other
unfolding delayed_safe_distance.other.t_stop_def other.t_stop_def other.s'_def
movement.t_stop_def other.p'_def
by (auto simp add: field_split_simps)
lemma delayed_other_s_eq:
assumes "0 \<le> t"
shows "delayed_safe_distance.other.s t = other.s (t + \<delta>)"
proof (cases "\<delta> \<le> other.t_stop")
assume 1: "\<delta> \<le> other.t_stop"
have "t + \<delta> \<le> other.t_stop \<or> other.t_stop < t + \<delta>" by auto
moreover
{ assume "t + \<delta> \<le> other.t_stop"
hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p t"
using delayed_other_t_stop_eq [OF 1] assms
unfolding delayed_safe_distance.other.s_def by auto
also have "... = other.p (t + \<delta>)"
unfolding movement.p_def other.s_def other.s'_def other.p'_def
using pos_react 1
by (auto simp add: power2_eq_square field_split_simps)
also have "... = other.s (t + \<delta>)"
unfolding other.s_def
using assms pos_react \<open>t + \<delta> \<le> other.t_stop\<close> by auto
finally have "delayed_safe_distance.other.s t = other.s (t + \<delta>)" by auto }
moreover
{ assume "other.t_stop < t + \<delta>"
hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p_max"
using delayed_other_t_stop_eq [OF 1] assms delayed_safe_distance.other.t_stop_nonneg
unfolding delayed_safe_distance.other.s_def by auto
also have "... = other.p_max"
unfolding movement.p_max_eq other.s_def other.s'_def other.p_def other.p'_def
using pos_react 1 decelerate_other
by (auto simp add: power2_eq_square field_split_simps)
also have "... = other.s (t + \<delta>)"
unfolding other.s_def
using assms pos_react \<open>other.t_stop < t + \<delta>\<close> by auto
finally have "delayed_safe_distance.other.s t = other.s (t + \<delta>)" by auto }
ultimately show ?thesis by auto
next
assume "\<not> \<delta> \<le> other.t_stop"
hence "other.t_stop < \<delta>" by auto
hence "other.s' \<delta> = 0" and "other.s \<delta> = other.p_max"
unfolding other.s'_def other.s_def using pos_react by auto
hence "delayed_safe_distance.other.s t = delayed_safe_distance.other.p_max"
unfolding delayed_safe_distance.other.s_def using assms decelerate_other
by (auto simp add:movement.p_max_eq movement.p_def movement.t_stop_def)
also have "... = other.p_max"
unfolding movement.p_max_eq using \<open>other.s' \<delta> = 0\<close> \<open>other.s \<delta> = other.p_max\<close>
using other.p_max_eq by auto
also have "... = other.s (t + \<delta>)"
unfolding other.s_def using pos_react assms \<open>other.t_stop < \<delta>\<close> by auto
finally show "delayed_safe_distance.other.s t = other.s (t + \<delta>)" by auto
qed
lemma translate_collision_range:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
shows "delayed_safe_distance.collision {0 ..} \<longleftrightarrow> collision_react {\<delta> ..}"
proof
assume "delayed_safe_distance.collision {0 ..}"
then obtain t where eq: "ego2.s t = delayed_safe_distance.other.s t" and "0 \<le> t"
unfolding delayed_safe_distance.collision_def by auto
have "ego2.s t = (ego2.s \<circ> \<tau>) (t + \<delta>)" unfolding comp_def \<tau>_def by auto
also have "... = u (t + \<delta>)" unfolding u_def using \<open>0 \<le> t\<close> pos_react
by (auto simp: \<tau>_def ego2.init_s)
finally have left:"ego2.s t = u (t + \<delta>)" by auto
have right: "delayed_safe_distance.other.s t = other.s (t + \<delta>)"
using delayed_other_s_eq pos_react \<open>0 \<le> t\<close> by auto
with eq and left have "u (t + \<delta>) = other.s (t + \<delta>)" by auto
moreover have "\<delta> \<le> t + \<delta>" using \<open>0 \<le> t\<close> by auto
ultimately show "collision_react {\<delta> ..}" unfolding collision_react_def by auto
next
assume "collision_react {\<delta> ..}"
hence "collision_react {\<delta> <..}" using collision_react_strict[OF assms] by simp
then obtain t where eq: "u t = other.s t" and "\<delta> < t"
unfolding collision_react_def by auto
moreover hence "u t = (ego2.s \<circ> \<tau>) t" unfolding u_def using pos_react by auto
moreover have "other.s t = delayed_safe_distance.other.s (t - \<delta>)"
using delayed_other_s_eq \<open>\<delta> < t\<close> by auto
ultimately have "ego2.s (t - \<delta>) = delayed_safe_distance.other.s (t - \<delta>)"
unfolding comp_def \<tau>_def by auto
with \<open>\<delta> < t\<close> show "delayed_safe_distance.collision {0 ..}"
unfolding delayed_safe_distance.collision_def by auto
qed
theorem cond_3r_2:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
assumes "other.s \<delta> \<le> u_max"
shows "collision_react {0 ..} \<longleftrightarrow>
(a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
proof -
have "collision_react {0 ..} \<longleftrightarrow> collision_react {\<delta> ..}" by (rule collision_after_delta[OF assms(1) assms(2)])
also have "... \<longleftrightarrow> delayed_safe_distance.collision {0 ..}" by (simp add: translate_collision_range[OF assms(1) assms(2)])
also have "... \<longleftrightarrow> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
by (rule delayed_cond3'[OF assms(3) assms(2)])
finally show "collision_react {0 ..} \<longleftrightarrow> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
by auto
qed
lemma sd_2r_correct_for_3r_2:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r"
assumes "other.s \<delta> \<le> u_max"
assumes "\<not> (a\<^sub>o > a\<^sub>e \<and> other.s' \<delta> < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * other.s' \<delta> < 0)"
shows "no_collision_react {0..}"
proof -
from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto
hence "s\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto
hence "s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto
hence "s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1/2 * a\<^sub>o * (-v\<^sub>o / a\<^sub>o)\<^sup>2 > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e"
using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto
hence "other.s_stop > u_max" unfolding other.s_def using u_max_eq other.t_stop_def
using ego.q_def other.p_def other.p_max_def other.s_def other.s_t_stop by auto
thus ?thesis
using assms(2) assms(3) collision_after_delta cond_1r delayed_cond3' translate_collision_range by linarith
qed
lemma sd2_at_most_sd4:
assumes "a\<^sub>o > a\<^sub>e"
shows "safe_distance_2r \<le> safe_distance_4r"
proof -
have "a\<^sub>o \<noteq> 0" and "a\<^sub>e \<noteq> 0" and "a\<^sub>o - a\<^sub>e \<noteq> 0" and "0 < 2 * (a\<^sub>o - a\<^sub>e)" using hyps assms(1) by auto
have "0 \<le> (- v\<^sub>e * a\<^sub>o + v\<^sub>o * a\<^sub>e + a\<^sub>o * a\<^sub>e * \<delta>) * (- v\<^sub>e * a\<^sub>o + v\<^sub>o * a\<^sub>e + a\<^sub>o * a\<^sub>e * \<delta>)"
(is "0 \<le> (?l1 + ?l2 + ?l3) * ?r") by auto
also have "... = v\<^sub>e\<^sup>2 * a\<^sub>o\<^sup>2 + v\<^sub>o\<^sup>2 * a\<^sub>e\<^sup>2 + a\<^sub>o\<^sup>2 * a\<^sub>e\<^sup>2 * \<delta>\<^sup>2 - 2 * v\<^sub>e * a\<^sub>o * v\<^sub>o * a\<^sub>e - 2 * a\<^sub>o\<^sup>2 * a\<^sub>e * \<delta> * v\<^sub>e + 2 * a\<^sub>o * a\<^sub>e\<^sup>2 * \<delta> * v\<^sub>o"
(is "?lhs = ?rhs")
by (auto simp add:algebra_simps power_def)
finally have "0 \<le> ?rhs" by auto
hence "(- v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o) * (a\<^sub>o * a\<^sub>e) \<le> (a\<^sub>o * a\<^sub>e * \<delta>\<^sup>2 - 2 * v\<^sub>e * v\<^sub>o - 2 * a\<^sub>o * \<delta> * v\<^sub>e + 2 * a\<^sub>e * \<delta> * v\<^sub>o) * (a\<^sub>o * a\<^sub>e)"
by (auto simp add: algebra_simps power_def)
hence "2 * v\<^sub>e * \<delta> * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e + v\<^sub>e\<^sup>2 + v\<^sub>o\<^sup>2 - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o \<le> v\<^sub>o\<^sup>2 + a\<^sub>o\<^sup>2 * \<delta>\<^sup>2 + v\<^sub>e\<^sup>2 + 2 * v\<^sub>o * \<delta> * a\<^sub>o - 2 * v\<^sub>e * v\<^sub>o - 2 * a\<^sub>o * \<delta> * v\<^sub>e - 2 * v\<^sub>o * \<delta> * a\<^sub>o + 2 * a\<^sub>e * \<delta> * v\<^sub>o - a\<^sub>o\<^sup>2 * \<delta>\<^sup>2 + a\<^sub>o * a\<^sub>e * \<delta>\<^sup>2 + 2 * v\<^sub>e * \<delta> * (a\<^sub>o - a\<^sub>e)"
by (auto simp add: ego2.decel other.decel)
hence "2 * v\<^sub>e * \<delta> * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 * a\<^sub>o / a\<^sub>e + v\<^sub>e\<^sup>2 + v\<^sub>o\<^sup>2 - v\<^sub>o\<^sup>2 * a\<^sub>e / a\<^sub>o \<le> (v\<^sub>o + \<delta> * a\<^sub>o - v\<^sub>e)\<^sup>2 - 2 * v\<^sub>o * \<delta> * a\<^sub>o + 2 * a\<^sub>e * \<delta> * v\<^sub>o - a\<^sub>o\<^sup>2 * \<delta>\<^sup>2 + a\<^sub>o * a\<^sub>e * \<delta>\<^sup>2 + 2 * v\<^sub>e * \<delta> * (a\<^sub>o - a\<^sub>e)"
by (auto simp add: algebra_simps power_def)
hence "v\<^sub>e * \<delta> * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * a\<^sub>o + v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * a\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * a\<^sub>e \<le> (v\<^sub>o + \<delta> * a\<^sub>o - v\<^sub>e)\<^sup>2 - v\<^sub>o * \<delta> * 2 * a\<^sub>o - v\<^sub>o * \<delta> * 2 * -a\<^sub>e - a\<^sub>o * \<delta>\<^sup>2 / 2 * 2 * a\<^sub>o - a\<^sub>o * \<delta>\<^sup>2 / 2 * 2 * -a\<^sub>e + v\<^sub>e * \<delta> * 2 * (a\<^sub>o - a\<^sub>e)"
(is "?lhs1 \<le> ?rhs1")
by (simp add: \<open>a\<^sub>o \<noteq> 0\<close> \<open>a\<^sub>e \<noteq> 0\<close> power2_eq_square algebra_simps)
hence "v\<^sub>e * \<delta> * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e * 2 * (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o * 2 * (a\<^sub>o - a\<^sub>e) \<le> (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) * 2 * (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> * 2 * (a\<^sub>o - a\<^sub>e) - a\<^sub>o * \<delta>\<^sup>2 / 2 * 2 * (a\<^sub>o - a\<^sub>e) + v\<^sub>e * \<delta> * 2 *(a\<^sub>o - a\<^sub>e)"
(is "?lhs2 \<le> ?rhs2")
proof -
assume "?lhs1 \<le> ?rhs1"
have "?lhs1 = ?lhs2" by (auto simp add:field_simps)
moreover
- have "?rhs1 = ?rhs2" using `a\<^sub>o - a\<^sub>e \<noteq> 0` by (auto simp add:field_simps)
- ultimately show ?thesis using `?lhs1 \<le> ?rhs1` by auto
+ have "?rhs1 = ?rhs2" using \<open>a\<^sub>o - a\<^sub>e \<noteq> 0\<close> by (auto simp add:field_simps)
+ ultimately show ?thesis using \<open>?lhs1 \<le> ?rhs1\<close> by auto
qed
hence "(v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o) * 2 * (a\<^sub>o - a\<^sub>e) \<le> ((v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>) * 2 *(a\<^sub>o - a\<^sub>e)"
by (simp add: algebra_simps)
hence "v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o \<le> (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>"
- using \<open>a\<^sub>o > a\<^sub>e\<close> mult_le_cancel_iff1[OF `0 < 2 * (a\<^sub>o - a\<^sub>e)`, of "(v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o)"
+ using \<open>a\<^sub>o > a\<^sub>e\<close> mult_le_cancel_iff1[OF \<open>0 < 2 * (a\<^sub>o - a\<^sub>e)\<close>, of "(v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o)"
"(v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>"] semiring_normalization_rules(18)
by (metis (no_types, lifting))
thus ?thesis using safe_distance_2r_def safe_distance_4r_def by auto
qed
lemma sd_4r_correct:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_4r"
assumes "other.s \<delta> \<le> u_max"
assumes "\<delta> \<le> other.t_stop"
assumes "a\<^sub>o > a\<^sub>e"
shows "no_collision_react {0..}"
proof -
from assms have "s\<^sub>o - s\<^sub>e > (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>"
unfolding safe_distance_4r_def by auto
hence "s\<^sub>o + v\<^sub>o * \<delta> + 1/2 * a\<^sub>o * \<delta>\<^sup>2 - s\<^sub>e - v\<^sub>e * \<delta> > (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)" by linarith
hence "other.s \<delta> - ego.q \<delta> > (other.s' \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)"
using assms(3) ego.q_def other.p_def other.s_def other.p'_def other.s'_def pos_react by auto
hence "other.s \<delta> - ego.q \<delta> > delayed_safe_distance.snd_safe_distance"
by (simp add: delayed_safe_distance.snd_safe_distance_def)
hence c: "\<not> (other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance)" by linarith
have "u_max < other.s_stop"
unfolding u_max_eq other.s_t_stop other.p_max_eq ego.q_def using assms(1) sd2_at_most_sd4[OF assms(4)]
unfolding safe_distance_4r_def safe_distance_2r_def by auto
consider "s\<^sub>o \<le> u_max" | "s\<^sub>o > u_max" by linarith
thus ?thesis
proof (cases)
case 1
- from cond_3r_2[OF this `u_max < other.s_stop` assms(2)] show ?thesis
+ from cond_3r_2[OF this \<open>u_max < other.s_stop\<close> assms(2)] show ?thesis
using c by auto
next
case 2
then show ?thesis using cond_1r by auto
qed
qed
text \<open>Irrelevant, this Safe Distance is unreachable in the checker.\<close>
definition safe_distance_5r :: real where
"safe_distance_5r = v\<^sub>e\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o + v\<^sub>e * \<delta>"
lemma sd_5r_correct:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_5r"
assumes "u_max < other.s_stop"
assumes "other.s \<delta> \<le> u_max"
assumes "\<delta> > other.t_stop"
shows "no_collision_react {0..}"
proof -
from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o + v\<^sub>e * \<delta>"
unfolding safe_distance_5r_def by auto
hence "s\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o - s\<^sub>e - v\<^sub>e * \<delta> > (0 - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)"
using assms(2-4) unfolding other.s_def other.s_t_stop
apply (auto simp: movement.p_t_stop split: if_splits)
using cond_1r cond_2r other.s_t_stop by linarith+
hence "other.s \<delta> - ego.q \<delta> > (other.s' \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e)"
using assms(2) assms(3) assms(4) other.s_def other.s_t_stop by auto
hence "other.s \<delta> - ego.q \<delta> > delayed_safe_distance.snd_safe_distance"
by (simp add: delayed_safe_distance.snd_safe_distance_def)
hence "\<not> (other.s \<delta> - ego.q \<delta> \<le> delayed_safe_distance.snd_safe_distance)" by linarith
thus ?thesis using assms(2) assms(3) cond_1r cond_3r_2 by linarith
qed
lemma translate_no_collision_range:
"delayed_safe_distance.no_collision {0 ..} \<longleftrightarrow> no_collision_react {\<delta> ..}"
proof
assume left: "delayed_safe_distance.no_collision {0 ..}"
show "no_collision_react {\<delta> ..}"
proof (unfold collision_react_def; simp; rule ballI)
fix t::real
assume "t \<in> {\<delta> ..}"
hence "\<delta> \<le> t" by simp
with pos_react have "0 \<le> t - \<delta>" by simp
with left have ineq: "ego2.s (t - \<delta>) \<noteq> delayed_safe_distance.other.s (t - \<delta>)"
unfolding delayed_safe_distance.collision_def by auto
have "ego2.s (t - \<delta>) = (ego2.s \<circ> \<tau>) t" unfolding comp_def \<tau>_def by auto
also have "... = u t" unfolding u_def using \<open>\<delta> \<le> t\<close> pos_react
by (auto simp: \<tau>_def ego2.init_s)
finally have "ego2.s (t - \<delta>) = u t" by auto
moreover have "delayed_safe_distance.other.s (t - \<delta>) = other.s t"
using delayed_other_s_eq pos_react \<open>\<delta> \<le> t\<close> by auto
ultimately show "u t \<noteq> other.s t" using ineq by auto
qed
next
assume right: "no_collision_react {\<delta> ..}"
show "delayed_safe_distance.no_collision {0 ..}"
proof (unfold delayed_safe_distance.collision_def; simp; rule ballI)
fix t ::real
assume "t \<in> {0 ..}"
hence "0 \<le> t" by auto
hence "\<delta> \<le> t + \<delta>" by auto
with right have ineq: "u (t + \<delta>) \<noteq> other.s (t + \<delta>)" unfolding collision_react_def by auto
have "u (t + \<delta>) = ego2.s t" unfolding u_def comp_def \<tau>_def
using \<open>0 \<le> t\<close> pos_react \<open>\<delta> \<le> t+ \<delta>\<close> by (auto simp add:ego2.init_s)
moreover have "other.s (t + \<delta>) = delayed_safe_distance.other.s t"
using delayed_other_s_eq[of t] using \<open>0 \<le> t\<close> by auto
ultimately show "ego2.s t \<noteq> delayed_safe_distance.other.s t" using ineq by auto
qed
qed
lemma delayed_cond1:
assumes "other.s \<delta> > u_max"
shows "delayed_safe_distance.no_collision {0 ..}"
proof -
have "ego2.s_stop = u_max" unfolding ego2.s_t_stop ego2.p_max_eq u_max_eq by auto
also have "... < other.s \<delta>" using assms by simp
finally have "ego2.s_stop < other.s \<delta>" by auto
thus "delayed_safe_distance.no_collision {0 ..}" by (simp add: delayed_safe_distance.cond_1)
qed
theorem cond_3r_3:
assumes "s\<^sub>o \<le> u_max"
assumes "u_max < other.s_stop"
assumes "other.s \<delta> > u_max"
shows "no_collision_react {0 ..}"
proof -
have eq: "{0 ..} = {0 .. \<delta>} \<union> {\<delta> ..}" using pos_react by auto
show ?thesis unfolding eq
proof (intro no_collision_union)
show "no_collision_react {0 .. \<delta>}" by (rule no_collision_react_initially[OF assms(1) assms(2)])
next
have "delayed_safe_distance.no_collision {0 ..}" by (rule delayed_cond1[OF assms(3)])
with translate_no_collision_range show "no_collision_react {\<delta> ..}" by auto
qed
qed
lemma sd_2r_correct_for_3r_3:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_2r"
assumes "other.s \<delta> > u_max"
shows "no_collision_react {0..}"
proof -
from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o" unfolding safe_distance_2r_def by auto
hence "s\<^sub>o - v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto
hence "s\<^sub>o - v\<^sub>o\<^sup>2 / a\<^sub>o + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto
hence "s\<^sub>o + v\<^sub>o * (- v\<^sub>o / a\<^sub>o) + 1/2 * a\<^sub>o * (-v\<^sub>o / a\<^sub>o)\<^sup>2 > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e"
using other.p_def other.p_max_def other.p_max_eq other.t_stop_def by auto
hence "other.s_stop > u_max" unfolding other.s_def using u_max_eq other.t_stop_def
using ego.q_def other.p_def other.p_max_def other.s_def other.s_t_stop by auto
thus ?thesis
using assms(2) cond_1r cond_3r_3 by linarith
qed
lemma sd_3r_correct:
assumes "s\<^sub>o - s\<^sub>e > safe_distance_3r"
assumes "\<delta> \<le> other.t_stop"
shows "no_collision_react {0 ..}"
proof -
from assms have "s\<^sub>o - s\<^sub>e > v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2" unfolding safe_distance_3r_def by auto
hence "s\<^sub>o + v\<^sub>o * \<delta> + 1/2 * a\<^sub>o * \<delta>\<^sup>2 > s\<^sub>e + v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e" by auto
hence "other.s \<delta> > u_max" using other.s_def u_max_eq assms(2) ego.q_def other.p_def pos_react by auto
thus ?thesis using cond_1r cond_3r_3 delayed_other_s_stop_eq delayed_safe_distance.other.s0_le_s_stop by linarith
qed
lemma sd_2_at_least_sd_3:
assumes "\<delta> \<le> other.t_stop"
shows "safe_distance_3r \<ge> safe_distance_2r"
proof -
from assms have "\<delta> = other.t_stop \<or> \<delta> < other.t_stop" by auto
then have "safe_distance_3r = safe_distance_2r \<or> safe_distance_3r > safe_distance_2r"
proof (rule Meson.disj_forward)
assume "\<delta> = other.t_stop"
hence "\<delta> = - v\<^sub>o / a\<^sub>o" unfolding other.t_stop_def by auto
hence "- v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 = - v\<^sub>o * other.t_stop - 1/2 * a\<^sub>o * other.t_stop\<^sup>2" by (simp add: movement.t_stop_def)
thus "safe_distance_3r = safe_distance_2r"
using other.p_def other.p_max_def other.p_max_eq safe_distance_2r_def safe_distance_3r_def by auto
next
assume "\<delta> < other.t_stop"
hence "\<delta> < - v\<^sub>o / a\<^sub>o" unfolding other.t_stop_def by auto
hence "0 < v\<^sub>o + a\<^sub>o * \<delta>"
using other.decel other.p'_def other.p'_pos_iff by auto
hence "0 < v\<^sub>o + 1/2 * a\<^sub>o * (\<delta> + other.t_stop)" by (auto simp add:field_simps other.t_stop_def)
hence "0 > v\<^sub>o * (\<delta> - other.t_stop) + 1/2 * a\<^sub>o * (\<delta> + other.t_stop) * (\<delta> - other.t_stop)"
- using `\<delta> < other.t_stop` mult_less_cancel_left_neg[where c="\<delta> - other.t_stop" and a ="v\<^sub>o + 1 / 2 * a\<^sub>o * (\<delta> + other.t_stop)" and b="0"]
+ using \<open>\<delta> < other.t_stop\<close> mult_less_cancel_left_neg[where c="\<delta> - other.t_stop" and a ="v\<^sub>o + 1 / 2 * a\<^sub>o * (\<delta> + other.t_stop)" and b="0"]
by (auto simp add: field_simps)
hence " (\<delta> + other.t_stop) * (\<delta> - other.t_stop) = (\<delta>\<^sup>2 - other.t_stop\<^sup>2)"
by (simp add: power2_eq_square square_diff_square_factored)
hence "0 > v\<^sub>o * (\<delta> - other.t_stop) + 1/2 * a\<^sub>o * (\<delta>\<^sup>2 - other.t_stop\<^sup>2)"
by (metis (no_types, hide_lams) \<open>v\<^sub>o * (\<delta> - other.t_stop) + 1 / 2 * a\<^sub>o * (\<delta> + other.t_stop) * (\<delta> - other.t_stop) < 0\<close> divide_divide_eq_left divide_divide_eq_right times_divide_eq_left)
hence "0 > v\<^sub>o * \<delta> - v\<^sub>o * other.t_stop + 1/2 * a\<^sub>o * \<delta>\<^sup>2 - 1/2 * a\<^sub>o * other.t_stop\<^sup>2 "
by (simp add: right_diff_distrib)
hence "- v\<^sub>o * \<delta> - 1/2 * a\<^sub>o * \<delta>\<^sup>2 > - v\<^sub>o * (- v\<^sub>o / a\<^sub>o) - 1/2 * a\<^sub>o * (- v\<^sub>o / a\<^sub>o)\<^sup>2"
unfolding movement.t_stop_def by argo
thus "safe_distance_3r > safe_distance_2r"
using other.p_def other.p_max_def other.p_max_eq other.t_stop_def safe_distance_2r_def safe_distance_3r_def by auto
qed
thus ?thesis by auto
qed
end
subsection \<open>Checker Design\<close>
text \<open>
We define two checkers for different cases:
\<^item> one checker for the case that \<open>\<delta> \<le> other.t_stop (other.t_stop = - v\<^sub>o / a\<^sub>o)\<close>
\<^item> a second checker for the case that \<open>\<delta> > other.t_stop\<close>
\<close>
definition check_precond_r1 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> s\<^sub>o > s\<^sub>e \<and> 0 \<le> v\<^sub>e \<and> 0 \<le> v\<^sub>o \<and> a\<^sub>e < 0 \<and> a\<^sub>o < 0 \<and> 0 < \<delta> \<and> \<delta> \<le> - v\<^sub>o / a\<^sub>o"
definition safe_distance0 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = v\<^sub>e * \<delta> - v\<^sub>o * \<delta> - a\<^sub>o * \<delta>\<^sup>2 / 2"
definition safe_distance_1r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance_1r a\<^sub>e v\<^sub>e \<delta> = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / a\<^sub>e / 2"
definition safe_distance_2r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e + v\<^sub>o\<^sup>2 / 2 / a\<^sub>o"
definition safe_distance_4r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = (v\<^sub>o + a\<^sub>o * \<delta> - v\<^sub>e)\<^sup>2 / 2 / (a\<^sub>o - a\<^sub>e) - v\<^sub>o * \<delta> - 1 / 2 * a\<^sub>o * \<delta>\<^sup>2 + v\<^sub>e * \<delta>"
definition safe_distance_3r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = v\<^sub>e * \<delta> - v\<^sub>e\<^sup>2 / 2 / a\<^sub>e - v\<^sub>o * \<delta> - 1 / 2 * a\<^sub>o * \<delta>\<^sup>2"
definition checker_r1 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>;
vo_star = v\<^sub>o + a\<^sub>o * \<delta>;
t_stop_o_star = - vo_star / a\<^sub>o;
t_stop_e = - v\<^sub>e / a\<^sub>e;
safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>;
safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>;
safe_dist2 = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>;
safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>
in
if \<not> precond then False
else if distance > safe_dist0 \<or> distance > safe_dist3 then True
else if (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star) then distance > safe_dist2
else distance > safe_dist1"
theorem checker_r1_correctness:
"(checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and>
safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})"
proof
assume asm: "checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
have pre: "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
proof (rule ccontr)
assume "\<not> check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
with asm show "False" unfolding checker_r1_def Let_def by auto
qed
from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>"
by (unfold_locales) (auto simp add: check_precond_r1_def)
interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
rewrites "sdn.distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>" and
"sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_4r = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_3r = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
proof -
from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>" by auto
next
show "safe_distance_normal.distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> "
unfolding safe_distance_normal.distance0_def[OF sdn'] safe_distance0_def by auto
next
show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \<delta> = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto
next
show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto
next
show "safe_distance_normal.safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> "
unfolding safe_distance_normal.safe_distance_4r_def[OF sdn'] safe_distance_4r_def by auto
next
show "safe_distance_normal.safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_3r_def[OF sdn'] safe_distance_3r_def by auto
qed
have "0 < \<delta>" and "\<delta> \<le> - v\<^sub>o / a\<^sub>o" using pre unfolding check_precond_r1_def by auto
define so_delta where "so_delta = s\<^sub>o + v\<^sub>o * \<delta> + a\<^sub>o * \<delta>\<^sup>2 / 2"
define q_e_delta where "q_e_delta \<equiv> s\<^sub>e + v\<^sub>e * \<delta>"
define u_stop_e where "u_stop_e \<equiv> q_e_delta - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)"
define vo_star where "vo_star = v\<^sub>o + a\<^sub>o * \<delta>"
define t_stop_o_star where "t_stop_o_star \<equiv> - vo_star / a\<^sub>o"
define t_stop_e where "t_stop_e = - v\<^sub>e / a\<^sub>e"
define distance where "distance \<equiv> s\<^sub>o - s\<^sub>e"
define distance0 where "distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
define safe_dist2 where "safe_dist2 \<equiv> safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist1 where "safe_dist1 \<equiv> safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist3 where "safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
note abb = so_delta_def q_e_delta_def u_stop_e_def vo_star_def t_stop_o_star_def t_stop_e_def
distance_def safe_dist2_def safe_dist1_def safe_dist0_def safe_dist3_def distance0_def
consider "distance > safe_dist0" | "distance > safe_dist3" | "distance \<le> safe_dist0 \<and> distance \<le> safe_dist3"
by linarith
hence "sdn.no_collision_react {0..}"
proof (cases)
case 1
then show ?thesis using sdn.sd_1r_correct unfolding abb by auto
next
case 2
hence pre2: "distance > distance0" using sdn.distance0_at_most_sd3r unfolding abb by auto
hence "sdn.u \<delta> < sdn.other.s \<delta>" using pre unfolding sdn.u_def sdn.ego.q_def
sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def
by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r1_def `sdn.u \<delta> < sdn.other.s \<delta>`)
+ by (unfold_locales) (auto simp add:check_precond_r1_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
show ?thesis using sdr.sd_3r_correct 2 pre unfolding check_precond_r1_def abb sdn.other.t_stop_def
by auto
next
case 3
hence "distance \<le> safe_dist3" by auto
hence "sdn.other.s \<delta> \<le> sdn.u_max" using pre unfolding check_precond_r1_def sdn.other.s_def sdn.other.t_stop_def
sdn.other.p_def sdn.u_max_eq sdn.ego.q_def abb sdn.safe_distance_3r_def by auto
have " (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star) \<or> \<not> (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star) "
by auto
moreover
{ assume cond: "(a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star)"
with 3 pre have "distance > safe_dist2" using asm unfolding checker_r1_def
Let_def abb by auto
with sdn.distance0_at_most_sd4r have "distance > distance0" unfolding abb using cond by auto
hence "sdn.u \<delta> < sdn.other.s \<delta>" using pre unfolding sdn.u_def sdn.ego.q_def
sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def
by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r1_def `sdn.u \<delta> < sdn.other.s \<delta>`)
- from sdr.sd_4r_correct[OF _ `sdn.other.s \<delta> \<le> sdn.u_max`] `distance > safe_dist2`
+ by (unfold_locales) (auto simp add:check_precond_r1_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
+ from sdr.sd_4r_correct[OF _ \<open>sdn.other.s \<delta> \<le> sdn.u_max\<close>] \<open>distance > safe_dist2\<close>
have ?thesis using pre cond unfolding check_precond_r1_def sdn.other.t_stop_def abb by auto }
moreover
{ assume not_cond: "\<not> (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star)"
with 3 pre have "distance > safe_dist1" using asm unfolding checker_r1_def
Let_def abb by auto
with sdn.dist0_sd2r_1 have "distance > distance0" using pre not_cond unfolding check_precond_r1_def
sdn.other.t_stop_def sdn.other.s'_def sdn.other.p'_def abb by (auto simp add:field_simps)
hence "sdn.u \<delta> < sdn.other.s \<delta>" using pre unfolding sdn.u_def sdn.ego.q_def
sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def
by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r1_def `sdn.u \<delta> < sdn.other.s \<delta>`)
- from sdr.sd_2r_correct_for_3r_2[OF _ `sdn.other.s \<delta> \<le> sdn.u_max`] not_cond `distance > safe_dist1`
+ by (unfold_locales) (auto simp add:check_precond_r1_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
+ from sdr.sd_2r_correct_for_3r_2[OF _ \<open>sdn.other.s \<delta> \<le> sdn.u_max\<close>] not_cond \<open>distance > safe_dist1\<close>
have ?thesis using pre unfolding abb sdn.other.s'_def check_precond_r1_def sdn.other.t_stop_def sdn.other.p'_def
by (auto simp add:field_simps) }
ultimately show ?thesis by auto
qed
with pre show " check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> sdn.no_collision_react {0..}" by auto
next
assume "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..}"
hence pre: "check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>" and as2: "safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..}"
by auto
show "checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> "
proof (rule ccontr)
assume as1: "\<not> checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
from pre have "0 < \<delta>" and "\<delta> \<le> - v\<^sub>o / a\<^sub>o" unfolding check_precond_r1_def by auto
define so_delta where "so_delta = s\<^sub>o + v\<^sub>o * \<delta> + a\<^sub>o * \<delta>\<^sup>2 / 2"
define q_e_delta where "q_e_delta \<equiv> s\<^sub>e + v\<^sub>e * \<delta>"
define u_stop_e where "u_stop_e \<equiv> q_e_delta - v\<^sub>e\<^sup>2 / (2 * a\<^sub>e)"
define vo_star where "vo_star \<equiv> v\<^sub>o + a\<^sub>o * \<delta>"
define t_stop_o_star where "t_stop_o_star \<equiv> - vo_star / a\<^sub>o"
define t_stop_e where "t_stop_e \<equiv> - v\<^sub>e / a\<^sub>e"
define distance where "distance \<equiv> s\<^sub>o - s\<^sub>e"
define distance0 where "distance0 \<equiv> safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist0 where "safe_dist0 \<equiv> safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
define safe_dist2 where "safe_dist2 \<equiv> safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist1 where "safe_dist1 \<equiv> safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist3 where "safe_dist3 \<equiv> safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
note abb = so_delta_def q_e_delta_def u_stop_e_def vo_star_def t_stop_o_star_def t_stop_e_def
distance_def safe_dist2_def safe_dist1_def safe_dist0_def safe_dist3_def distance0_def
from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>"
by (unfold_locales) (auto simp add: check_precond_r1_def)
interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
rewrites "sdn.distance0 = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>" and
"sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_4r = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_3r = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
proof -
from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>" by auto
next
show "safe_distance_normal.distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance0 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> "
unfolding safe_distance_normal.distance0_def[OF sdn'] safe_distance0_def by auto
next
show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \<delta> = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto
next
show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto
next
show "safe_distance_normal.safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> "
unfolding safe_distance_normal.safe_distance_4r_def[OF sdn'] safe_distance_4r_def by auto
next
show "safe_distance_normal.safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_3r_def[OF sdn'] safe_distance_3r_def by auto
qed
have "\<not> distance > distance0 \<or> distance > distance0" by auto
moreover
{ assume "\<not> distance > distance0"
hence "distance \<le> distance0" by auto
with sdn.cond_3r_1' have "sdn.collision_react {0..\<delta>}" using pre unfolding check_precond_r1_def abb
sdn.other.t_stop_def by auto
with sdn.collision_react_subset have "sdn.collision_react {0..}" by auto
with as2 have "False" by auto }
moreover
{ assume if2: "distance > distance0"
have "\<not> (distance > safe_dist0 \<or> distance > safe_dist3)"
proof (rule ccontr)
assume "\<not> \<not> (safe_dist0 < distance \<or> safe_dist3 < distance)"
hence "(safe_dist0 < distance \<or> safe_dist3 < distance)" by auto
with as1 show "False" using pre if2 unfolding checker_r1_def Let_def abb
by auto
qed
hence if31: "distance \<le> safe_dist0" and if32: "distance \<le> safe_dist3" by auto
have "sdn.u \<delta> < sdn.other.s \<delta>" using if2 pre unfolding sdn.u_def sdn.ego.q_def
sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def abb check_precond_r1_def sdn.distance0_def
by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r1_def `sdn.u \<delta> < sdn.other.s \<delta>`)
+ by (unfold_locales) (auto simp add:check_precond_r1_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
have " s\<^sub>o \<le> sdn.u_max" using if31 unfolding sdn.u_max_eq sdn.ego.q_def abb
sdn.safe_distance_1r_def by auto
have "sdn.other.s \<delta> \<le> sdn.u_max" using if32 pre unfolding sdn.other.s_def check_precond_r1_def
sdn.other.t_stop_def sdn.other.p_def sdn.u_max_eq sdn.ego.q_def abb sdn.safe_distance_3r_def
by auto
consider "(a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star)" |
"\<not> (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star)" by auto
hence "False"
proof (cases)
case 1
hence rest_conjunct:"(a\<^sub>e < a\<^sub>o \<and> sdn.other.s' \<delta> < v\<^sub>e \<and> v\<^sub>e - a\<^sub>e / a\<^sub>o * sdn.other.s' \<delta> < 0)"
using pre unfolding check_precond_r1_def unfolding sdn.other.s'_def sdn.other.t_stop_def
sdn.other.p'_def abb by (auto simp add:field_simps)
from 1 have "distance \<le> safe_dist2" using as1 pre if2 if31 if32 unfolding checker_r1_def
Let_def abb by auto
hence cond_f: "sdn.other.s \<delta> - sdn.ego.q \<delta> \<le> sdr.delayed_safe_distance.snd_safe_distance"
using pre unfolding check_precond_r1_def sdn.other.s_def sdn.other.t_stop_def sdn.other.p_def
sdn.ego.q_def sdr.delayed_safe_distance.snd_safe_distance_def using sdn.other.s'_def[of "\<delta>"]
unfolding sdn.other.t_stop_def sdn.other.p'_def abb sdn.safe_distance_4r_def
by auto
have "distance > safe_dist1 \<or> distance \<le> safe_dist1" by auto
moreover
{ assume "distance > safe_dist1"
hence "sdn.u_max < sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop
sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps)
- from sdr.cond_3r_2[OF `s\<^sub>o \<le> sdn.u_max` this `sdn.other.s \<delta> \<le> sdn.u_max`]
+ from sdr.cond_3r_2[OF \<open>s\<^sub>o \<le> sdn.u_max\<close> this \<open>sdn.other.s \<delta> \<le> sdn.u_max\<close>]
have "sdn.collision_react {0..}" using cond_f rest_conjunct by auto
with as2 have "False" by auto }
moreover
{ assume "distance \<le> safe_dist1"
hence "sdn.u_max \<ge> sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop
sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps)
with sdn.cond_2r[OF this] have "sdn.collision_react {0..}" by auto
with as2 have "False" by auto }
ultimately show ?thesis by auto
next
case 2
hence "distance \<le> safe_dist1" using as1 pre if2 if31 if32 unfolding checker_r1_def
Let_def abb by auto
hence "sdn.u_max \<ge> sdn.other.s_stop" unfolding sdn.u_max_eq sdn.ego.q_def sdn.other.s_t_stop
sdn.other.p_max_eq abb sdn.safe_distance_2r_def by (auto simp add:field_simps)
with sdn.cond_2r[OF this] have "sdn.collision_react {0..}" by auto
with as2 show "False" by auto
qed }
ultimately show "False" by auto
qed
qed
definition check_precond_r2 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> s\<^sub>o > s\<^sub>e \<and> 0 \<le> v\<^sub>e \<and> 0 \<le> v\<^sub>o \<and> a\<^sub>e < 0 \<and> a\<^sub>o < 0 \<and> 0 < \<delta> \<and> \<delta> > - v\<^sub>o / a\<^sub>o"
definition safe_distance0_2 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
"safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = v\<^sub>e * \<delta> + 1 / 2 * v\<^sub>o\<^sup>2 / a\<^sub>o"
definition checker_r2 :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>;
safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>;
safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>
in
if \<not> precond then False
else if distance > safe_dist0 then True
else distance > safe_dist1"
theorem checker_r2_correctness:
"(checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and>
safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})"
proof
assume asm: "checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
have pre: "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
proof (rule ccontr)
assume "\<not> check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
with asm show "False" unfolding checker_r2_def Let_def by auto
qed
from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>"
by (unfold_locales) (auto simp add: check_precond_r2_def)
interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
rewrites "sdn.distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>" and
"sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
proof -
from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>" by auto
next
show "safe_distance_normal.distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.distance0_2_def[OF sdn'] safe_distance0_2_def by auto
next
show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \<delta> = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto
next
show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto
qed
have "0 < \<delta>" and "\<delta> > - v\<^sub>o / a\<^sub>o" using pre unfolding check_precond_r2_def by auto
define distance where "distance \<equiv> s\<^sub>o - s\<^sub>e"
define distance0_2 where "distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
define safe_dist1 where "safe_dist1 \<equiv> safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
note abb = distance_def safe_dist1_def safe_dist0_def distance0_2_def
consider "distance > safe_dist0" | "distance \<le> safe_dist0"
by linarith
hence "sdn.no_collision_react {0..}"
proof (cases)
case 1
then show ?thesis using sdn.sd_1r_correct unfolding abb by auto
next
case 2
hence "(s\<^sub>o \<le> sdn.u_max)" using distance_def safe_dist0_def sdn.sd_1r_eq by linarith
with 2 pre have "distance > safe_dist1" using asm unfolding checker_r2_def Let_def abb by auto
with sdn.dist0_sd2r_2 have "distance > distance0_2" using abb \<open>- v\<^sub>o / a\<^sub>o < \<delta>\<close> by auto
hence "sdn.u \<delta> < sdn.other.s \<delta>" using abb sdn.distance0_2_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def by auto
- have "sdn.u_max < sdn.other.s \<delta>" using abb sdn.sd2r_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def `distance > safe_dist1` by auto
+ have "sdn.u_max < sdn.other.s \<delta>" using abb sdn.sd2r_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def \<open>distance > safe_dist1\<close> by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r2_def `sdn.u \<delta> < sdn.other.s \<delta>`)
- from sdr.sd_2r_correct_for_3r_3[OF] `distance > safe_dist1` `sdn.u \<delta> < sdn.other.s \<delta>` `sdn.u_max < sdn.other.s \<delta>`
+ by (unfold_locales) (auto simp add:check_precond_r2_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
+ from sdr.sd_2r_correct_for_3r_3[OF] \<open>distance > safe_dist1\<close> \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close> \<open>sdn.u_max < sdn.other.s \<delta>\<close>
show ?thesis using pre unfolding abb sdn.other.s'_def check_precond_r2_def sdn.other.t_stop_def sdn.other.p'_def
by (auto simp add:field_simps)
qed
with pre show " check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> sdn.no_collision_react {0..}" by auto
next
assume "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..}"
hence pre: "check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>" and as2: "safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..}"
by auto
show "checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
proof (rule ccontr)
assume as1: "\<not> checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
from pre have "0 < \<delta>" and "\<delta> > - v\<^sub>o / a\<^sub>o" unfolding check_precond_r2_def by auto
define distance where "distance \<equiv> s\<^sub>o - s\<^sub>e"
define distance0_2 where "distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
define safe_dist0 where "safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
define safe_dist1 where "safe_dist1 \<equiv> safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
note abb = distance_def safe_dist1_def safe_dist0_def distance0_2_def
from pre have sdn': "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>"
by (unfold_locales) (auto simp add: check_precond_r2_def)
interpret sdn: safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
rewrites "sdn.distance0_2 = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>" and
"sdn.safe_distance_1r = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>" and
"sdn.safe_distance_2r = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
proof -
from sdn' show "safe_distance_normal a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>" by auto
next
show "safe_distance_normal.distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance0_2 v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.distance0_2_def[OF sdn'] safe_distance0_2_def by auto
next
show "safe_distance_normal.safe_distance_1r a\<^sub>e v\<^sub>e \<delta> = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>"
unfolding safe_distance_normal.safe_distance_1r_def[OF sdn'] safe_distance_1r_def by auto
next
show "safe_distance_normal.safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta> = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>"
unfolding safe_distance_normal.safe_distance_2r_def[OF sdn'] safe_distance_2r_def by auto
qed
have "\<not> distance > distance0_2 \<or> distance > distance0_2" by auto
moreover
{ assume "\<not> distance > distance0_2"
hence "distance \<le> distance0_2" by auto
with sdn.cond_3r_1'_2 have "sdn.collision_react {0..\<delta>}" using pre unfolding check_precond_r2_def abb sdn.other.t_stop_def by auto
with sdn.collision_react_subset have "sdn.collision_react {0..}" by auto
with as2 have "False" by auto }
moreover
{ assume if2: "distance > distance0_2"
have "\<not> (distance > safe_dist0)"
proof (rule ccontr)
assume "\<not> \<not> (safe_dist0 < distance)"
hence "(safe_dist0 < distance)" by auto
with as1 show "False" using pre if2 unfolding checker_r2_def Let_def abb by auto
qed
hence if3: "distance \<le> safe_dist0" by auto
with pre have "distance \<le> safe_dist1" using as1 unfolding checker_r2_def Let_def abb by auto
have "sdn.u \<delta> < sdn.other.s \<delta>" using abb if2 sdn.distance0_2_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def by auto
from pre interpret sdr: safe_distance_no_collsion_delta a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta>
- by (unfold_locales) (auto simp add:check_precond_r2_def `sdn.u \<delta> < sdn.other.s \<delta>`)
- have "sdn.u_max \<ge> sdn.other.s \<delta>" using abb sdn.sd2r_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def `distance \<le> safe_dist1` by auto
- with `\<delta> > - v\<^sub>o / a\<^sub>o` have "sdn.u_max \<ge> sdn.other.s_stop"
+ by (unfold_locales) (auto simp add:check_precond_r2_def \<open>sdn.u \<delta> < sdn.other.s \<delta>\<close>)
+ have "sdn.u_max \<ge> sdn.other.s \<delta>" using abb sdn.sd2r_eq \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> sdn.other.t_stop_def \<open>distance \<le> safe_dist1\<close> by auto
+ with \<open>\<delta> > - v\<^sub>o / a\<^sub>o\<close> have "sdn.u_max \<ge> sdn.other.s_stop"
using sdn.other.s_mono sdn.other.t_stop_nonneg sdn.other.p_t_stop sdn.other.p_zero sdn.other.t_stop_def
apply (auto simp: sdn.other.s_def movement.t_stop_def split: if_splits)
using sdn.other.p_zero by auto
hence "sdn.collision_react {0..}" using sdn.cond_2r by auto
with as2 have "False" by auto }
ultimately show "False" by auto
qed
qed
text \<open>Combine the two checkers into one.\<close>
definition check_precond_r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> s\<^sub>o > s\<^sub>e \<and> 0 \<le> v\<^sub>e \<and> 0 \<le> v\<^sub>o \<and> a\<^sub>e < 0 \<and> a\<^sub>o < 0 \<and> 0 < \<delta>"
definition checker_r :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> bool" where
"checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<equiv>
let distance = s\<^sub>o - s\<^sub>e;
precond = check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>;
vo_star = v\<^sub>o + a\<^sub>o * \<delta>;
t_stop_o_star = -vo_star / a\<^sub>o;
t_stop_e = -v\<^sub>e / a\<^sub>e;
t_stop_o = -v\<^sub>o / a\<^sub>o;
safe_dist0 = safe_distance_1r a\<^sub>e v\<^sub>e \<delta>;
safe_dist1 = safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>;
safe_dist2 = safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>;
safe_dist3 = safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>
in
if \<not> precond then False
else if distance > safe_dist0 then True
else if \<delta> \<le> t_stop_o \<and> distance > safe_dist3 then True
else if \<delta> \<le> t_stop_o \<and> (a\<^sub>o > a\<^sub>e \<and> vo_star < v\<^sub>e \<and> t_stop_e < t_stop_o_star) then distance > safe_dist2
else distance > safe_dist1"
theorem checker_eq_1:
"checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> \<le> - v\<^sub>o / a\<^sub>o \<longleftrightarrow> checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
proof -
have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> \<le> - v\<^sub>o / a\<^sub>o \<longleftrightarrow> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>
\<and> (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \<delta>
\<or> s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>
\<or> (((a\<^sub>o > a\<^sub>e \<and> v\<^sub>o + a\<^sub>o * \<delta> < v\<^sub>e \<and> - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o) \<longrightarrow> s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)
\<and> (\<not> (a\<^sub>o > a\<^sub>e \<and> v\<^sub>o + a\<^sub>o * \<delta> < v\<^sub>e \<and> - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o) \<longrightarrow> s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)))
\<and> \<delta> \<le> - v\<^sub>o / a\<^sub>o" using checker_r_def by metis
also have "... \<longleftrightarrow> check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>
\<and> (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \<delta>
\<or> s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>
\<or> (((a\<^sub>o > a\<^sub>e \<and> v\<^sub>o + a\<^sub>o * \<delta> < v\<^sub>e \<and> - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o) \<longrightarrow> s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)
\<and> (\<not> (a\<^sub>o > a\<^sub>e \<and> v\<^sub>o + a\<^sub>o * \<delta> < v\<^sub>e \<and> - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o) \<longrightarrow> s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)))"
by (auto simp add:check_precond_r_def check_precond_r1_def)
also have "... \<longleftrightarrow> checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>" by (metis checker_r1_def)
finally show ?thesis by auto
qed
theorem checker_eq_2:
"checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> > - v\<^sub>o / a\<^sub>o \<longleftrightarrow> checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>"
proof -
have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> > - v\<^sub>o / a\<^sub>o \<longleftrightarrow> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> (\<not> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<or>
s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \<delta> \<or>
(\<delta> \<le> - v\<^sub>o / a\<^sub>o \<and> s\<^sub>o - s\<^sub>e > safe_distance_3r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>) \<or>
(\<delta> \<le> - v\<^sub>o / a\<^sub>o \<and> a\<^sub>o > a\<^sub>e \<and> v\<^sub>o + a\<^sub>o * \<delta> < v\<^sub>e \<and> - v\<^sub>e / a\<^sub>e < - (v\<^sub>o + a\<^sub>o * \<delta>) / a\<^sub>o \<and> s\<^sub>o - s\<^sub>e > safe_distance_4r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>) \<or>
s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>) \<and> \<delta> > - v\<^sub>o / a\<^sub>o" unfolding checker_r_def Let_def if_splits by auto
also have
"... \<longleftrightarrow> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>
\<and> (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \<delta> \<or> s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)
\<and> \<delta> > - v\<^sub>o / a\<^sub>o" by (auto simp add:HOL.disjE)
also have
"... \<longleftrightarrow> check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>
\<and> (s\<^sub>o - s\<^sub>e > safe_distance_1r a\<^sub>e v\<^sub>e \<delta> \<or> s\<^sub>o - s\<^sub>e > safe_distance_2r a\<^sub>e v\<^sub>e a\<^sub>o v\<^sub>o \<delta>)"
by (auto simp add:check_precond_r_def check_precond_r2_def)
also have "... \<longleftrightarrow> checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>" by (auto simp add:checker_r2_def Let_def if_splits)
thus ?thesis using calculation by auto
qed
theorem checker_r_correctness:
"(checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})"
proof -
have "checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<longleftrightarrow> (checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> \<le> - v\<^sub>o / a\<^sub>o) \<or> (checker_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> \<delta> > - v\<^sub>o / a\<^sub>o)" by auto
also have "... \<longleftrightarrow> checker_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<or> checker_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta>" using checker_eq_1 checker_eq_2 by auto
also have "... \<longleftrightarrow> (check_precond_r1 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})
\<or> (check_precond_r2 s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})"
using checker_r1_correctness checker_r2_correctness by auto
also have "... \<longleftrightarrow> (\<delta> \<le> - v\<^sub>o / a\<^sub>o \<and> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})
\<or> (\<delta> > - v\<^sub>o / a\<^sub>o \<and> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..})"
by (auto simp add:check_precond_r_def check_precond_r1_def check_precond_r2_def)
also have "... \<longleftrightarrow> check_precond_r s\<^sub>e v\<^sub>e a\<^sub>e s\<^sub>o v\<^sub>o a\<^sub>o \<delta> \<and> safe_distance_normal.no_collision_react a\<^sub>e v\<^sub>e s\<^sub>e a\<^sub>o v\<^sub>o s\<^sub>o \<delta> {0..}"
by auto
finally show ?thesis by auto
qed
end
\ No newline at end of file
diff --git a/thys/Syntax_Independent_Logic/Deduction.thy b/thys/Syntax_Independent_Logic/Deduction.thy
--- a/thys/Syntax_Independent_Logic/Deduction.thy
+++ b/thys/Syntax_Independent_Logic/Deduction.thy
@@ -1,2136 +1,2136 @@
chapter \<open>Deduction\<close>
(*<*)
theory Deduction imports Syntax
begin
(*>*)
text \<open>We formalize deduction in a logical system that (shallowly) embeds
intuitionistic logic connectives and quantifiers over a signature containing
the numerals.\<close>
section \<open>Positive Logic Deduction\<close>
locale Deduct =
Syntax_with_Numerals_and_Connectives
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and num
and eql cnj imp all exi
+
fixes
\<comment> \<open>Provability of numeric formulas:\<close>
prv :: "'fmla \<Rightarrow> bool"
\<comment> \<open>Hilbert-style system for intuitionistic logic over $\longrightarrow$,$\land$,$\forall$,$\exists$.
($\bot$, $\lnot$ and $\lor$ will be included later.)
Hilbert-style is preferred since it requires the least amount of infrastructure.
(Later, natural deduction rules will also be defined.)\<close>
assumes
\<comment> \<open>Propositional rules and axioms. There is a single propositional rule, modus ponens.\<close>
\<comment> \<open>The modus ponens rule:\<close>
prv_imp_mp:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp \<phi> \<chi>) \<Longrightarrow> prv \<phi> \<Longrightarrow> prv \<chi>"
and
\<comment> \<open>The propositional intuitionitic axioms:\<close>
prv_imp_imp_triv:
"\<And>\<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp \<phi> (imp \<chi> \<phi>))"
and
prv_imp_trans:
"\<And> \<phi> \<chi> \<psi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
prv (imp (imp \<phi> (imp \<chi> \<psi>))
(imp (imp \<phi> \<chi>) (imp \<phi> \<psi>)))"
and
prv_imp_cnjL:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp (cnj \<phi> \<chi>) \<phi>)"
and
prv_imp_cnjR:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp (cnj \<phi> \<chi>) \<chi>)"
and
prv_imp_cnjI:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp \<phi> (imp \<chi> (cnj \<phi> \<chi>)))"
and
(* *)
\<comment> \<open>Predicate calculus (quantifier) rules and axioms\<close>
\<comment> \<open>The rules of universal and existential generalization:\<close>
prv_all_imp_gen:
"\<And> x \<phi> \<chi>. x \<notin> Fvars \<phi> \<Longrightarrow> prv (imp \<phi> \<chi>) \<Longrightarrow> prv (imp \<phi> (all x \<chi>))"
and
prv_exi_imp_gen:
"\<And> x \<phi> \<chi>. x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
x \<notin> Fvars \<chi> \<Longrightarrow> prv (imp \<phi> \<chi>) \<Longrightarrow> prv (imp (exi x \<phi>) \<chi>)"
and
\<comment> \<open>Two quantifier instantiation axioms:\<close>
prv_all_inst:
"\<And> x \<phi> t.
x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow>
prv (imp (all x \<phi>) (subst \<phi> t x))"
and
prv_exi_inst:
"\<And> x \<phi> t.
x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow>
prv (imp (subst \<phi> t x) (exi x \<phi>))"
and
\<comment> \<open>The equality axioms:\<close>
prv_eql_refl:
"\<And> x. x \<in> var \<Longrightarrow>
prv (eql (Var x) (Var x))"
and
prv_eql_subst:
"\<And> \<phi> x y.
x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow>
prv ((imp (eql (Var x) (Var y))
(imp \<phi> (subst \<phi> (Var y) x))))"
begin
subsection \<open>Properties of the propositional fragment\<close>
lemma prv_imp_triv:
assumes phi: "\<phi> \<in> fmla" and psi: "\<psi> \<in> fmla"
shows "prv \<psi> \<Longrightarrow> prv (imp \<phi> \<psi>)"
by (meson prv_imp_imp_triv prv_imp_mp imp phi psi)
lemma prv_imp_refl:
assumes phi: "\<phi> \<in> fmla"
shows "prv (imp \<phi> \<phi>)"
by (metis prv_imp_imp_triv prv_imp_mp prv_imp_trans imp phi)
lemma prv_imp_refl2: "\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> \<phi> = \<psi> \<Longrightarrow> prv (imp \<phi> \<psi>)"
using prv_imp_refl by auto
lemma prv_cnjI:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla"
shows "prv \<phi> \<Longrightarrow> prv \<chi> \<Longrightarrow> prv (cnj \<phi> \<chi>)"
by (meson cnj prv_imp_cnjI prv_imp_mp imp phi chi)
lemma prv_cnjEL:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla"
shows "prv (cnj \<phi> \<chi>) \<Longrightarrow> prv \<phi>"
using chi phi prv_imp_cnjL prv_imp_mp by blast
lemma prv_cnjER:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla"
shows "prv (cnj \<phi> \<chi>) \<Longrightarrow> prv \<chi>"
using chi phi prv_imp_cnjR prv_imp_mp by blast
lemma prv_prv_imp_trans:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla" and psi: "\<psi> \<in> fmla"
assumes 1: "prv (imp \<phi> \<chi>)" and 2: "prv (imp \<chi> \<psi>)"
shows "prv (imp \<phi> \<psi>)"
proof-
have "prv (imp \<phi> (imp \<chi> \<psi>))" by (simp add: 2 chi prv_imp_triv phi psi)
thus ?thesis by (metis 1 chi prv_imp_mp prv_imp_trans imp phi psi)
qed
lemma prv_imp_trans1:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla" and psi: "\<psi> \<in> fmla"
shows "prv (imp (imp \<chi> \<psi>) (imp (imp \<phi> \<chi>) (imp \<phi> \<psi>)))"
by (meson chi prv_prv_imp_trans prv_imp_imp_triv prv_imp_trans imp phi psi)
lemma prv_imp_com:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla" and psi: "\<psi> \<in> fmla"
assumes "prv (imp \<phi> (imp \<chi> \<psi>))"
shows "prv (imp \<chi> (imp \<phi> \<psi>))"
by (metis (no_types) assms prv_prv_imp_trans prv_imp_imp_triv prv_imp_mp prv_imp_trans imp)
lemma prv_imp_trans2:
assumes phi: "\<phi> \<in> fmla" and chi: "\<chi> \<in> fmla" and psi: "\<psi> \<in> fmla"
shows "prv (imp (imp \<phi> \<chi>) (imp (imp \<chi> \<psi>) (imp \<phi> \<psi>)))"
using prv_imp_mp prv_imp_trans prv_imp_trans1 prv_imp_imp_triv
by (meson chi prv_imp_com imp phi psi)
lemma prv_imp_cnj:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (imp \<phi> \<psi>) \<Longrightarrow> prv (imp \<phi> \<chi>) \<Longrightarrow> prv (imp \<phi> (cnj \<psi> \<chi>))"
proof -
assume "prv (imp \<phi> \<psi>)"
moreover
assume "prv (imp \<phi> \<chi>)"
then have "prv (imp \<phi> (imp \<psi> f))" if "prv (imp \<chi> f)" "f \<in> fmla" for f
using that by (metis (no_types) assms imp prv_imp_imp_triv prv_prv_imp_trans)
moreover have "prv (imp \<phi> (imp \<psi> \<psi>)) \<Longrightarrow> prv (imp \<phi> (imp \<phi> \<psi>))"
using \<open>prv (imp \<phi> \<psi>)\<close> by (metis (no_types) assms(1,3) imp prv_imp_com prv_prv_imp_trans)
ultimately show ?thesis
by (metis (no_types) assms cnj imp prv_imp_cnjI prv_imp_com prv_imp_mp prv_imp_trans)
qed
lemma prv_imp_imp_com:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows
"prv (imp (imp \<phi> (imp \<chi> \<psi>))
(imp \<chi> (imp \<phi> \<psi>)))"
by (metis (no_types) assms
prv_prv_imp_trans prv_imp_com prv_imp_imp_triv prv_imp_trans imp)
lemma prv_cnj_imp_monoR2:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
assumes "prv (imp \<phi> (imp \<chi> \<psi>))"
shows "prv (imp (cnj \<phi> \<chi>) \<psi>)"
proof -
have "prv (imp (cnj \<phi> \<chi>) (cnj \<phi> \<chi>))"
using prv_imp_refl by (blast intro: assms(1-3))
then have "prv (imp (imp (cnj \<phi> \<chi>) (imp (cnj \<phi> \<chi>) \<psi>)) (imp (cnj \<phi> \<chi>) \<psi>))"
by (metis (no_types) cnj imp assms(1-3) prv_imp_com prv_imp_mp prv_imp_trans)
then show ?thesis
by (metis (no_types) imp cnj assms prv_imp_cnjL prv_imp_cnjR prv_imp_com prv_imp_mp prv_prv_imp_trans)
qed
lemma prv_imp_imp_imp_cnj:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows
"prv (imp (imp \<phi> (imp \<chi> \<psi>))
(imp (cnj \<phi> \<chi>) \<psi>))"
proof-
have "prv (imp \<phi> (imp (imp \<phi> (imp \<chi> \<psi>)) (imp \<chi> \<psi>)))"
by (simp add: assms prv_imp_com prv_imp_refl)
hence "prv (imp \<phi> (imp \<chi> (imp (imp \<phi> (imp \<chi> \<psi>)) \<psi>)))"
by (metis (no_types, lifting) assms prv_prv_imp_trans prv_imp_imp_com imp)
hence "prv (imp (cnj \<phi> \<chi>)
(imp (imp \<phi> (imp \<chi> \<psi>)) \<psi>))"
by (simp add: assms prv_cnj_imp_monoR2)
thus ?thesis using assms prv_imp_com prv_imp_mp by (meson cnj imp)
qed
lemma prv_imp_cnj_imp:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows
"prv (imp (imp (cnj \<phi> \<chi>) \<psi>)
(imp \<phi> (imp \<chi> \<psi>)))"
by (metis (no_types) assms cnj prv_prv_imp_trans prv_imp_cnjI prv_imp_com prv_imp_trans2 imp)
lemma prv_cnj_imp:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
assumes "prv (imp (cnj \<phi> \<chi>) \<psi>)"
shows "prv (imp \<phi> (imp \<chi> \<psi>))"
using assms prv_imp_cnj_imp prv_imp_mp by (meson cnj imp)
text \<open>Monotonicy of conjunction w.r.t. implication:\<close>
lemma prv_cnj_imp_monoR:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (imp (imp \<phi> \<chi>) (imp (imp \<phi> \<psi>) (imp \<phi> (cnj \<chi> \<psi>))))"
by (meson assms cnj imp prv_cnj_imp prv_cnj_imp_monoR2 prv_imp_cnj prv_imp_cnjL prv_imp_cnjR)
lemma prv_imp_cnj3L:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (imp \<phi>1 \<chi>) (imp (cnj \<phi>1 \<phi>2) \<chi>))"
using assms prv_imp_cnjL prv_imp_mp prv_imp_trans2
by (metis cnj imp)
lemma prv_imp_cnj3R:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (imp \<phi>2 \<chi>) (imp (cnj \<phi>1 \<phi>2) \<chi>))"
using prv_imp_cnjR prv_imp_mp prv_imp_trans2
by (metis assms cnj imp)
lemma prv_cnj_imp_mono:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (imp \<phi>1 \<chi>1) (imp (imp \<phi>2 \<chi>2) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
proof-
have "prv (imp (imp (cnj \<phi>1 \<phi>2) \<chi>1) (imp (imp (cnj \<phi>1 \<phi>2) \<chi>2) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
using prv_cnj_imp_monoR[of "cnj \<phi>1 \<phi>2" \<chi>1 \<chi>2] assms by auto
hence "prv (imp (imp \<phi>1 \<chi>1) (imp (imp (cnj \<phi>1 \<phi>2) \<chi>2) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
by (metis (no_types) imp cnj assms prv_imp_cnj3L prv_prv_imp_trans)
hence "prv (imp (imp (cnj \<phi>1 \<phi>2) \<chi>2) (imp (imp \<phi>1 \<chi>1) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
using prv_imp_com assms by (meson cnj imp)
hence "prv (imp (imp \<phi>2 \<chi>2) (imp (imp \<phi>1 \<chi>1) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
using prv_imp_cnj3R prv_imp_mp prv_imp_trans1
by (metis (no_types) assms cnj prv_prv_imp_trans imp)
thus ?thesis using prv_imp_com assms
by (meson cnj imp)
qed
lemma prv_cnj_mono:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (imp \<phi>1 \<chi>1)" and "prv (imp \<phi>2 \<chi>2)"
shows "prv (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))"
using assms prv_cnj_imp_mono prv_imp_mp
by (metis (mono_tags) cnj prv_prv_imp_trans prv_imp_cnj prv_imp_cnjL prv_imp_cnjR)
lemma prv_cnj_imp_monoR4:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi>1 \<in> fmla" and "\<psi>2 \<in> fmla"
shows
"prv (imp (imp \<phi> (imp \<chi> \<psi>1))
(imp (imp \<phi> (imp \<chi> \<psi>2)) (imp \<phi> (imp \<chi> (cnj \<psi>1 \<psi>2)))))"
by (simp add: assms prv_cnj_imp prv_imp_cnj prv_imp_cnjL prv_imp_cnjR prv_cnj_imp_monoR2)
lemma prv_imp_cnj4:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi>1 \<in> fmla" and "\<psi>2 \<in> fmla"
shows "prv (imp \<phi> (imp \<chi> \<psi>1)) \<Longrightarrow> prv (imp \<phi> (imp \<chi> \<psi>2)) \<Longrightarrow> prv (imp \<phi> (imp \<chi> (cnj \<psi>1 \<psi>2)))"
by (simp add: assms prv_cnj_imp prv_imp_cnj prv_cnj_imp_monoR2)
lemma prv_cnj_imp_monoR5:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (imp \<phi>1 \<chi>1) (imp (imp \<phi>2 \<chi>2) (imp \<phi>1 (imp \<phi>2 (cnj \<chi>1 \<chi>2)))))"
proof-
have "prv (imp (imp \<phi>1 \<chi>1) (imp (imp \<phi>2 \<chi>2) (imp (cnj \<phi>1 \<phi>2) (cnj \<chi>1 \<chi>2))))"
using prv_cnj_imp_mono[of \<phi>1 \<phi>2 \<chi>1 \<chi>2] assms by auto
hence "prv (imp (imp \<phi>1 \<chi>1) (imp (cnj \<phi>1 \<phi>2) (imp (imp \<phi>2 \<chi>2) (cnj \<chi>1 \<chi>2))))"
by (metis (no_types, lifting) assms cnj imp prv_imp_imp_com prv_prv_imp_trans)
hence "prv (imp (imp \<phi>1 \<chi>1) (imp \<phi>1 (imp \<phi>2 (imp (imp \<phi>2 \<chi>2) (cnj \<chi>1 \<chi>2)))))"
using prv_imp_cnj_imp prv_imp_mp prv_imp_trans2
by (metis (mono_tags) assms cnj prv_prv_imp_trans imp)
hence 1: "prv (imp (imp \<phi>1 \<chi>1) (imp \<phi>1 (imp (imp \<phi>2 \<chi>2) (imp \<phi>2 (cnj \<chi>1 \<chi>2)))))"
using prv_cnj_imp prv_imp_cnjR prv_imp_mp prv_imp_trans1
by (metis (no_types) assms cnj prv_cnj_imp_monoR prv_prv_imp_trans prv_imp_imp_triv imp)
thus ?thesis
by (metis (no_types, lifting) assms cnj imp prv_prv_imp_trans prv_imp_imp_com)
qed
lemma prv_imp_cnj5:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (imp \<phi>1 \<chi>1)" and "prv (imp \<phi>2 \<chi>2)"
shows "prv (imp \<phi>1 (imp \<phi>2 (cnj \<chi>1 \<chi>2)))"
by (simp add: assms prv_cnj_imp prv_cnj_mono)
text \<open>Properties of formula equivalence:\<close>
lemma prv_eqv_imp:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (eqv \<phi> \<chi>) (eqv \<chi> \<phi>))"
by (simp add: assms prv_imp_cnj prv_imp_cnjL prv_imp_cnjR eqv_def)
lemma prv_eqv_eqv:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (eqv (eqv \<phi> \<chi>) (eqv \<chi> \<phi>))"
using assms prv_cnjI prv_eqv_imp eqv_def by auto
lemma prv_imp_eqvEL:
"\<phi>1 \<in> fmla \<Longrightarrow> \<phi>2 \<in> fmla \<Longrightarrow> prv (eqv \<phi>1 \<phi>2) \<Longrightarrow> prv (imp \<phi>1 \<phi>2)"
unfolding eqv_def by (meson cnj imp prv_imp_cnjL prv_imp_mp)
lemma prv_imp_eqvER:
"\<phi>1 \<in> fmla \<Longrightarrow> \<phi>2 \<in> fmla \<Longrightarrow> prv (eqv \<phi>1 \<phi>2) \<Longrightarrow> prv (imp \<phi>2 \<phi>1)"
unfolding eqv_def by (meson cnj imp prv_imp_cnjR prv_imp_mp)
lemma prv_eqv_imp_trans:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (imp (eqv \<phi> \<chi>) (imp (eqv \<chi> \<psi>) (eqv \<phi> \<psi>)))"
proof-
have "prv (imp (eqv \<phi> \<chi>) (imp (imp \<chi> \<psi>) (imp \<phi> \<psi>)))"
using assms prv_imp_cnjL prv_imp_mp prv_imp_trans2 eqv_def
by (metis prv_imp_cnj3L eqv imp)
hence "prv (imp (eqv \<phi> \<chi>) (imp (eqv \<chi> \<psi>) (imp \<phi> \<psi>)))"
using prv_imp_cnjL prv_imp_mp prv_imp_trans2 eqv_def
by (metis (no_types) assms prv_imp_cnj3L prv_imp_com eqv imp)
hence 1: "prv (imp (cnj (eqv \<phi> \<chi>) (eqv \<chi> \<psi>)) (imp \<phi> \<psi>))"
using prv_cnj_imp_monoR2
by (simp add: assms(1) assms(2) assms(3))
have "prv (imp (eqv \<phi> \<chi>) (imp (imp \<psi> \<chi>) (imp \<psi> \<phi>)))"
using prv_imp_cnjR prv_imp_mp prv_imp_trans1 eqv_def
by (metis assms prv_cnj_imp_monoR2 prv_imp_triv imp)
hence "prv (imp (eqv \<phi> \<chi>) (imp (eqv \<chi> \<psi>) (imp \<psi> \<phi>)))"
by (metis assms cnj eqv_def imp prv_imp_cnj3R prv_prv_imp_trans)
hence 2: "prv (imp (cnj (eqv \<phi> \<chi>) (eqv \<chi> \<psi>)) (imp \<psi> \<phi>))"
using prv_cnj_imp_monoR2
by (metis (no_types, lifting) assms eqv imp)
have "prv (imp (cnj (eqv \<phi> \<chi>) (eqv \<chi> \<psi>)) (eqv \<phi> \<psi>))"
using 1 2 using assms prv_imp_cnj by (auto simp: eqv_def[of \<phi> \<psi>])
thus ?thesis
by (simp add: assms prv_cnj_imp)
qed
lemma prv_eqv_cnj_trans:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (imp (cnj (eqv \<phi> \<chi>) (eqv \<chi> \<psi>)) (eqv \<phi> \<psi>))"
by (simp add: assms prv_eqv_imp_trans prv_cnj_imp_monoR2)
lemma prv_eqvI:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (imp \<phi> \<chi>)" and "prv (imp \<chi> \<phi>)"
shows "prv (eqv \<phi> \<chi>)"
by (simp add: assms eqv_def prv_cnjI)
text \<open>Formula equivalence is a congruence (i.e., an equivalence that
is compatible with the other connectives):\<close>
lemma prv_eqv_refl: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv \<phi> \<phi>)"
by (simp add: prv_cnjI prv_imp_refl eqv_def)
lemma prv_eqv_sym:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (eqv \<phi> \<chi>) \<Longrightarrow> prv (eqv \<chi> \<phi>)"
using assms prv_cnjI prv_imp_cnjL prv_imp_cnjR prv_imp_mp eqv_def
by (meson prv_eqv_imp eqv)
lemma prv_eqv_trans:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (eqv \<phi> \<chi>) \<Longrightarrow> prv (eqv \<chi> \<psi>) \<Longrightarrow> prv (eqv \<phi> \<psi>)"
using assms prv_cnjI prv_cnj_imp_monoR2 prv_imp_mp prv_imp_trans1 prv_imp_imp_triv eqv_def
by (metis prv_prv_imp_trans prv_imp_cnjL prv_imp_cnjR eqv imp)
lemma imp_imp_compat_eqvL:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (eqv \<phi>1 \<phi>2) (eqv (imp \<phi>1 \<chi>) (imp \<phi>2 \<chi>)))"
proof -
have f: "prv (imp (eqv \<phi>1 \<phi>2) (eqv (imp \<phi>1 \<chi>) (imp \<phi>2 \<chi>)))"
if "prv (imp (eqv \<phi>1 \<phi>2) (imp (imp \<phi>2 \<chi>) (imp \<phi>1 \<chi>)))" "prv (imp (eqv \<phi>1 \<phi>2) (imp (imp \<phi>1 \<chi>) (imp \<phi>2 \<chi>)))"
using assms that prv_imp_cnj by (auto simp: eqv_def)
moreover have "(prv (imp (eqv \<phi>1 \<phi>2) (imp \<phi>1 \<phi>2)) \<and> prv (imp (eqv \<phi>1 \<phi>2) (imp \<phi>2 \<phi>1)))"
by (simp add: assms eqv_def prv_imp_cnjL prv_imp_cnjR)
ultimately show ?thesis
by (metis (no_types) assms eqv imp prv_imp_trans2 prv_prv_imp_trans)
qed
lemma imp_imp_compat_eqvR:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (eqv \<chi>1 \<chi>2) (eqv (imp \<phi> \<chi>1) (imp \<phi> \<chi>2)))"
by (simp add: assms prv_cnj_mono prv_imp_trans1 eqv_def)
lemma imp_imp_compat_eqv:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (eqv \<phi>1 \<phi>2) (imp (eqv \<chi>1 \<chi>2) (eqv (imp \<phi>1 \<chi>1) (imp \<phi>2 \<chi>2))))"
proof-
have "prv (imp (eqv \<phi>1 \<phi>2) (imp (eqv \<chi>1 \<chi>2) (cnj (eqv (imp \<phi>1 \<chi>1) (imp \<phi>2 \<chi>1))
(eqv (imp \<phi>2 \<chi>1) (imp \<phi>2 \<chi>2)))))"
using prv_imp_cnj5
[OF _ _ _ _ imp_imp_compat_eqvL[of \<phi>1 \<phi>2 \<chi>1] imp_imp_compat_eqvR[of \<phi>2 \<chi>1 \<chi>2]] assms by auto
hence "prv (imp (cnj (eqv \<phi>1 \<phi>2) (eqv \<chi>1 \<chi>2)) (cnj (eqv (imp \<phi>1 \<chi>1) (imp \<phi>2 \<chi>1))
(eqv (imp \<phi>2 \<chi>1) (imp \<phi>2 \<chi>2))))"
by(simp add: assms prv_cnj_imp_monoR2)
hence "prv (imp (cnj (eqv \<phi>1 \<phi>2) (eqv \<chi>1 \<chi>2)) (eqv (imp \<phi>1 \<chi>1) (imp \<phi>2 \<chi>2)))"
using assms prv_eqv_cnj_trans[of "imp \<phi>1 \<chi>1" "imp \<phi>2 \<chi>1" "imp \<phi>2 \<chi>2"]
using prv_imp_mp prv_imp_trans2
by (metis (no_types) cnj prv_prv_imp_trans eqv imp)
thus ?thesis using assms prv_cnj_imp by auto
qed
lemma imp_compat_eqvL:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)"
shows "prv (eqv (imp \<phi>1 \<chi>) (imp \<phi>2 \<chi>))"
using assms prv_imp_mp imp_imp_compat_eqvL by (meson eqv imp)
lemma imp_compat_eqvR:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (eqv \<chi>1 \<chi>2)"
shows "prv (eqv (imp \<phi> \<chi>1) (imp \<phi> \<chi>2))"
using assms prv_imp_mp imp_imp_compat_eqvR by (meson eqv imp)
lemma imp_compat_eqv:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)" and "prv (eqv \<chi>1 \<chi>2)"
shows "prv (eqv (imp \<phi>1 \<chi>1) (imp \<phi>2 \<chi>2))"
using assms prv_imp_mp imp_imp_compat_eqv by (metis eqv imp)
(* *)
lemma imp_cnj_compat_eqvL:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (eqv \<phi>1 \<phi>2) (eqv (cnj \<phi>1 \<chi>) (cnj \<phi>2 \<chi>)))"
proof -
have "prv (imp (imp (imp \<phi>2 \<phi>1) (imp (cnj \<phi>2 \<chi>) (cnj \<phi>1 \<chi>)))
(imp (cnj (imp \<phi>1 \<phi>2) (imp \<phi>2 \<phi>1)) (cnj (imp (cnj \<phi>1 \<chi>) (cnj \<phi>2 \<chi>))
(imp (cnj \<phi>2 \<chi>) (cnj \<phi>1 \<chi>)))))"
by (metis (no_types) imp cnj assms prv_imp_mp assms prv_cnj_imp_mono prv_imp_com prv_imp_refl)
then show ?thesis
by (metis (no_types) imp cnj assms prv_imp_mp assms eqv_def prv_cnj_imp_mono prv_imp_com prv_imp_refl)
qed
lemma imp_cnj_compat_eqvR:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (eqv \<chi>1 \<chi>2) (eqv (cnj \<phi> \<chi>1) (cnj \<phi> \<chi>2)))"
by (simp add: assms prv_cnj_mono prv_imp_cnj3R prv_imp_cnj4 prv_imp_cnjL prv_imp_triv eqv_def)
lemma imp_cnj_compat_eqv:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
shows "prv (imp (eqv \<phi>1 \<phi>2) (imp (eqv \<chi>1 \<chi>2) (eqv (cnj \<phi>1 \<chi>1) (cnj \<phi>2 \<chi>2))))"
proof-
have "prv (imp (eqv \<phi>1 \<phi>2) (imp (eqv \<chi>1 \<chi>2) (cnj (eqv (cnj \<phi>1 \<chi>1) (cnj \<phi>2 \<chi>1))
(eqv (cnj \<phi>2 \<chi>1) (cnj \<phi>2 \<chi>2)))))"
using prv_imp_cnj5
[OF _ _ _ _ imp_cnj_compat_eqvL[of \<phi>1 \<phi>2 \<chi>1] imp_cnj_compat_eqvR[of \<phi>2 \<chi>1 \<chi>2]] assms by auto
hence "prv (imp (cnj (eqv \<phi>1 \<phi>2) (eqv \<chi>1 \<chi>2)) (cnj (eqv (cnj \<phi>1 \<chi>1) (cnj \<phi>2 \<chi>1))
(eqv (cnj \<phi>2 \<chi>1) (cnj \<phi>2 \<chi>2))))"
by(simp add: assms prv_cnj_imp_monoR2)
hence "prv (imp (cnj (eqv \<phi>1 \<phi>2) (eqv \<chi>1 \<chi>2)) (eqv (cnj \<phi>1 \<chi>1) (cnj \<phi>2 \<chi>2)))"
using assms prv_eqv_cnj_trans[of "cnj \<phi>1 \<chi>1" "cnj \<phi>2 \<chi>1" "cnj \<phi>2 \<chi>2"]
using prv_imp_mp prv_imp_trans2
by (metis (no_types) cnj prv_prv_imp_trans eqv)
thus ?thesis using assms prv_cnj_imp by auto
qed
lemma cnj_compat_eqvL:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)"
shows "prv (eqv (cnj \<phi>1 \<chi>) (cnj \<phi>2 \<chi>))"
using assms prv_imp_mp imp_cnj_compat_eqvL by (meson eqv cnj)
lemma cnj_compat_eqvR:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (eqv \<chi>1 \<chi>2)"
shows "prv (eqv (cnj \<phi> \<chi>1) (cnj \<phi> \<chi>2))"
using assms prv_imp_mp imp_cnj_compat_eqvR by (meson eqv cnj)
lemma cnj_compat_eqv:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)" and "prv (eqv \<chi>1 \<chi>2)"
shows "prv (eqv (cnj \<phi>1 \<chi>1) (cnj \<phi>2 \<chi>2))"
using assms prv_imp_mp imp_cnj_compat_eqv by (metis eqv imp cnj)
lemma prv_eqv_prv:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv \<phi>" and "prv (eqv \<phi> \<chi>)"
shows "prv \<chi>"
by (metis assms prv_imp_cnjL prv_imp_mp eqv eqv_def imp)
lemma prv_eqv_prv_rev:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv \<phi>" and "prv (eqv \<chi> \<phi>)"
shows "prv \<chi>"
using assms prv_eqv_prv prv_eqv_sym by blast
lemma prv_imp_eqv_transi:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (imp \<phi> \<chi>1)" and "prv (eqv \<chi>1 \<chi>2)"
shows "prv (imp \<phi> \<chi>2)"
by (meson assms imp imp_compat_eqvR prv_eqv_prv)
lemma prv_imp_eqv_transi_rev:
assumes "\<phi> \<in> fmla" and "\<chi>1 \<in> fmla" and "\<chi>2 \<in> fmla"
assumes "prv (imp \<phi> \<chi>2)" and "prv (eqv \<chi>1 \<chi>2)"
shows "prv (imp \<phi> \<chi>1)"
by (meson assms prv_eqv_sym prv_imp_eqv_transi)
lemma prv_eqv_imp_transi:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)" and "prv (imp \<phi>2 \<chi>)"
shows "prv (imp \<phi>1 \<chi>)"
by (meson assms prv_imp_eqv_transi prv_imp_refl prv_prv_imp_trans)
lemma prv_eqv_imp_transi_rev:
assumes "\<phi>1 \<in> fmla" and "\<phi>2 \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (eqv \<phi>1 \<phi>2)" and "prv (imp \<phi>1 \<chi>)"
shows "prv (imp \<phi>2 \<chi>)"
by (meson assms prv_eqv_imp_transi prv_eqv_sym)
lemma prv_imp_monoL: "\<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
prv (imp \<chi> \<psi>) \<Longrightarrow> prv (imp (imp \<phi> \<chi>) (imp \<phi> \<psi>))"
by (meson imp prv_imp_mp prv_imp_trans1)
lemma prv_imp_monoR: "\<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
prv (imp \<psi> \<chi>) \<Longrightarrow> prv (imp (imp \<chi> \<phi>) (imp \<psi> \<phi>))"
by (meson imp prv_imp_mp prv_imp_trans2)
text \<open>More properties involving conjunction:\<close>
lemma prv_cnj_com_imp:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla"
shows "prv (imp (cnj \<phi> \<chi>) (cnj \<chi> \<phi>))"
by (simp add: prv_imp_cnj prv_imp_cnjL prv_imp_cnjR)
lemma prv_cnj_com:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla"
shows "prv (eqv (cnj \<phi> \<chi>) (cnj \<chi> \<phi>))"
by (simp add: prv_cnj_com_imp prv_eqvI)
lemma prv_cnj_assoc_imp1:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (imp (cnj \<phi> (cnj \<chi> \<psi>)) (cnj (cnj \<phi> \<chi>) \<psi>))"
by (simp add: prv_cnj_imp_monoR2 prv_imp_cnj prv_imp_cnjL prv_imp_cnjR prv_imp_triv)
lemma prv_cnj_assoc_imp2:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (imp (cnj (cnj \<phi> \<chi>) \<psi>) (cnj \<phi> (cnj \<chi> \<psi>)))"
proof -
have "prv (imp (cnj \<phi> \<chi>) (imp \<psi> \<phi>)) \<and> cnj \<chi> \<psi> \<in> fmla \<and> cnj \<phi> \<chi> \<in> fmla"
by (meson \<chi> \<phi> \<psi> cnj imp prv_cnj_imp_monoR2 prv_imp_imp_triv prv_prv_imp_trans)
then show ?thesis
using \<chi> \<phi> \<psi> cnj imp prv_cnj_imp_monoR2 prv_imp_cnj4 prv_imp_cnjI prv_imp_triv by presburger
qed
lemma prv_cnj_assoc:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (eqv (cnj \<phi> (cnj \<chi> \<psi>)) (cnj (cnj \<phi> \<chi>) \<psi>))"
by (simp add: prv_cnj_assoc_imp1 prv_cnj_assoc_imp2 prv_eqvI)
lemma prv_cnj_com_imp3:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla" "\<phi>3 \<in> fmla"
shows "prv (imp (cnj \<phi>1 (cnj \<phi>2 \<phi>3))
(cnj \<phi>2 (cnj \<phi>1 \<phi>3)))"
by (simp add: assms prv_cnj_imp_monoR2 prv_imp_cnj prv_imp_cnjL prv_imp_refl prv_imp_triv)
subsection \<open>Properties involving quantifiers\<close>
text \<open>Fundamental properties:\<close>
lemma prv_allE:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t \<in> trm"
shows "prv (all x \<phi>) \<Longrightarrow> prv (subst \<phi> t x)"
using assms prv_all_inst prv_imp_mp by (meson subst all)
lemma prv_exiI:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t \<in> trm"
shows "prv (subst \<phi> t x) \<Longrightarrow> prv (exi x \<phi>)"
using assms prv_exi_inst prv_imp_mp by (meson subst exi)
lemma prv_imp_imp_exi:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "x \<notin> Fvars \<phi>"
shows "prv (imp (exi x (imp \<phi> \<chi>)) (imp \<phi> (exi x \<chi>)))"
using assms imp exi Fvars_exi Fvars_imp Un_iff assms prv_exi_imp_gen prv_exi_inst prv_imp_mp
prv_imp_trans1 member_remove remove_def subst_same_Var by (metis (full_types) Var)
lemma prv_imp_exi:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
shows "x \<notin> Fvars \<phi> \<Longrightarrow> prv (exi x (imp \<phi> \<chi>)) \<Longrightarrow> prv (imp \<phi> (exi x \<chi>))"
using assms prv_imp_imp_exi prv_imp_mp by (meson exi imp)
lemma prv_exi_imp:
assumes x: "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "x \<notin> Fvars \<chi>" and d: "prv (all x (imp \<phi> \<chi>))"
shows "prv (imp (exi x \<phi>) \<chi>)"
proof-
have "prv (imp \<phi> \<chi>)" using prv_allE[of x _ "Var x", of "imp \<phi> \<chi>"] assms by simp
thus ?thesis using assms prv_exi_imp_gen by blast
qed
lemma prv_all_imp:
assumes x: "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "x \<notin> Fvars \<phi>" and "prv (all x (imp \<phi> \<chi>))"
shows "prv (imp \<phi> (all x \<chi>))"
proof-
have "prv (imp \<phi> \<chi>)" using prv_allE[of x _ "Var x", of "imp \<phi> \<chi>"] assms by simp
thus ?thesis using assms prv_all_imp_gen by blast
qed
lemma prv_exi_inst_same:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "x \<in> var"
shows "prv (imp \<phi> (exi x \<phi>))"
proof-
have 0: "\<phi> = subst \<phi> (Var x) x" using assms by simp
show ?thesis
apply(subst 0)
using assms by (intro prv_exi_inst) auto
qed
lemma prv_exi_cong:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "x \<in> var"
and "prv (imp \<phi> \<chi>)"
shows "prv (imp (exi x \<phi>) (exi x \<chi>))"
proof-
have 0: "prv (imp \<chi> (exi x \<chi>))" using assms prv_exi_inst_same by auto
show ?thesis
using assms apply(intro prv_exi_imp_gen)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal using "0" exi prv_prv_imp_trans by blast .
qed
lemma prv_exi_congW:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "x \<in> var"
and "prv (imp \<phi> \<chi>)" "prv (exi x \<phi>)"
shows "prv (exi x \<chi>)"
by (meson exi assms prv_exi_cong prv_imp_mp)
lemma prv_all_cong:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "x \<in> var"
and "prv (imp \<phi> \<chi>)"
shows "prv (imp (all x \<phi>) (all x \<chi>))"
proof-
have 0: "prv (imp (all x \<phi>) \<chi>)" using assms
by (metis Var all prv_all_inst prv_prv_imp_trans subst_same_Var)
show ?thesis by (simp add: "0" assms prv_all_imp_gen)
qed
lemma prv_all_congW:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "x \<in> var"
and "prv (imp \<phi> \<chi>)" "prv (all x \<phi>)"
shows "prv (all x \<chi>)"
by (meson all assms prv_all_cong prv_imp_mp)
text \<open>Quantifiers versus free variables and substitution:\<close>
lemma exists_no_Fvars: "\<exists> \<phi>. \<phi> \<in> fmla \<and> prv \<phi> \<and> Fvars \<phi> = {}"
proof-
obtain n where [simp]: "n \<in> num" using numNE by blast
show ?thesis
by (intro exI[of _ "imp (eql n n) (eql n n)"]) (simp add: prv_imp_refl)
qed
lemma prv_all_gen:
assumes "x \<in> var" and "\<phi> \<in> fmla"
assumes "prv \<phi>" shows "prv (all x \<phi>)"
using assms prv_all_imp_gen prv_imp_mp prv_imp_triv exists_no_Fvars by blast
lemma all_subst_rename_prv:
"\<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow>
y \<notin> Fvars \<phi> \<Longrightarrow> prv (all x \<phi>) \<Longrightarrow> prv (all y (subst \<phi> (Var y) x))"
by (simp add: prv_allE prv_all_gen)
lemma allE_id:
assumes "y \<in> var" and "\<phi> \<in> fmla"
assumes "prv (all y \<phi>)"
shows "prv \<phi>"
using assms prv_allE by (metis Var subst_same_Var)
lemma prv_subst:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t \<in> trm"
shows "prv \<phi> \<Longrightarrow> prv (subst \<phi> t x)"
by (simp add: assms prv_allE prv_all_gen)
lemma prv_rawpsubst:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "prv \<phi>"
shows "prv (rawpsubst \<phi> txs)"
using assms apply (induct txs arbitrary: \<phi>)
subgoal by auto
subgoal for tx txs \<phi> by (cases tx) (auto intro: prv_subst) .
lemma prv_psubst:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "prv \<phi>"
shows "prv (psubst \<phi> txs)"
proof-
define us where us: "us \<equiv> getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by auto
subgoal by (fastforce simp: image_iff)
by auto
show ?thesis using assms us_facts unfolding psubst_def
by (auto simp: Let_def us[symmetric] intro!: prv_rawpsubst rawpsubst dest!: set_zip_D)
qed
lemma prv_eqv_rawpsubst:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> snd ` set txs \<subseteq> var \<Longrightarrow> fst ` set txs \<subseteq> trm \<Longrightarrow> prv (eqv \<phi> \<psi>) \<Longrightarrow>
prv (eqv (rawpsubst \<phi> txs) (rawpsubst \<psi> txs))"
by (metis eqv prv_rawpsubst rawpsubst_eqv)
lemma prv_eqv_psubst:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> snd ` set txs \<subseteq> var \<Longrightarrow> fst ` set txs \<subseteq> trm \<Longrightarrow> prv (eqv \<phi> \<psi>) \<Longrightarrow>
distinct (map snd txs) \<Longrightarrow>
prv (eqv (psubst \<phi> txs) (psubst \<psi> txs))"
by (metis eqv prv_psubst psubst_eqv)
lemma prv_all_imp_trans:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (all x (imp \<phi> \<chi>)) \<Longrightarrow> prv (all x (imp \<chi> \<psi>)) \<Longrightarrow> prv (all x (imp \<phi> \<psi>))"
by (metis Var assms prv_allE prv_all_gen prv_prv_imp_trans imp subst_same_Var)
lemma prv_all_imp_cnj:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (all x (imp \<phi> (imp \<psi> \<chi>))) \<Longrightarrow> prv (all x (imp (cnj \<psi> \<phi>) \<chi>))"
by (metis Var assms cnj prv_allE prv_all_gen prv_imp_com prv_cnj_imp_monoR2 imp subst_same_Var)
lemma prv_all_imp_cnj_rev:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "\<chi> \<in> fmla" and "\<psi> \<in> fmla"
shows "prv (all x (imp (cnj \<phi> \<psi>) \<chi>)) \<Longrightarrow> prv (all x (imp \<phi> (imp \<psi> \<chi>)))"
by (metis (full_types) Var assms cnj prv_allE prv_all_gen prv_cnj_imp imp subst_same_Var)
subsection \<open>Properties concerning equality\<close>
lemma prv_eql_reflT:
assumes t: "t \<in> trm"
shows "prv (eql t t)"
proof-
obtain x where x: "x \<in> var" using var_NE by auto
show ?thesis using prv_subst[OF x _ t prv_eql_refl[OF x]] x t by simp
qed
lemma prv_eql_subst_trm:
assumes xx: "x \<in> var" and \<phi>: "\<phi> \<in> fmla" and "t1 \<in> trm" and "t2 \<in> trm"
shows "prv ((imp (eql t1 t2)
(imp (subst \<phi> t1 x) (subst \<phi> t2 x))))"
proof-
have "finite ({x} \<union> FvarsT t1 \<union> FvarsT t2 \<union> Fvars \<phi>)" (is "finite ?A")
by (simp add: assms finite_FvarsT finite_Fvars)
then obtain y where y: "y \<in> var" and "y \<notin> ?A"
by (meson finite_subset infinite_var subset_iff)
hence xy: "x \<noteq> y" and yt1: "y \<notin> FvarsT t1" and yt2: "y \<notin> FvarsT t2" and y\<phi>: "y \<notin> Fvars \<phi>" by auto
have x: "x \<notin> Fvars (subst \<phi> (Var y) x)" using xy y assms by simp
hence 1: "prv (imp (eql t1 (Var y)) (imp (subst \<phi> t1 x) (subst \<phi> (Var y) x)))"
using xy y assms prv_subst[OF xx _ _ prv_eql_subst[OF xx y \<phi>]] by simp
have yy: "y \<notin> Fvars (subst \<phi> t1 x)" using yt1 y\<phi> assms by simp
from prv_subst[OF y _ _ 1, of t2]
show ?thesis using xy yt1 yt2 y\<phi> x xx y yy assms by auto
qed
lemma prv_eql_subst_trm2:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t1 \<in> trm" and "t2 \<in> trm"
assumes "prv (eql t1 t2)"
shows "prv (imp (subst \<phi> t1 x) (subst \<phi> t2 x))"
by (meson assms eql imp local.subst prv_eql_subst_trm prv_imp_mp)
lemma prv_eql_sym:
assumes [simp]: "t1 \<in> trm" "t2 \<in> trm"
shows "prv (imp (eql t1 t2) (eql t2 t1))"
proof-
obtain x where x[simp]: "x \<in> var" "x \<notin> FvarsT t1"
by (meson assms finite_FvarsT finite_subset infinite_var subsetI)
thus ?thesis using prv_eql_subst_trm[of x "eql (Var x) t1" t1 t2, simplified]
by (meson assms eql imp prv_eql_reflT prv_imp_com prv_imp_mp)
qed
lemma prv_prv_eql_sym: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> prv (eql t1 t2) \<Longrightarrow> prv (eql t2 t1)"
by (meson eql prv_eql_sym prv_imp_mp)
lemma prv_all_eql:
assumes "x \<in> var" and "y \<in> var" and "\<phi> \<in> fmla" and "t1 \<in> trm" and "t2 \<in> trm"
shows "prv (all x ((imp (eql t1 t2)
(imp (subst \<phi> t2 y) (subst \<phi> t1 y)))))"
by (meson subst assms prv_all_gen prv_prv_imp_trans prv_eql_subst_trm prv_eql_sym eql imp)
lemma prv_eql_subst_trm_rev:
assumes "t1 \<in> trm" and "t2 \<in> trm" and "\<phi> \<in> fmla" and "y \<in> var"
shows
"prv ((imp (eql t1 t2)
(imp (subst \<phi> t2 y) (subst \<phi> t1 y))))"
using assms prv_all_eql prv_all_inst prv_imp_mp subst_same_Var
by (meson subst prv_prv_imp_trans prv_eql_subst_trm prv_eql_sym eql imp)
lemma prv_eql_subst_trm_rev2:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t1 \<in> trm" and "t2 \<in> trm"
assumes "prv (eql t1 t2)"
shows "prv (imp (subst \<phi> t2 x) (subst \<phi> t1 x))"
by (meson assms eql imp local.subst prv_eql_subst_trm_rev prv_imp_mp)
lemma prv_eql_subst_trm_eqv:
assumes "x \<in> var" and "\<phi> \<in> fmla" and "t1 \<in> trm" and "t2 \<in> trm"
assumes "prv (eql t1 t2)"
shows "prv (eqv (subst \<phi> t1 x) (subst \<phi> t2 x))"
using assms prv_eql_subst_trm2[OF assms] prv_eql_subst_trm_rev2[OF assms]
prv_eqvI by auto
lemma prv_eql_subst_trm_id:
assumes "y \<in> var" "\<phi> \<in> fmla" and "n \<in> num"
shows "prv (imp (eql (Var y) n) (imp \<phi> (subst \<phi> n y)))"
using assms prv_eql_subst_trm
by (metis Var in_num subst_same_Var)
lemma prv_eql_subst_trm_id_back:
assumes "y \<in> var" "\<phi> \<in> fmla" and "n \<in> num"
shows "prv (imp (eql (Var y) n) (imp (subst \<phi> n y) \<phi>))"
by (metis Var assms in_num prv_eql_subst_trm_rev subst_same_Var)
lemma prv_eql_subst_trm_id_rev:
assumes "y \<in> var" "\<phi> \<in> fmla" and "n \<in> num"
shows "prv (imp (eql n (Var y)) (imp \<phi> (subst \<phi> n y)))"
using assms prv_eql_subst_trm_rev by (metis Var in_num subst_same_Var)
lemma prv_eql_subst_trm_id_back_rev:
assumes "y \<in> var" "\<phi> \<in> fmla" and "n \<in> num"
shows "prv (imp (eql n (Var y)) (imp (subst \<phi> n y) \<phi>))"
by (metis (full_types) Var assms in_num prv_eql_subst_trm subst_same_Var)
lemma prv_eql_imp_trans_rev:
assumes t1[simp]: "t1 \<in> trm" and t2[simp]: "t2 \<in> trm" and t3[simp]: "t3 \<in> trm"
shows "prv (imp (eql t1 t2) (imp (eql t1 t3) (eql t2 t3)))"
proof-
obtain y1 where y1[simp]: "y1 \<in> var" and "y1 \<notin> FvarsT t1 \<union> FvarsT t2 \<union> FvarsT t3"
using finite_FvarsT finite_subset infinite_var subsetI t1 t2 t3
by (metis (full_types) infinite_Un)
hence [simp]: "y1 \<notin> FvarsT t1" "y1 \<notin> FvarsT t2" "y1 \<notin> FvarsT t3" by auto
obtain y2 where y2[simp]: "y2 \<in> var" and "y2 \<notin> FvarsT t1 \<union> FvarsT t2 \<union> FvarsT t3 \<union> {y1}"
using finite_FvarsT finite_subset infinite_var subsetI t1 t2 t3
by (metis (full_types) finite_insert infinite_Un insert_is_Un)
hence [simp]: "y2 \<notin> FvarsT t1" "y2 \<notin> FvarsT t2" "y2 \<notin> FvarsT t3" "y1 \<noteq> y2" by auto
have 0: "prv (imp (eql (Var y1) (Var y2))
(imp (eql (Var y1) t3) (eql (Var y2) t3)))"
using prv_eql_subst[OF y1 y2, of "eql (Var y1) t3"] by simp
note 1 = prv_subst[OF y1 _ t1 0, simplified]
show ?thesis using prv_subst[OF y2 _ t2 1, simplified] .
qed
lemma prv_eql_imp_trans:
assumes t1[simp]: "t1 \<in> trm" and t2[simp]: "t2 \<in> trm" and t3[simp]: "t3 \<in> trm"
shows "prv (imp (eql t1 t2) (imp (eql t2 t3) (eql t1 t3)))"
by (meson eql imp prv_eql_sym prv_eql_imp_trans_rev prv_prv_imp_trans t1 t2 t3)
lemma prv_eql_trans:
assumes t1[simp]: "t1 \<in> trm" and t2[simp]: "t2 \<in> trm" and t3[simp]: "t3 \<in> trm"
and "prv (eql t1 t2)" and "prv (eql t2 t3)"
shows "prv (eql t1 t3)"
by (meson assms cnj eql prv_cnjI prv_cnj_imp_monoR2 prv_eql_imp_trans prv_imp_mp)
subsection \<open>The equivalence between soft substitution and substitution\<close>
lemma prv_subst_imp_softSubst:
assumes [simp,intro!]: "x \<in> var" "t \<in> trm" "\<phi> \<in> fmla" "x \<notin> FvarsT t"
shows "prv (imp (subst \<phi> t x) (softSubst \<phi> t x))"
unfolding softSubst_def by (rule prv_imp_exi)
(auto simp: prv_eql_reflT prv_imp_cnj prv_imp_refl prv_imp_triv subst_compose_same
intro: prv_exiI[of _ _ t])
lemma prv_subst_implies_softSubst:
assumes "x \<in> var" "t \<in> trm" "\<phi> \<in> fmla"
and "x \<notin> FvarsT t"
and "prv (subst \<phi> t x)"
shows "prv (softSubst \<phi> t x)"
using assms prv_subst_imp_softSubst
by (metis Var cnj eql exi subst prv_imp_mp softSubst_def)
lemma prv_softSubst_imp_subst:
assumes [simp,intro!]: "x \<in> var" "t \<in> trm" "\<phi> \<in> fmla" "x \<notin> FvarsT t"
shows "prv (imp (softSubst \<phi> t x) (subst \<phi> t x))"
unfolding softSubst_def apply(rule prv_exi_imp_gen)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by (metis Var assms(1-3) eql subst prv_cnj_imp_monoR2 prv_eql_subst_trm subst_same_Var) .
lemma prv_softSubst_implies_subst:
assumes "x \<in> var" "t \<in> trm" "\<phi> \<in> fmla"
and "x \<notin> FvarsT t"
and "prv (softSubst \<phi> t x)"
shows "prv (subst \<phi> t x)"
by (metis Var assms cnj eql exi local.subst prv_imp_mp prv_softSubst_imp_subst softSubst_def)
lemma prv_softSubst_eqv_subst:
assumes [simp,intro!]: "x \<in> var" "t \<in> trm" "\<phi> \<in> fmla" "x \<notin> FvarsT t"
shows "prv (eqv (softSubst \<phi> t x) (subst \<phi> t x))"
by (metis Var assms cnj eql exi local.subst prv_eqvI prv_softSubst_imp_subst prv_subst_imp_softSubst softSubst_def)
end \<comment> \<open>context @{locale Deduct}\<close>
section \<open>Deduction Considering False\<close>
locale Deduct_with_False =
Syntax_with_Connectives_False
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
+
Deduct
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and num
and prv
+
assumes
prv_fls[simp,intro]: "\<And>\<phi>. prv (imp fls \<phi>)"
begin
subsection \<open>Basic properties of False (fls)\<close>
lemma prv_expl:
assumes "\<phi> \<in> fmla"
assumes "prv fls"
shows "prv \<phi>"
using assms prv_imp_mp by blast
lemma prv_cnjR_fls: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (cnj fls \<phi>) fls)"
by (simp add: prv_eqvI prv_imp_cnjL)
lemma prv_cnjL_fls: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (cnj \<phi> fls) fls)"
by (simp add: prv_eqvI prv_imp_cnjR)
subsection \<open>Properties involving negation\<close>
text \<open>Recall that negation has been defined from implication and False.\<close>
lemma prv_imp_neg_fls:
assumes "\<phi> \<in> fmla"
shows "prv (imp \<phi> (imp (neg \<phi>) fls))"
using assms prv_imp_com prv_imp_refl neg_def by auto
lemma prv_neg_fls:
assumes "\<phi> \<in> fmla"
assumes "prv \<phi>" and "prv (neg \<phi>)"
shows "prv fls"
by (metis assms prv_imp_mp fls neg_def)
lemma prv_imp_neg_neg:
assumes "\<phi> \<in> fmla"
shows "prv (imp \<phi> (neg (neg \<phi>)))"
using assms prv_imp_neg_fls neg_def by auto
lemma prv_neg_neg:
assumes "\<phi> \<in> fmla"
assumes "prv \<phi>"
shows "prv (neg (neg \<phi>))"
by (metis assms prv_imp_mp prv_imp_neg_fls neg neg_def)
lemma prv_imp_imp_neg_rev:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
shows "prv (imp (imp \<phi> \<chi>)
(imp (neg \<chi>) (neg \<phi>)))"
unfolding neg_def using prv_imp_trans2[OF assms fls] .
lemma prv_imp_neg_rev:
assumes "\<phi> \<in> fmla" and "\<chi> \<in> fmla"
assumes "prv (imp \<phi> \<chi>)"
shows "prv (imp (neg \<chi>) (neg \<phi>))"
by (meson assms imp neg prv_imp_imp_neg_rev prv_imp_mp)
lemma prv_eqv_neg_prv_fls:
"\<phi> \<in> fmla \<Longrightarrow>
prv (eqv \<phi> (neg \<phi>)) \<Longrightarrow> prv fls"
by (metis cnj fls neg neg_def prv_cnj_imp_monoR2 prv_eqv_imp_transi prv_imp_cnj prv_imp_eqvER prv_imp_mp prv_imp_neg_rev)
lemma prv_eqv_eqv_neg_prv_fls:
"\<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (eqv \<phi> \<chi>) \<Longrightarrow> prv (eqv \<phi> (neg \<chi>))\<Longrightarrow> prv fls"
by (meson neg prv_eqv_neg_prv_fls prv_eqv_sym prv_eqv_trans)
lemma prv_eqv_eqv_neg_prv_fls2:
"\<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (eqv \<phi> \<chi>) \<Longrightarrow> prv (eqv \<chi> (neg \<phi>))\<Longrightarrow> prv fls"
by (simp add: prv_eqv_eqv_neg_prv_fls prv_eqv_sym)
lemma prv_neg_imp_imp_trans:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "\<psi> \<in> fmla"
and "prv (neg \<phi>)"
and "prv (imp \<chi> (imp \<psi> \<phi>))"
shows "prv (imp \<chi> (neg \<psi>))"
unfolding neg_def
by (metis assms cnj fls neg_def prv_cnj_imp prv_cnj_imp_monoR2 prv_prv_imp_trans)
lemma prv_imp_neg_imp_neg_imp:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla"
and "prv (neg \<phi>)"
shows "prv ((imp \<chi> (neg (imp \<chi> \<phi>))))"
by (metis assms fls imp neg_def prv_imp_com prv_imp_monoL)
lemma prv_prv_neg_imp_neg:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla"
and "prv \<phi>" and "prv \<chi>"
shows "prv (neg (imp \<phi> (neg \<chi>)))"
by (meson assms imp neg prv_imp_mp prv_imp_neg_imp_neg_imp prv_neg_neg)
lemma prv_imp_neg_imp_cnjL:
assumes "\<phi> \<in> fmla" "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "prv (imp \<phi> (neg \<phi>1))"
shows "prv (imp \<phi> (neg (cnj \<phi>1 \<phi>2)))"
unfolding neg_def by (metis assms cnj fls neg neg_def prv_imp_cnj3L prv_prv_imp_trans)
lemma prv_imp_neg_imp_cnjR:
assumes "\<phi> \<in> fmla" "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "prv (imp \<phi> (neg \<phi>2))"
shows "prv (imp \<phi> (neg (cnj \<phi>1 \<phi>2)))"
unfolding neg_def by (metis assms cnj fls neg neg_def prv_imp_cnj3R prv_prv_imp_trans)
text \<open>Negation versus quantifiers:\<close>
lemma prv_all_neg_imp_neg_exi:
assumes x: "x \<in> var" and \<phi>: "\<phi> \<in> fmla"
shows "prv (imp (all x (neg \<phi>)) (neg (exi x \<phi>)))"
proof-
have 0: "prv (imp (all x (neg \<phi>)) (imp \<phi> fls))"
using prv_all_inst[OF x, of "neg \<phi>" "Var x",simplified] assms by (auto simp: neg_def)
have 1: "prv (imp \<phi> (imp (all x (neg \<phi>)) fls))"
using 0 by (simp add: assms prv_imp_com)
have 2: "prv (imp (exi x \<phi>) (imp (all x (neg \<phi>)) fls))"
using 1 assms by (intro prv_exi_imp_gen) auto
thus ?thesis by (simp add: assms neg_def prv_imp_com)
qed
lemma prv_neg_exi_imp_all_neg:
assumes x: "x \<in> var" and \<phi>: "\<phi> \<in> fmla"
shows "prv (imp (neg (exi x \<phi>)) (all x (neg \<phi>)))"
using assms
by (intro prv_all_imp_gen[of x "neg (exi x \<phi>)"])
(auto simp: prv_exi_inst_same prv_imp_neg_rev)
lemma prv_neg_exi_eqv_all_neg:
assumes x: "x \<in> var" and \<phi>: "\<phi> \<in> fmla"
shows "prv (eqv (neg (exi x \<phi>)) (all x (neg \<phi>)))"
by (simp add: \<phi> prv_all_neg_imp_neg_exi prv_eqvI prv_neg_exi_imp_all_neg x)
lemma prv_neg_exi_implies_all_neg:
assumes x: "x \<in> var" and \<phi>: "\<phi> \<in> fmla" and "prv (neg (exi x \<phi>))"
shows "prv (all x (neg \<phi>))"
by (meson \<phi> all assms(3) exi neg prv_imp_mp prv_neg_exi_imp_all_neg x)
lemma prv_neg_neg_exi:
assumes "x \<in> var" "\<phi> \<in> fmla"
and "prv (neg \<phi>)"
shows "prv (neg (exi x \<phi>))"
using assms neg_def prv_exi_imp_gen by auto
lemma prv_exi_imp_exiI:
assumes [simp]: "x \<in> var" "y \<in> var" "\<phi> \<in> fmla" and yx: "y = x \<or> y \<notin> Fvars \<phi>"
shows "prv (imp (exi x \<phi>) (exi y (subst \<phi> (Var y) x)))"
proof(cases "y = x")
case [simp]: False hence [simp]: "x \<noteq> y" by blast
show ?thesis apply(rule prv_exi_imp_gen)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
using yx
by (auto intro!: prv_imp_exi prv_exiI[of _ _ "Var x"]
simp: prv_imp_refl2)
qed(simp add: yx prv_imp_refl)
lemma prv_imp_neg_allI:
assumes "\<phi> \<in> fmla" "\<chi> \<in> fmla" "t \<in> trm" "x \<in> var"
and "prv (imp \<phi> (neg (subst \<chi> t x)))"
shows "prv (imp \<phi> (neg (all x \<chi>)))"
by (meson all assms subst neg prv_all_inst prv_imp_neg_rev prv_prv_imp_trans)
lemma prv_imp_neg_allWI:
assumes "\<chi> \<in> fmla" "t \<in> trm" "x \<in> var"
and "prv (neg (subst \<chi> t x))"
shows "prv (neg (all x \<chi>))"
by (metis all assms fls subst neg_def prv_all_inst prv_prv_imp_trans)
subsection \<open>Properties involving True (tru)\<close>
lemma prv_imp_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (imp \<phi> tru)"
by (simp add: neg_def prv_imp_triv tru_def)
lemma prv_tru_imp: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (imp tru \<phi>) \<phi>)"
by (metis imp neg_def prv_eqvI prv_fls prv_imp_com prv_imp_imp_triv prv_imp_mp prv_imp_refl tru tru_def)
lemma prv_fls_neg_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv fls (neg tru))"
using neg_def prv_eqvI prv_neg_neg tru_def by auto
lemma prv_tru_neg_fls: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv tru (neg fls))"
by (simp add: prv_eqv_refl tru_def)
lemma prv_cnjR_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (cnj tru \<phi>) \<phi>)"
by (simp add: prv_eqvI prv_imp_cnj prv_imp_cnjR prv_imp_refl prv_imp_tru)
lemma prv_cnjL_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (cnj \<phi> tru) \<phi>)"
by (simp add: prv_eqvI prv_imp_cnj prv_imp_cnjL prv_imp_refl prv_imp_tru)
subsection \<open>Property of set-based conjunctions\<close>
text \<open>These are based on properties of the auxiliary list conjunctions.\<close>
lemma prv_lcnj_imp_in:
assumes "set \<phi>s \<subseteq> fmla"
and "\<phi> \<in> set \<phi>s"
shows "prv (imp (lcnj \<phi>s) \<phi>)"
proof-
have "\<phi> \<in> fmla" using assms by auto
thus ?thesis using assms
by (induct \<phi>s arbitrary: \<phi>)
(auto simp : prv_imp_cnjL prv_cnj_imp_monoR2 prv_imp_triv)
qed
lemma prv_lcnj_imp:
assumes "\<chi> \<in> fmla" and "set \<phi>s \<subseteq> fmla"
and "\<phi> \<in> set \<phi>s" and "prv (imp \<phi> \<chi>)"
shows "prv (imp (lcnj \<phi>s) \<chi>)"
using assms lcnj prv_lcnj_imp_in prv_prv_imp_trans by blast
lemma prv_imp_lcnj:
assumes "\<chi> \<in> fmla" and "set \<phi>s \<subseteq> fmla"
and "\<And>\<phi>. \<phi> \<in> set \<phi>s \<Longrightarrow> prv (imp \<chi> \<phi>)"
shows "prv (imp \<chi> (lcnj \<phi>s))"
using assms
by (induct \<phi>s arbitrary: \<chi>) (auto simp: prv_imp_tru prv_imp_com prv_imp_cnj)
lemma prv_lcnj_imp_inner:
assumes "\<phi> \<in> fmla" "set \<phi>1s \<subseteq> fmla" "set \<phi>2s \<subseteq> fmla"
shows "prv (imp (cnj \<phi> (lcnj (\<phi>1s @ \<phi>2s))) (lcnj (\<phi>1s @ \<phi> # \<phi>2s)))"
using assms proof(induction \<phi>1s)
case (Cons \<phi>1 \<phi>1s)
have [intro!]: "set (\<phi>1s @ \<phi>2s) \<subseteq> fmla" "set (\<phi>1s @ \<phi> # \<phi>2s) \<subseteq> fmla" using Cons by auto
have 0: "prv (imp (cnj \<phi> (cnj \<phi>1 (lcnj (\<phi>1s @ \<phi>2s))))
(cnj \<phi>1 (cnj \<phi> (lcnj (\<phi>1s @ \<phi>2s)))))"
using Cons by (intro prv_cnj_com_imp3) fastforce+
have 1: "prv (imp (cnj \<phi>1 (cnj \<phi> (lcnj (\<phi>1s @ \<phi>2s))))
(cnj \<phi>1 (lcnj (\<phi>1s @ \<phi> # \<phi>2s))))"
using Cons by (intro prv_cnj_mono) (auto simp add: prv_imp_refl)
show ?case using prv_prv_imp_trans[OF _ _ _ 0 1] Cons by auto
qed(simp add: prv_imp_refl)
lemma prv_lcnj_imp_remdups:
assumes "set \<phi>s \<subseteq> fmla"
shows "prv (imp (lcnj (remdups \<phi>s)) (lcnj \<phi>s))"
using assms apply(induct \<phi>s)
by (auto simp: prv_imp_cnj prv_lcnj_imp_in prv_cnj_mono prv_imp_refl)
lemma prv_lcnj_mono:
assumes \<phi>1s: "set \<phi>1s \<subseteq> fmla" and "set \<phi>2s \<subseteq> set \<phi>1s"
shows "prv (imp (lcnj \<phi>1s) (lcnj \<phi>2s))"
proof-
define \<phi>2s' where \<phi>2s': "\<phi>2s' = remdups \<phi>2s"
have A: "set \<phi>2s' \<subseteq> set \<phi>1s" "distinct \<phi>2s'" unfolding \<phi>2s' using assms by auto
have B: "prv (imp (lcnj \<phi>1s) (lcnj \<phi>2s'))"
using \<phi>1s A proof(induction \<phi>1s arbitrary: \<phi>2s')
case (Cons \<phi>1 \<phi>1s \<phi>2ss)
show ?case proof(cases "\<phi>1 \<in> set \<phi>2ss")
case True
then obtain \<phi>2ss1 \<phi>2ss2 where \<phi>2ss: "\<phi>2ss = \<phi>2ss1 @ \<phi>1 # \<phi>2ss2"
by (meson split_list)
define \<phi>2s where \<phi>2s: "\<phi>2s \<equiv> \<phi>2ss1 @ \<phi>2ss2"
- have nin: "\<phi>1 \<notin> set \<phi>2s" using \<phi>2ss `distinct \<phi>2ss` unfolding \<phi>2s by auto
+ have nin: "\<phi>1 \<notin> set \<phi>2s" using \<phi>2ss \<open>distinct \<phi>2ss\<close> unfolding \<phi>2s by auto
have [intro!]: "set \<phi>2s \<subseteq> fmla" unfolding \<phi>2s using \<phi>2ss Cons by auto
have 0: "prv (imp (cnj \<phi>1 (lcnj \<phi>2s)) (lcnj \<phi>2ss))"
unfolding \<phi>2s \<phi>2ss using Cons \<phi>2ss
by (intro prv_lcnj_imp_inner) auto
have 1: "prv (imp (lcnj \<phi>1s) (lcnj \<phi>2s))"
using nin Cons.prems True \<phi>2s \<phi>2ss by (intro Cons.IH) auto
have 2: "prv (imp (cnj \<phi>1 (lcnj \<phi>1s)) (cnj \<phi>1 (lcnj \<phi>2s)))"
using Cons \<phi>2ss \<phi>2s 1 by (intro prv_cnj_mono) (fastforce simp add: prv_imp_refl)+
show ?thesis
using Cons by (auto intro!: prv_prv_imp_trans[OF _ _ _ 2 0])
next
case False
then show ?thesis
by (meson Cons lcnj prv_imp_lcnj prv_lcnj_imp_in subset_iff)
qed
qed(simp add: prv_imp_refl)
have C: "prv (imp (lcnj \<phi>2s') (lcnj \<phi>2s))"
unfolding \<phi>2s' using assms by (intro prv_lcnj_imp_remdups) auto
show ?thesis using A assms by (intro prv_prv_imp_trans[OF _ _ _ B C]) auto
qed
lemma prv_lcnj_eqv:
assumes "set \<phi>1s \<subseteq> fmla" and "set \<phi>2s = set \<phi>1s"
shows "prv (eqv (lcnj \<phi>1s) (lcnj \<phi>2s))"
using assms prv_lcnj_mono by (intro prv_eqvI) auto
lemma prv_lcnj_mono_imp:
assumes "set \<phi>1s \<subseteq> fmla" "set \<phi>2s \<subseteq> fmla" and "\<forall> \<phi>2 \<in> set \<phi>2s. \<exists> \<phi>1 \<in> set \<phi>1s. prv (imp \<phi>1 \<phi>2)"
shows "prv (imp (lcnj \<phi>1s) (lcnj \<phi>2s))"
using assms apply(intro prv_imp_lcnj)
subgoal by auto
subgoal by auto
subgoal using prv_lcnj_imp by blast .
text \<open>Set-based conjunction commutes with substitution only up to provably equivalence:\<close>
lemma prv_subst_scnj:
assumes "F \<subseteq> fmla" "finite F" "t \<in> trm" "x \<in> var"
shows "prv (eqv (subst (scnj F) t x) (scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)))"
using assms unfolding scnj_def by (fastforce intro!: prv_lcnj_eqv)
lemma prv_imp_subst_scnj:
assumes "F \<subseteq> fmla" "finite F" "t \<in> trm" "x \<in> var"
shows "prv (imp (subst (scnj F) t x) (scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)))"
using prv_subst_scnj[OF assms] assms by (intro prv_imp_eqvEL) auto
lemma prv_subst_scnj_imp:
assumes "F \<subseteq> fmla" "finite F" "t \<in> trm" "x \<in> var"
shows "prv (imp (scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)) (subst (scnj F) t x))"
using prv_subst_scnj[OF assms] assms by (intro prv_imp_eqvER) auto
lemma prv_scnj_imp_in:
assumes "F \<subseteq> fmla" "finite F"
and "\<phi> \<in> F"
shows "prv (imp (scnj F) \<phi>)"
unfolding scnj_def using assms by (intro prv_lcnj_imp_in) auto
lemma prv_scnj_imp:
assumes "\<chi> \<in> fmla" and "F \<subseteq> fmla" "finite F"
and "\<phi> \<in> F" and "prv (imp \<phi> \<chi>)"
shows "prv (imp (scnj F) \<chi>)"
unfolding scnj_def using assms by (intro prv_lcnj_imp) auto
lemma prv_imp_scnj:
assumes "\<chi> \<in> fmla" and "F \<subseteq> fmla" "finite F"
and "\<And>\<phi>. \<phi> \<in> F \<Longrightarrow> prv (imp \<chi> \<phi>)"
shows "prv (imp \<chi> (scnj F))"
unfolding scnj_def using assms by (intro prv_imp_lcnj) auto
lemma prv_scnj_mono:
assumes "F1 \<subseteq> fmla" and "F2 \<subseteq> F1" and "finite F1"
shows "prv (imp (scnj F1) (scnj F2))"
unfolding scnj_def using assms apply (intro prv_lcnj_mono)
subgoal by auto
subgoal by (metis asList infinite_super) .
lemma prv_scnj_mono_imp:
assumes "F1 \<subseteq> fmla" "F2 \<subseteq> fmla" "finite F1" "finite F2"
and "\<forall> \<phi>2 \<in> F2. \<exists> \<phi>1 \<in> F1. prv (imp \<phi>1 \<phi>2)"
shows "prv (imp (scnj F1) (scnj F2))"
unfolding scnj_def using assms by (intro prv_lcnj_mono_imp) auto
text \<open>Commutation with parallel substitution:\<close>
lemma prv_imp_scnj_insert:
assumes "F \<subseteq> fmla" and "finite F" and "\<phi> \<in> fmla"
shows "prv (imp (scnj (insert \<phi> F)) (cnj \<phi> (scnj F)))"
using assms apply (intro prv_imp_cnj)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by (auto intro: prv_imp_refl prv_scnj_imp)
subgoal by (auto intro: prv_scnj_mono) .
lemma prv_implies_scnj_insert:
assumes "F \<subseteq> fmla" and "finite F" and "\<phi> \<in> fmla"
and "prv (scnj (insert \<phi> F))"
shows "prv (cnj \<phi> (scnj F))"
by (meson assms cnj finite.insertI insert_subset prv_imp_mp prv_imp_scnj_insert scnj)
lemma prv_imp_cnj_scnj:
assumes "F \<subseteq> fmla" and "finite F" and "\<phi> \<in> fmla"
shows "prv (imp (cnj \<phi> (scnj F)) (scnj (insert \<phi> F)))"
using assms
by (auto intro!: prv_imp_scnj prv_imp_cnjL
simp: prv_cnj_imp_monoR2 prv_imp_triv prv_scnj_imp_in subset_iff)
lemma prv_implies_cnj_scnj:
assumes "F \<subseteq> fmla" and "finite F" and "\<phi> \<in> fmla"
and "prv (cnj \<phi> (scnj F))"
shows "prv (scnj (insert \<phi> F))"
by (meson assms cnj finite.insertI insert_subset prv_imp_cnj_scnj prv_imp_mp scnj)
lemma prv_eqv_scnj_insert:
assumes "F \<subseteq> fmla" and "finite F" and "\<phi> \<in> fmla"
shows "prv (eqv (scnj (insert \<phi> F)) (cnj \<phi> (scnj F)))"
by (simp add: assms prv_eqvI prv_imp_cnj_scnj prv_imp_scnj_insert)
lemma prv_scnj1_imp:
"\<phi> \<in> fmla \<Longrightarrow> prv (imp (scnj {\<phi>}) \<phi>)"
using prv_scnj_imp_in by auto
lemma prv_imp_scnj1:
"\<phi> \<in> fmla \<Longrightarrow> prv (imp \<phi> (scnj {\<phi>}))"
using prv_imp_refl prv_imp_scnj by fastforce
lemma prv_scnj2_imp_cnj:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> prv (imp (scnj {\<phi>,\<psi>}) (cnj \<phi> \<psi>))"
using prv_imp_cnj prv_scnj_imp_in by auto
lemma prv_cnj_imp_scnj2:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> prv (imp (cnj \<phi> \<psi>) (scnj {\<phi>,\<psi>}))"
using prv_imp_cnjL prv_imp_cnjR prv_imp_scnj by fastforce
lemma prv_imp_imp_scnj2:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> prv (imp \<phi> (imp \<psi> (scnj {\<phi>,\<psi>})))"
using prv_cnj_imp_scnj2 prv_cnj_imp by auto
(* *)
lemma prv_rawpsubst_scnj:
assumes "F \<subseteq> fmla" "finite F"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "prv (eqv (rawpsubst (scnj F) txs) (scnj ((\<lambda>\<phi>. rawpsubst \<phi> txs) ` F)))"
using assms proof(induction txs arbitrary: F)
case (Cons tx txs)
then obtain t x where tx[simp]: "tx = (t,x)" by (cases tx) auto
hence [simp]: "t \<in> trm" "x \<in> var" using Cons.prems by auto
have 0: "(\<lambda>\<phi>. rawpsubst (subst \<phi> t x) txs) ` F =
(\<lambda>\<phi>. rawpsubst \<phi> txs) ` ((\<lambda>\<phi>. subst \<phi> t x) ` F)"
using Cons.prems by auto
have "prv (eqv (subst (scnj F) t x)
(scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)))"
using Cons.prems by (intro prv_subst_scnj) auto
hence "prv (eqv (rawpsubst (subst (scnj F) t x) txs)
(rawpsubst (scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)) txs))"
using Cons.prems by (intro prv_eqv_rawpsubst) auto
moreover
have "prv (eqv (rawpsubst (scnj ((\<lambda>\<phi>. subst \<phi> t x) ` F)) txs)
(scnj ((\<lambda>\<phi>. rawpsubst (subst \<phi> t x) txs) ` F)))"
unfolding 0 using Cons.prems by (intro Cons.IH) auto
ultimately show ?case
using Cons.prems apply - by (rule prv_eqv_trans) (auto intro!: rawpsubst)
qed(auto simp: image_def prv_eqv_refl)[]
lemma prv_psubst_scnj:
assumes "F \<subseteq> fmla" "finite F"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "prv (eqv (psubst (scnj F) txs) (scnj ((\<lambda>\<phi>. psubst \<phi> txs) ` F)))"
proof-
define us where us: "us \<equiv> getFrN (map snd txs) (map fst txs) [scnj F] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> \<Union> (Fvars ` F) = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[scnj F]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[scnj F]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[scnj F]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[scnj F]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
define vs where vs: "vs \<equiv> \<lambda> \<phi>. getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
hence vss: "\<And>\<phi>. vs \<phi> = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)" by auto
{fix \<phi> assume "\<phi> \<in> F" hence "\<phi> \<in> fmla" using assms by auto
hence "set (vs \<phi>) \<subseteq> var \<and>
set (vs \<phi>) \<inter> Fvars \<phi> = {} \<and>
set (vs \<phi>) \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {} \<and>
set (vs \<phi>) \<inter> snd ` (set txs) = {} \<and>
length (vs \<phi>) = length txs \<and>
distinct (vs \<phi>)"
using assms unfolding vs
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply (intro conjI)
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
} note vs_facts = this
have [simp]: "\<And> x f. f \<in> F \<Longrightarrow> x \<in> set (vs f) \<Longrightarrow> x \<in> var"
using vs_facts
by (meson subsetD)
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs = "\<lambda> \<phi>. zip (map fst txs) (vs \<phi>)"
let ?vxs = "\<lambda> \<phi>. zip (map Var (vs \<phi>)) (map snd txs)"
let ?c = "rawpsubst (scnj F) ?uxs"
have c: "prv (eqv ?c
(scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)))"
using assms us_facts by (intro prv_rawpsubst_scnj) (auto intro!: rawpsubstT dest!: set_zip_D)
hence "prv (eqv (rawpsubst ?c ?tus)
(rawpsubst (scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)) ?tus))"
using assms us_facts
by (intro prv_eqv_rawpsubst) (auto intro!: rawpsubst dest!: set_zip_D)
moreover
have "prv (eqv (rawpsubst (scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)) ?tus)
(scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F))))"
using assms us_facts
by (intro prv_rawpsubst_scnj) (auto intro!: rawpsubst dest!: set_zip_D)
ultimately
have 0: "prv (eqv (rawpsubst ?c ?tus)
(scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F))))"
using assms us_facts apply - by (rule prv_eqv_trans) (auto intro!: rawpsubst dest!: set_zip_D)
moreover
have "prv (eqv (scnj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)))
(scnj ((\<lambda>\<phi>. rawpsubst (rawpsubst \<phi> (?vxs \<phi>)) (?tvs \<phi>)) ` F)))"
using assms us_facts vs_facts apply(intro prv_eqvI)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal apply(rule prv_scnj_mono_imp)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by auto
subgoal by auto
subgoal apply clarsimp
subgoal for \<phi> apply(rule bexI[of _ \<phi>]) apply(rule prv_imp_refl2)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (rule rawpsubst_compose_freshVar2)
(auto intro!: rawpsubst dest!: set_zip_D) . . .
subgoal apply(rule prv_scnj_mono_imp)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal apply clarsimp
subgoal for \<phi> apply(rule bexI[of _ \<phi>]) apply(rule prv_imp_refl2)
apply (auto intro!: rawpsubst dest!: set_zip_D)
apply(rule rawpsubst_compose_freshVar2)
apply (auto intro!: rawpsubst dest!: set_zip_D) . . . .
ultimately
have "prv (eqv (rawpsubst (rawpsubst (scnj F) ?uxs) ?tus)
(scnj ((\<lambda>\<phi>. rawpsubst (rawpsubst \<phi> (?vxs \<phi>)) (?tvs \<phi>)) ` F)))"
using assms us_facts
apply - by (rule prv_eqv_trans) (auto intro!: rawpsubst dest!: set_zip_D)
thus ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vss)
qed
lemma prv_imp_psubst_scnj:
assumes "F \<subseteq> fmla" "finite F" "snd ` set txs \<subseteq> var" "fst ` set txs \<subseteq> trm"
and "distinct (map snd txs)"
shows "prv (imp (psubst (scnj F) txs) (scnj ((\<lambda>\<phi>. psubst \<phi> txs) ` F)))"
using prv_psubst_scnj[OF assms] assms apply(intro prv_imp_eqvEL) by auto
lemma prv_psubst_scnj_imp:
assumes "F \<subseteq> fmla" "finite F" "snd ` set txs \<subseteq> var" "fst ` set txs \<subseteq> trm"
and "distinct (map snd txs)"
shows "prv (imp (scnj ((\<lambda>\<phi>. psubst \<phi> txs) ` F)) (psubst (scnj F) txs))"
using prv_psubst_scnj[OF assms] assms apply(intro prv_imp_eqvER) by auto
subsection \<open>Consistency and $\omega$-consistency\<close>
definition consistent :: bool where
"consistent \<equiv> \<not> prv fls"
lemma consistent_def2: "consistent \<longleftrightarrow> (\<exists>\<phi> \<in> fmla. \<not> prv \<phi>)"
unfolding consistent_def using prv_expl by blast
lemma consistent_def3: "consistent \<longleftrightarrow> (\<forall>\<phi> \<in> fmla. \<not> (prv \<phi> \<and> prv (neg \<phi>)))"
unfolding consistent_def using prv_neg_fls neg_def by auto
(* Omega-consistency: *)
definition \<omega>consistent :: bool where
"\<omega>consistent \<equiv>
\<forall> \<phi> \<in> fmla. \<forall> x \<in> var. Fvars \<phi> = {x} \<longrightarrow>
(\<forall> n \<in> num. prv (neg (subst \<phi> n x)))
\<longrightarrow>
\<not> prv (neg (neg (exi x \<phi>)))"
text \<open>The above particularly strong version of @{term \<omega>consistent} is used for the sake of working without
assuming classical logic. It of course implies the more standard formulations for classical logic:\<close>
definition \<omega>consistentStd1 :: bool where
"\<omega>consistentStd1 \<equiv>
\<forall> \<phi> \<in> fmla. \<forall> x \<in> var. Fvars \<phi> = {x} \<longrightarrow>
(\<forall> n \<in> num. prv (neg (subst \<phi> n x))) \<longrightarrow> \<not> prv (exi x \<phi>)"
definition \<omega>consistentStd2 :: bool where
"\<omega>consistentStd2 \<equiv>
\<forall> \<phi> \<in> fmla. \<forall> x \<in> var. Fvars \<phi> = {x} \<longrightarrow>
(\<forall> n \<in> num. prv (subst \<phi> n x)) \<longrightarrow> \<not> prv (exi x (neg \<phi>))"
lemma \<omega>consistent_impliesStd1:
"\<omega>consistent \<Longrightarrow>
\<omega>consistentStd1"
unfolding \<omega>consistent_def \<omega>consistentStd1_def using prv_neg_neg by blast
lemma \<omega>consistent_impliesStd2:
"\<omega>consistent \<Longrightarrow>
\<omega>consistentStd2"
by (auto dest!: \<omega>consistent_impliesStd1 bspec[of _ _ "neg _"]
simp: \<omega>consistentStd1_def \<omega>consistentStd2_def prv_neg_neg)
text \<open>In the presence of classical logic deduction, the stronger condition is
equivalent to the standard ones:\<close>
lemma \<omega>consistent_iffStd1:
assumes "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> prv (imp (neg (neg \<phi>)) \<phi>)"
shows "\<omega>consistent \<longleftrightarrow> \<omega>consistentStd1"
apply standard
subgoal using \<omega>consistent_impliesStd1 by auto
subgoal unfolding \<omega>consistentStd1_def \<omega>consistent_def
by (meson assms exi neg prv_imp_mp) .
lemma \<omega>consistent_iffStd2:
assumes "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> prv (imp (neg (neg \<phi>)) \<phi>)"
shows "\<omega>consistent \<longleftrightarrow> \<omega>consistentStd2"
unfolding \<omega>consistent_iffStd1[OF assms, simplified]
\<omega>consistentStd1_def \<omega>consistentStd2_def apply safe
subgoal for \<phi> x
by (auto simp: prv_neg_neg dest: bspec[of _ _ "neg _"])
subgoal for \<phi> x
using prv_exi_congW prv_imp_neg_fls
by (auto simp: neg_def prv_neg_neg dest!: bspec[of _ _ "neg _"]) .
text \<open>$\omega$-consistency implies consistency:\<close>
lemma \<omega>consistentStd1_implies_consistent:
assumes "\<omega>consistentStd1"
shows "consistent"
unfolding consistent_def
proof safe
assume pf: "prv fls"
then obtain x where x: "x \<in> var" "x \<notin> Fvars fls"
using finite_Fvars getFresh by auto
let ?fls = "cnj (fls) (eql (Var x) (Var x))"
have 0: "\<forall> n \<in> num. prv (neg (subst ?fls n x))" and 1: "prv (exi x fls)"
using x fls by (auto simp: pf prv_expl)
have 2: "\<not> prv (exi x ?fls)" using 0 fls x assms
unfolding \<omega>consistentStd1_def by simp
show False using 1 2 consistent_def consistent_def2 pf x(1) by blast
qed
lemma \<omega>consistentStd2_implies_consistent:
assumes "\<omega>consistentStd2"
shows "consistent"
unfolding consistent_def
proof safe
assume pf: "prv fls"
then obtain x where x: "x \<in> var" "x \<notin> Fvars fls"
using finite_Fvars getFresh by auto
let ?fls = "cnj (fls) (eql (Var x) (Var x))"
have 0: "\<forall> n \<in> num. prv (subst ?fls n x)" and 1: "prv (exi x (neg ?fls))"
using x fls by (auto simp: pf prv_expl)
have 2: "\<not> prv (exi x (neg ?fls))" using 0 fls x assms
unfolding \<omega>consistentStd2_def by auto
show False using 1 2 by simp
qed
corollary \<omega>consistent_implies_consistent:
assumes "\<omega>consistent"
shows "consistent"
by (simp add: \<omega>consistentStd2_implies_consistent \<omega>consistent_impliesStd2 assms)
end \<comment> \<open>context @{locale Deduct_with_False}\<close>
section \<open>Deduction Considering False and Disjunction\<close>
locale Deduct_with_False_Disj =
Syntax_with_Connectives_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
+
Deduct_with_False
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
num
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
+
assumes
prv_dsj_impL:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp \<phi> (dsj \<phi> \<chi>))"
and
prv_dsj_impR:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
prv (imp \<chi> (dsj \<phi> \<chi>))"
and
prv_imp_dsjE:
"\<And> \<phi> \<chi> \<psi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
prv (imp (imp \<phi> \<psi>) (imp (imp \<chi> \<psi>) (imp (dsj \<phi> \<chi>) \<psi>)))"
begin
lemma prv_imp_dsjEE:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
assumes " prv (imp \<phi> \<psi>)" and "prv (imp \<chi> \<psi>)"
shows "prv (imp (dsj \<phi> \<chi>) \<psi>)"
by (metis assms dsj imp prv_imp_dsjE prv_imp_mp)
lemma prv_dsj_cases:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla" "\<chi> \<in> fmla"
and "prv (dsj \<phi>1 \<phi>2)" and "prv (imp \<phi>1 \<chi>)" and "prv (imp \<phi>2 \<chi>)"
shows "prv \<chi>"
by (meson assms dsj prv_imp_dsjEE prv_imp_mp)
subsection \<open>Disjunction vs. disjunction\<close>
lemma prv_dsj_com_imp:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla"
shows "prv (imp (dsj \<phi> \<chi>) (dsj \<chi> \<phi>))"
by (metis \<chi> \<phi> dsj imp prv_dsj_impL prv_dsj_impR prv_imp_dsjE prv_imp_mp)
lemma prv_dsj_com:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla"
shows "prv (eqv (dsj \<phi> \<chi>) (dsj \<chi> \<phi>))"
by (simp add: prv_dsj_com_imp prv_eqvI)
lemma prv_dsj_assoc_imp1:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (imp (dsj \<phi> (dsj \<chi> \<psi>)) (dsj (dsj \<phi> \<chi>) \<psi>))"
proof -
have f1: "\<And>f fa fb. f \<notin> fmla \<or> \<not> prv (imp fa fb) \<or> fb \<notin> fmla \<or> fa \<notin> fmla \<or> prv (imp fa (dsj fb f))"
by (meson dsj prv_dsj_impL prv_prv_imp_trans)
have "prv (imp \<phi> (dsj \<phi> \<chi>))"
by (simp add: prv_dsj_impL)
then show ?thesis
using f1 \<chi> \<phi> \<psi> dsj prv_dsj_impR prv_imp_dsjEE by presburger
qed
lemma prv_dsj_assoc_imp2:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (imp (dsj (dsj \<phi> \<chi>) \<psi>) (dsj \<phi> (dsj \<chi> \<psi>)))"
proof -
have f1: "\<forall>f fa fb. (((prv (imp f (dsj fa fb)) \<or> \<not> prv (imp f (dsj fb fa))) \<or> f \<notin> fmla) \<or> fa \<notin> fmla) \<or> fb \<notin> fmla"
by (meson dsj prv_dsj_com_imp prv_prv_imp_trans)
have "\<forall>f fa fb. (((prv (imp f (dsj fa fb)) \<or> \<not> prv (imp f fa)) \<or> fa \<notin> fmla) \<or> f \<notin> fmla) \<or> fb \<notin> fmla"
by (meson dsj prv_dsj_impL prv_prv_imp_trans)
then show ?thesis
using f1 by (metis \<chi> \<phi> \<psi> dsj prv_dsj_impR prv_imp_dsjEE)
qed
lemma prv_dsj_assoc:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>[simp]: "\<chi> \<in> fmla" and \<psi>[simp]: "\<psi> \<in> fmla"
shows "prv (eqv (dsj \<phi> (dsj \<chi> \<psi>)) (dsj (dsj \<phi> \<chi>) \<psi>))"
by (simp add: prv_dsj_assoc_imp1 prv_dsj_assoc_imp2 prv_eqvI)
lemma prv_dsj_com_imp3:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla" "\<phi>3 \<in> fmla"
shows "prv (imp (dsj \<phi>1 (dsj \<phi>2 \<phi>3))
(dsj \<phi>2 (dsj \<phi>1 \<phi>3)))"
proof -
have "\<forall>f fa fb. (((prv (imp f (dsj fb fa)) \<or> \<not> prv (imp f fa)) \<or> fa \<notin> fmla) \<or> f \<notin> fmla) \<or> fb \<notin> fmla"
by (meson dsj prv_dsj_impR prv_prv_imp_trans)
then show ?thesis
by (meson assms(1) assms(2) assms(3) dsj prv_dsj_impL prv_dsj_impR prv_imp_dsjEE)
qed
lemma prv_dsj_mono:
"\<phi>1 \<in> fmla \<Longrightarrow> \<phi>2 \<in> fmla \<Longrightarrow> \<chi>1 \<in> fmla \<Longrightarrow> \<chi>2 \<in> fmla \<Longrightarrow>
prv (imp \<phi>1 \<chi>1) \<Longrightarrow> prv (imp \<phi>2 \<chi>2) \<Longrightarrow> prv (imp (dsj \<phi>1 \<phi>2) (dsj \<chi>1 \<chi>2))"
by (meson dsj prv_dsj_impL prv_dsj_impR prv_imp_dsjEE prv_prv_imp_trans)
subsection \<open>Disjunction vs. conjunction\<close>
lemma prv_cnj_dsj_distrib1:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (imp (cnj \<phi> (dsj \<chi>1 \<chi>2)) (dsj (cnj \<phi> \<chi>1) (cnj \<phi> \<chi>2)))"
by (simp add: prv_cnj_imp prv_cnj_imp_monoR2 prv_dsj_impL prv_dsj_impR prv_imp_com prv_imp_dsjEE)
lemma prv_cnj_dsj_distrib2:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (imp (dsj (cnj \<phi> \<chi>1) (cnj \<phi> \<chi>2)) (cnj \<phi> (dsj \<chi>1 \<chi>2)))"
by (simp add: prv_cnj_mono prv_dsj_impL prv_dsj_impR prv_imp_dsjEE prv_imp_refl)
lemma prv_cnj_dsj_distrib:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (eqv (cnj \<phi> (dsj \<chi>1 \<chi>2)) (dsj (cnj \<phi> \<chi>1) (cnj \<phi> \<chi>2)))"
by (simp add: prv_cnj_dsj_distrib1 prv_cnj_dsj_distrib2 prv_eqvI)
lemma prv_dsj_cnj_distrib1:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (imp (dsj \<phi> (cnj \<chi>1 \<chi>2)) (cnj (dsj \<phi> \<chi>1) (dsj \<phi> \<chi>2)))"
by (simp add: prv_cnj_mono prv_dsj_impL prv_dsj_impR prv_imp_cnj prv_imp_dsjEE)
lemma prv_dsj_cnj_distrib2:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (imp (cnj (dsj \<phi> \<chi>1) (dsj \<phi> \<chi>2)) (dsj \<phi> (cnj \<chi>1 \<chi>2)))"
proof -
have "\<forall>f fa fb. (((prv (imp f (imp fb fa)) \<or> \<not> prv (imp f fa)) \<or> fa \<notin> fmla) \<or> f \<notin> fmla) \<or> fb \<notin> fmla"
by (meson imp prv_imp_imp_triv prv_prv_imp_trans)
then show ?thesis
by (metis \<chi>1 \<chi>2 \<phi> cnj dsj imp prv_cnj_imp prv_cnj_imp_monoR2 prv_dsj_impL prv_dsj_impR
prv_imp_com prv_imp_dsjEE)
qed
lemma prv_dsj_cnj_distrib:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and \<chi>1[simp]: "\<chi>1 \<in> fmla" and \<chi>2[simp]: "\<chi>2 \<in> fmla"
shows "prv (eqv (dsj \<phi> (cnj \<chi>1 \<chi>2)) (cnj (dsj \<phi> \<chi>1) (dsj \<phi> \<chi>2)))"
by (simp add: prv_dsj_cnj_distrib1 prv_dsj_cnj_distrib2 prv_eqvI)
subsection \<open>Disjunction vs. True and False\<close>
lemma prv_dsjR_fls: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (dsj fls \<phi>) \<phi>)"
by (simp add: prv_dsj_impR prv_eqvI prv_imp_dsjEE prv_imp_refl)
lemma prv_dsjL_fls: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (dsj \<phi> fls) \<phi>)"
by (simp add: prv_dsj_impL prv_eqvI prv_imp_dsjEE prv_imp_refl)
lemma prv_dsjR_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (dsj tru \<phi>) tru)"
by (simp add: prv_dsj_impL prv_eqvI prv_imp_tru)
lemma prv_dsjL_tru: "\<phi> \<in> fmla \<Longrightarrow> prv (eqv (dsj \<phi> tru) tru)"
by (simp add: prv_dsj_impR prv_eqvI prv_imp_tru)
subsection \<open>Set-based disjunctions\<close>
text \<open>Just like for conjunctions, these are based on properties of the auxiliary
list disjunctions.\<close>
lemma prv_imp_ldsj_in:
assumes "set \<phi>s \<subseteq> fmla"
and "\<phi> \<in> set \<phi>s"
shows "prv (imp \<phi> (ldsj \<phi>s))"
proof-
have "\<phi> \<in> fmla" using assms by auto
thus ?thesis
using assms apply(induct \<phi>s arbitrary: \<phi>)
subgoal by auto
subgoal by (simp add: prv_dsj_impL)
(meson dsj ldsj prv_dsj_impL prv_dsj_impR prv_prv_imp_trans) .
qed
lemma prv_imp_ldsj:
assumes "\<chi> \<in> fmla" and "set \<phi>s \<subseteq> fmla"
and "\<phi> \<in> set \<phi>s" and "prv (imp \<chi> \<phi>)"
shows "prv (imp \<chi> (ldsj \<phi>s))"
using assms ldsj prv_imp_ldsj_in prv_prv_imp_trans by blast
lemma prv_ldsj_imp:
assumes "\<chi> \<in> fmla" and "set \<phi>s \<subseteq> fmla"
and "\<And>\<phi>. \<phi> \<in> set \<phi>s \<Longrightarrow> prv (imp \<phi> \<chi>)"
shows "prv (imp (ldsj \<phi>s) \<chi>)"
using assms
by (induct \<phi>s arbitrary: \<chi>)
(auto simp add: prv_imp_tru prv_imp_com prv_imp_dsjEE)
lemma prv_ldsj_imp_inner:
assumes "\<phi> \<in> fmla" "set \<phi>1s \<subseteq> fmla" "set \<phi>2s \<subseteq> fmla"
shows "prv (imp (ldsj (\<phi>1s @ \<phi> # \<phi>2s)) (dsj \<phi> (ldsj (\<phi>1s @ \<phi>2s))))"
using assms proof(induction \<phi>1s)
case (Cons \<phi>1 \<phi>1s)
have [intro!]: "set (\<phi>1s @ \<phi>2s) \<subseteq> fmla" "set (\<phi>1s @ \<phi> # \<phi>2s) \<subseteq> fmla" using Cons by auto
have 0: "prv (imp (dsj \<phi>1 (dsj \<phi> (ldsj (\<phi>1s @ \<phi>2s))))
(dsj \<phi> (dsj \<phi>1 (ldsj (\<phi>1s @ \<phi>2s)))))"
using Cons by (intro prv_dsj_com_imp3) fastforce+
have 1: "prv (imp (dsj \<phi>1 (ldsj (\<phi>1s @ \<phi> # \<phi>2s)))
(dsj \<phi>1 (dsj \<phi> (ldsj (\<phi>1s @ \<phi>2s)))))"
using Cons by (intro prv_dsj_mono) (auto simp add: prv_imp_refl)
show ?case using prv_prv_imp_trans[OF _ _ _ 1 0] Cons by auto
qed(simp add: prv_imp_refl)
lemma prv_ldsj_imp_remdups:
assumes "set \<phi>s \<subseteq> fmla"
shows "prv (imp (ldsj \<phi>s) (ldsj (remdups \<phi>s)))"
using assms apply(induct \<phi>s)
subgoal by auto
subgoal by (metis ldsj prv_imp_ldsj_in prv_ldsj_imp set_remdups) .
lemma prv_ldsj_mono:
assumes \<phi>2s: "set \<phi>2s \<subseteq> fmla" and "set \<phi>1s \<subseteq> set \<phi>2s"
shows "prv (imp (ldsj \<phi>1s) (ldsj \<phi>2s))"
proof-
define \<phi>1s' where \<phi>1s': "\<phi>1s' = remdups \<phi>1s"
have A: "set \<phi>1s' \<subseteq> set \<phi>2s" "distinct \<phi>1s'" unfolding \<phi>1s' using assms by auto
have B: "prv (imp (ldsj \<phi>1s') (ldsj \<phi>2s))"
using \<phi>2s A proof(induction \<phi>2s arbitrary: \<phi>1s')
case (Cons \<phi>2 \<phi>2s \<phi>1ss)
show ?case proof(cases "\<phi>2 \<in> set \<phi>1ss")
case True
then obtain \<phi>1ss1 \<phi>1ss2 where \<phi>1ss: "\<phi>1ss = \<phi>1ss1 @ \<phi>2 # \<phi>1ss2"
by (meson split_list)
define \<phi>1s where \<phi>1s: "\<phi>1s \<equiv> \<phi>1ss1 @ \<phi>1ss2"
- have nin: "\<phi>2 \<notin> set \<phi>1s" using \<phi>1ss `distinct \<phi>1ss` unfolding \<phi>1s by auto
+ have nin: "\<phi>2 \<notin> set \<phi>1s" using \<phi>1ss \<open>distinct \<phi>1ss\<close> unfolding \<phi>1s by auto
have [intro!,simp]: "set \<phi>1s \<subseteq> fmla" unfolding \<phi>1s using \<phi>1ss Cons by auto
have 0: "prv (imp (ldsj \<phi>1ss) (dsj \<phi>2 (ldsj \<phi>1s)))"
unfolding \<phi>1s \<phi>1ss
apply(rule prv_ldsj_imp_inner) using Cons \<phi>1ss by auto
have 1: "prv (imp (ldsj \<phi>1s) (ldsj \<phi>2s))" apply(rule Cons.IH) using nin Cons.prems True
using \<phi>1s \<phi>1ss by auto
have 2: "prv (imp (dsj \<phi>2 (ldsj \<phi>1s)) (dsj \<phi>2 (ldsj \<phi>2s)))"
using Cons \<phi>1ss \<phi>1s 1 apply(intro prv_dsj_mono)
using prv_imp_refl by auto blast+
show ?thesis using Cons by (auto intro!: prv_prv_imp_trans[OF _ _ _ 0 2])
next
case False
then show ?thesis using Cons
by (meson ldsj order.trans prv_imp_ldsj_in prv_ldsj_imp subset_code(1))
qed
qed(simp add: prv_imp_refl)
have C: "prv (imp (ldsj \<phi>1s) (ldsj \<phi>1s'))"
unfolding \<phi>1s' using assms by (intro prv_ldsj_imp_remdups) auto
show ?thesis using A assms by (intro prv_prv_imp_trans[OF _ _ _ C B]) auto
qed
lemma prv_ldsj_eqv:
assumes "set \<phi>1s \<subseteq> fmla" and "set \<phi>2s = set \<phi>1s"
shows "prv (eqv (ldsj \<phi>1s) (ldsj \<phi>2s))"
using assms prv_ldsj_mono by (intro prv_eqvI) auto
lemma prv_ldsj_mono_imp:
assumes "set \<phi>1s \<subseteq> fmla" "set \<phi>2s \<subseteq> fmla" and "\<forall> \<phi>1 \<in> set \<phi>1s. \<exists> \<phi>2 \<in> set \<phi>2s. prv (imp \<phi>1 \<phi>2)"
shows "prv (imp (ldsj \<phi>1s) (ldsj \<phi>2s))"
using assms apply(intro prv_ldsj_imp)
subgoal by auto
subgoal by auto
subgoal using prv_imp_ldsj by blast .
text \<open>Just like set-based conjunction, set-based disjunction commutes with substitution
only up to provably equivalence:\<close>
lemma prv_subst_sdsj:
"F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
prv (eqv (subst (sdsj F) t x) (sdsj ((\<lambda>\<phi>. subst \<phi> t x) ` F)))"
unfolding sdsj_def by (fastforce intro!: prv_ldsj_eqv)
lemma prv_imp_sdsj_in:
assumes "\<phi> \<in> fmla" and "F \<subseteq> fmla" "finite F"
and "\<phi> \<in> F"
shows "prv (imp \<phi> (sdsj F))"
unfolding sdsj_def using assms by (intro prv_imp_ldsj_in) auto
lemma prv_imp_sdsj:
assumes "\<chi> \<in> fmla" and "F \<subseteq> fmla" "finite F"
and "\<phi> \<in> F" and "prv (imp \<chi> \<phi>)"
shows "prv (imp \<chi> (sdsj F))"
unfolding sdsj_def using assms by (intro prv_imp_ldsj) auto
lemma prv_sdsj_imp:
assumes "\<chi> \<in> fmla" and "F \<subseteq> fmla" "finite F"
and "\<And>\<phi>. \<phi> \<in> F \<Longrightarrow> prv (imp \<phi> \<chi>)"
shows "prv (imp (sdsj F) \<chi>)"
unfolding sdsj_def using assms by (intro prv_ldsj_imp) auto
lemma prv_sdsj_mono:
assumes "F2 \<subseteq> fmla" and "F1 \<subseteq> F2" and "finite F2"
shows "prv (imp (sdsj F1) (sdsj F2))"
unfolding sdsj_def using assms apply(intro prv_ldsj_mono)
subgoal by auto
subgoal by (metis asList infinite_super) .
lemma prv_sdsj_mono_imp:
assumes "F1 \<subseteq> fmla" "F2 \<subseteq> fmla" "finite F1" "finite F2"
and "\<forall> \<phi>1 \<in> F1. \<exists> \<phi>2 \<in> F2. prv (imp \<phi>1 \<phi>2)"
shows "prv (imp (sdsj F1) (sdsj F2))"
unfolding sdsj_def using assms by (intro prv_ldsj_mono_imp) auto
lemma prv_sdsj_cases:
assumes "F \<subseteq> fmla" "finite F" "\<psi> \<in> fmla"
and "prv (sdsj F)" and "\<And> \<phi>. \<phi> \<in> F \<Longrightarrow> prv (imp \<phi> \<psi>)"
shows "prv \<psi>"
by (meson assms prv_imp_mp prv_sdsj_imp sdsj)
lemma prv_sdsj1_imp:
"\<phi> \<in> fmla \<Longrightarrow> prv (imp (sdsj {\<phi>}) \<phi>)"
using prv_imp_refl prv_sdsj_imp by fastforce
lemma prv_imp_sdsj1:
"\<phi> \<in> fmla \<Longrightarrow> prv (imp \<phi> (sdsj {\<phi>}))"
using prv_imp_refl prv_imp_sdsj by fastforce
lemma prv_sdsj2_imp_dsj:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> prv (imp (sdsj {\<phi>,\<psi>}) (dsj \<phi> \<psi>))"
using prv_dsj_impL prv_dsj_impR prv_sdsj_imp by fastforce
lemma prv_dsj_imp_sdsj2:
"\<phi> \<in> fmla \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> prv (imp (dsj \<phi> \<psi>) (sdsj {\<phi>,\<psi>}))"
by (simp add: prv_imp_dsjEE prv_imp_sdsj_in)
text \<open>Commutation with parallel substitution:\<close>
lemma prv_rawpsubst_sdsj:
assumes "F \<subseteq> fmla" "finite F"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "prv (eqv (rawpsubst (sdsj F) txs) (sdsj ((\<lambda>\<phi>. rawpsubst \<phi> txs) ` F)))"
using assms proof(induction txs arbitrary: F)
case (Cons tx txs)
then obtain t x where tx[simp]: "tx = (t,x)" by (cases tx) auto
hence [simp]: "t \<in> trm" "x \<in> var" using Cons.prems by auto
have 0: "(\<lambda>\<phi>. rawpsubst (subst \<phi> t x) txs) ` F =
(\<lambda>\<phi>. rawpsubst \<phi> txs) ` ((\<lambda>\<phi>. subst \<phi> t x) ` F)"
using Cons.prems by auto
have "prv (eqv (subst (sdsj F) t x)
(sdsj ((\<lambda>\<phi>. subst \<phi> t x) ` F)))"
using Cons.prems by (intro prv_subst_sdsj) auto
hence "prv (eqv (rawpsubst (subst (sdsj F) t x) txs)
(rawpsubst (sdsj ((\<lambda>\<phi>. subst \<phi> t x) ` F)) txs))"
using Cons.prems by (intro prv_eqv_rawpsubst) auto
moreover
have "prv (eqv (rawpsubst (sdsj ((\<lambda>\<phi>. subst \<phi> t x) ` F)) txs)
(sdsj ((\<lambda>\<phi>. rawpsubst (subst \<phi> t x) txs) ` F)))"
unfolding 0 using Cons.prems by (intro Cons.IH) auto
ultimately show ?case
using Cons.prems apply - by (rule prv_eqv_trans) (auto intro!: rawpsubst)
qed(auto simp: image_def prv_eqv_refl)[]
lemma prv_psubst_sdsj:
assumes "F \<subseteq> fmla" "finite F"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "prv (eqv (psubst (sdsj F) txs) (sdsj ((\<lambda>\<phi>. psubst \<phi> txs) ` F)))"
proof-
define us where us: "us \<equiv> getFrN (map snd txs) (map fst txs) [sdsj F] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> \<Union> (Fvars ` F) = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[sdsj F]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[sdsj F]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[sdsj F]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[sdsj F]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
define vs where vs: "vs \<equiv> \<lambda> \<phi>. getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
hence vss: "\<And>\<phi>. vs \<phi> = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)" by auto
{fix \<phi> assume "\<phi> \<in> F" hence "\<phi> \<in> fmla" using assms by auto
hence "set (vs \<phi>) \<subseteq> var \<and>
set (vs \<phi>) \<inter> Fvars \<phi> = {} \<and>
set (vs \<phi>) \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {} \<and>
set (vs \<phi>) \<inter> snd ` (set txs) = {} \<and>
length (vs \<phi>) = length txs \<and>
distinct (vs \<phi>)"
using assms unfolding vs
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply(intro conjI)
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
} note vs_facts = this
have [simp]: "\<And> x f. f \<in> F \<Longrightarrow> x \<in> set (vs f) \<Longrightarrow> x \<in> var"
using vs_facts by (meson subsetD)
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs = "\<lambda> \<phi>. zip (map fst txs) (vs \<phi>)"
let ?vxs = "\<lambda> \<phi>. zip (map Var (vs \<phi>)) (map snd txs)"
let ?c = "rawpsubst (sdsj F) ?uxs"
have c: "prv (eqv ?c
(sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)))"
using assms us_facts
by (intro prv_rawpsubst_sdsj) (auto intro!: rawpsubstT dest!: set_zip_D)
hence "prv (eqv (rawpsubst ?c ?tus)
(rawpsubst (sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)) ?tus))"
using assms us_facts by (intro prv_eqv_rawpsubst) (auto intro!: rawpsubst dest!: set_zip_D)
moreover
have "prv (eqv (rawpsubst (sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)) ?tus)
(sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F))))"
using assms us_facts by (intro prv_rawpsubst_sdsj) (auto intro!: rawpsubst dest!: set_zip_D)
ultimately
have 0: "prv (eqv (rawpsubst ?c ?tus)
(sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F))))"
using assms us_facts apply- by (rule prv_eqv_trans) (auto intro!: rawpsubst dest!: set_zip_D)
moreover
have "prv (eqv (sdsj ((\<lambda>\<phi>. rawpsubst \<phi> ?tus) ` ((\<lambda>\<phi>. rawpsubst \<phi> ?uxs) ` F)))
(sdsj ((\<lambda>\<phi>. rawpsubst (rawpsubst \<phi> (?vxs \<phi>)) (?tvs \<phi>)) ` F)))"
using assms us_facts vs_facts apply(intro prv_eqvI)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal apply(rule prv_sdsj_mono_imp)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by auto
subgoal by auto
subgoal apply clarsimp
subgoal for \<phi> apply(rule bexI[of _ \<phi>]) apply(rule prv_imp_refl2)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (rule rawpsubst_compose_freshVar2)
(auto intro!: rawpsubst dest!: set_zip_D) . . .
subgoal apply(rule prv_sdsj_mono_imp)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal by (auto intro!: rawpsubst dest!: set_zip_D)
subgoal apply clarsimp
subgoal for \<phi> apply(rule bexI[of _ \<phi>]) apply(rule prv_imp_refl2)
apply (auto intro!: rawpsubst dest!: set_zip_D)
apply(rule rawpsubst_compose_freshVar2)
apply (auto intro!: rawpsubst dest!: set_zip_D) . . . .
ultimately
have "prv (eqv (rawpsubst (rawpsubst (sdsj F) ?uxs) ?tus)
(sdsj ((\<lambda>\<phi>. rawpsubst (rawpsubst \<phi> (?vxs \<phi>)) (?tvs \<phi>)) ` F)))"
using assms us_facts
apply- by (rule prv_eqv_trans) (auto intro!: rawpsubst dest!: set_zip_D)
thus ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vss)
qed
end \<comment> \<open>context @{locale Deduct_with_False_Disj}\<close>
section \<open>Deduction with Quantified Variable Renaming Included\<close>
locale Deduct_with_False_Disj_Rename =
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
+
Syntax_with_Connectives_Rename
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
section \<open>Deduction with PseudoOrder Axioms Included\<close>
text \<open>We assume a two-variable formula Lq that satisfies two axioms
resembling the properties of the strict or nonstrict ordering on naturals.
The choice of these axioms is motivated by an abstract account of Rosser's trick
to improve on Gödel's First Incompleteness Theorem, reported in our
CADE 2019 paper~\cite{DBLP:conf/cade/0001T19}.\<close>
locale Deduct_with_PseudoOrder =
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
+
Syntax_PseudoOrder
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
Lq
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and prv
and Lq
+
assumes
Lq_num:
"let LLq = (\<lambda> t1 t2. psubst Lq [(t1,zz), (t2,yy)]) in
\<forall> \<phi> \<in> fmla. \<forall> q \<in> num. Fvars \<phi> = {zz} \<and> (\<forall> p \<in> num. prv (subst \<phi> p zz))
\<longrightarrow> prv (all zz (imp (LLq (Var zz) q) \<phi>))"
and
Lq_num2:
"let LLq = (\<lambda> t1 t2. psubst Lq [(t1,zz), (t2,yy)]) in
\<forall> p \<in> num. \<exists> P \<subseteq> num. finite P \<and> prv (dsj (sdsj {eql (Var yy) r | r. r \<in> P}) (LLq p (Var yy)))"
begin
lemma LLq_num:
assumes "\<phi> \<in> fmla" "q \<in> num" "Fvars \<phi> = {zz}" "\<forall> p \<in> num. prv (subst \<phi> p zz)"
shows "prv (all zz (imp (LLq (Var zz) q) \<phi>))"
using assms Lq_num unfolding LLq_def by auto
lemma LLq_num2:
assumes "p \<in> num"
shows "\<exists> P \<subseteq> num. finite P \<and> prv (dsj (sdsj {eql (Var yy) r | r. r \<in> P}) (LLq p (Var yy)))"
using assms Lq_num2 unfolding LLq_def by auto
end \<comment> \<open>context @{locale Deduct_with_PseudoOrder}\<close>
(*<*)
end
(*>*)
\ No newline at end of file
diff --git a/thys/Syntax_Independent_Logic/Deduction_Q.thy b/thys/Syntax_Independent_Logic/Deduction_Q.thy
--- a/thys/Syntax_Independent_Logic/Deduction_Q.thy
+++ b/thys/Syntax_Independent_Logic/Deduction_Q.thy
@@ -1,1324 +1,1324 @@
chapter \<open>Deduction in a System Embedding the Intuitionistic Robinson Arithmetic\<close>
(*<*)
theory Deduction_Q imports Syntax_Arith Natural_Deduction
begin
(*>*)
text \<open>NB: Robinson arithmetic, also know as system Q, is Peano arithmetic without the
induction axiom schema.\<close>
section \<open>Natural Deduction with the Bounded Quantifiers\<close>
text \<open>We start by simply putting together deduction with the arithmetic syntax,
which allows us to reason about bounded quantifiers:\<close>
locale Deduct_with_False_Disj_Arith =
Syntax_Arith
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
zer suc pls tms
+
Deduct_with_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and zer suc pls tms
and prv
begin
lemma nprv_ballI:
"nprv (insert (LLq (Var x) t) F) \<phi> \<Longrightarrow>
F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
x \<notin> (\<Union>\<phi> \<in> F. Fvars \<phi>) \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow>
nprv F (ball x t \<phi>)"
unfolding ball_def
by(nprover2 r1: nprv_allI r2: nprv_impI)
lemma nprv_ballE_aux:
"nprv F (ball x t \<phi>) \<Longrightarrow> nprv F (LLq t1 t) \<Longrightarrow>
F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> x \<in> var \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow>
nprv F (subst \<phi> t1 x)"
unfolding ball_def
by(nprover3 r1: nprv_allE[of _ x "imp (LLq (Var x) t) \<phi>" t1]
r2: nprv_impE0[of "LLq t1 t" "subst \<phi> t1 x"]
r3: nprv_mono[of F])
lemma nprv_ballE:
"nprv F (ball x t \<phi>) \<Longrightarrow> nprv F (LLq t1 t) \<Longrightarrow> nprv (insert (subst \<phi> t1 x) F) \<psi> \<Longrightarrow>
F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> x \<in> var \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
x \<notin> FvarsT t \<Longrightarrow>
nprv F \<psi>"
by (meson atrm_trm local.subst nprv_ballE_aux nprv_cut rev_subsetD)
lemmas nprv_ballE0 = nprv_ballE[OF nprv_hyp _ _, simped]
lemmas nprv_ballE1 = nprv_ballE[OF _ nprv_hyp _, simped]
lemmas nprv_ballE2 = nprv_ballE[OF _ _ nprv_hyp, simped]
lemmas nprv_ballE01 = nprv_ballE[OF nprv_hyp nprv_hyp _, simped]
lemmas nprv_ballE02 = nprv_ballE[OF nprv_hyp _ nprv_hyp, simped]
lemmas nprv_ballE12 = nprv_ballE[OF _ nprv_hyp nprv_hyp, simped]
lemmas nprv_ballE012 = nprv_ballE[OF nprv_hyp nprv_hyp nprv_hyp, simped]
lemma nprv_bexiI:
"nprv F (subst \<phi> t1 x) \<Longrightarrow> nprv F (LLq t1 t) \<Longrightarrow>
F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> x \<in> var \<Longrightarrow>
x \<notin> FvarsT t \<Longrightarrow>
nprv F (bexi x t \<phi>)"
unfolding bexi_def
by (nprover2 r1: nprv_exiI[of _ _ t1] r2: nprv_cnjI)
lemma nprv_bexiE:
"nprv F (bexi x t \<phi>) \<Longrightarrow> nprv (insert (LLq (Var x) t) (insert \<phi> F)) \<psi> \<Longrightarrow>
F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow>
x \<notin> (\<Union>\<phi> \<in> F. Fvars \<phi>) \<Longrightarrow> x \<notin> Fvars \<psi> \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow>
nprv F \<psi>"
unfolding bexi_def
by (nprover2 r1: nprv_exiE[of _ x "cnj (LLq (Var x) t) \<phi>"] r2: nprv_cnjH)
lemmas nprv_bexiE0 = nprv_bexiE[OF nprv_hyp _, simped]
lemmas nprv_bexiE1 = nprv_bexiE[OF _ nprv_hyp, simped]
lemmas nprv_bexiE01 = nprv_bexiE[OF nprv_hyp nprv_hyp, simped]
end \<comment> \<open>context @{locale Deduct_with_False_Disj}\<close>
section \<open>Deduction with the Robinson Arithmetic Axioms\<close>
locale Deduct_Q =
Deduct_with_False_Disj_Arith
var trm fmla
Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
zer suc pls tms
prv
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
and zer suc pls tms
and prv
+
assumes
\<comment> \<open>The Q axioms are stated for some fixed variables;
we will prove more useful versions, for arbitrary terms substituting the variables.\<close>
prv_neg_zer_suc_var:
"prv (neg (eql zer (suc (Var xx))))"
and
prv_inj_suc_var:
"prv (imp (eql (suc (Var xx)) (suc (Var yy)))
(eql (Var xx) (Var yy)))"
and
prv_zer_dsj_suc_var:
"prv (dsj (eql (Var yy) zer)
(exi xx (eql (Var yy) (suc (Var xx)))))"
and
prv_pls_zer_var:
"prv (eql (pls (Var xx) zer) (Var xx))"
and
prv_pls_suc_var:
"prv (eql (pls (Var xx) (suc (Var yy)))
(suc (pls (Var xx) (Var yy))))"
and
prv_tms_zer_var:
"prv (eql (tms (Var xx) zer) zer)"
and
prv_tms_suc_var:
"prv (eql (tms (Var xx) (suc (Var yy)))
(pls (tms (Var xx) (Var yy)) (Var xx)))"
begin
text \<open>Rules for quantifiers that allow changing, on the fly, the bound variable with
one that is fresh for the proof context:\<close>
lemma nprv_allI_var:
assumes n1[simp]: "nprv F (subst \<phi> (Var y) x)"
and i[simp]: "F \<subseteq> fmla" "finite F" "\<phi> \<in> fmla" "x \<in> var" "y \<in> var"
and u: "y \<notin> (\<Union>\<phi> \<in> F. Fvars \<phi>)" and yx[simp]: "y = x \<or> y \<notin> Fvars \<phi>"
shows "nprv F (all x \<phi>)"
proof-
have [simp]: "\<And>\<phi>. \<phi> \<in> F \<Longrightarrow> y \<notin> Fvars \<phi>" using u by auto
show ?thesis
apply(subst all_rename2[of _ _ y])
subgoal by auto
subgoal by auto
subgoal by auto
subgoal using yx by auto
subgoal by (nrule r: nprv_allI) .
qed
lemma nprv_exiE_var:
assumes n: "nprv F (exi x \<phi>)"
and nn: "nprv (insert (subst \<phi> (Var y) x) F) \<psi>"
and 0: "F \<subseteq> fmla" "finite F" "\<phi> \<in> fmla" "x \<in> var" "y \<in> var" "\<psi> \<in> fmla"
and yx: "y = x \<or> y \<notin> Fvars \<phi>" "y \<notin> \<Union> (Fvars ` F)" "y \<notin> Fvars \<psi>"
shows "nprv F \<psi>"
proof-
have e: "exi x \<phi> = exi y (subst \<phi> (Var y) x)"
using 0 yx n nn by (subst exi_rename2[of _ _ y]) (auto simp: 0)
show ?thesis
using assms unfolding e
by (auto intro: nprv_exiE[of _ y "subst \<phi> (Var y) x"])
qed
(* The substitution closures of the variable-based axioms
(and the rulifications of the ones that are implications, negations or disjunctions). *)
lemma prv_neg_zer_suc:
assumes [simp]: "t \<in> atrm" shows "prv (neg (eql zer (suc t)))"
using prv_psubst[OF _ _ _ prv_neg_zer_suc_var, of "[(t,xx)]"]
by simp
lemma prv_neg_suc_zer:
assumes "t \<in> atrm" shows "prv (neg (eql (suc t) zer))"
by (metis assms atrm.simps atrm_imp_trm eql fls neg_def prv_eql_sym
prv_neg_zer_suc prv_prv_imp_trans zer_atrm)
(* Rulification: *)
lemmas nprv_zer_suc_contrE =
nprv_flsE[OF nprv_addImpLemmaE[OF prv_neg_zer_suc[unfolded neg_def]], OF _ _ nprv_hyp, simped, rotated]
lemmas nprv_zer_suc_contrE0 = nprv_zer_suc_contrE[OF nprv_hyp, simped]
(* A variation of the above, taking advantage of transitivity and symmetry: *)
lemma nprv_zer_suc_2contrE:
"nprv F (eql t zer) \<Longrightarrow> nprv F (eql t (suc t1)) \<Longrightarrow>
finite F \<Longrightarrow> F \<subseteq> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow>
nprv F \<phi>"
using nprv_eql_transI[OF nprv_eql_symI] nprv_zer_suc_contrE
by (meson atrm_imp_trm suc zer_atrm)
lemmas nprv_zer_suc_2contrE0 = nprv_zer_suc_2contrE[OF nprv_hyp _, simped]
lemmas nprv_zer_suc_2contrE1 = nprv_zer_suc_2contrE[OF _ nprv_hyp, simped]
lemmas nprv_zer_suc_2contrE01 = nprv_zer_suc_2contrE[OF nprv_hyp nprv_hyp, simped]
(* *)
lemma prv_inj_suc:
"t \<in> atrm \<Longrightarrow> t' \<in> atrm \<Longrightarrow>
prv (imp (eql (suc t) (suc t'))
(eql t t'))"
using prv_psubst[OF _ _ _ prv_inj_suc_var, of "[(t,xx),(t',yy)]"]
by simp
(* Rulification: *)
lemmas nprv_eql_sucI = nprv_addImpLemmaI[OF prv_inj_suc, simped, rotated 4]
lemmas nprv_eql_sucE = nprv_addImpLemmaE[OF prv_inj_suc, simped, rotated 2]
lemmas nprv_eql_sucE0 = nprv_eql_sucE[OF nprv_hyp _, simped]
lemmas nprv_eql_sucE1 = nprv_eql_sucE[OF _ nprv_hyp, simped]
lemmas nprv_eql_sucE01 = nprv_eql_sucE[OF nprv_hyp nprv_hyp, simped]
(* NB: Provable substitution closures of sentences in the presence of quantifiers do not go
very smoothly -- the main reason being that bound variable renaming is not assumed
to hold up to equality, but it (only follows that it) holds up to provability: *)
lemma prv_zer_dsj_suc:
assumes t[simp]: "t \<in> atrm" and x[simp]: "x \<in> var" "x \<notin> FvarsT t"
shows "prv (dsj (eql t zer)
(exi x (eql t (suc (Var x)))))"
proof-
define x' where x': "x' \<equiv> getFr [x,yy] [t] []"
have x'_facts[simp]: "x' \<in> var" "x' \<noteq> x" "x' \<noteq> yy" "x' \<notin> FvarsT t" unfolding x'
using getFr_FvarsT_Fvars[of "[x,yy]" "[t]" "[]"] by auto
have "prv (imp (exi xx (eql (Var yy) (suc (Var xx)))) (exi x' (eql (Var yy) (suc (Var x')))))"
by (auto intro!: prv_exi_imp prv_all_gen
simp: prv_exi_inst[of x' "eql (Var yy) (suc (Var x'))" "Var xx", simplified])
with prv_zer_dsj_suc_var
have 0: "prv (dsj (eql (Var yy) zer) (exi x' (eql (Var yy) (suc (Var x')))))"
by (elim prv_dsj_cases[rotated 3])
(auto intro: prv_dsj_impL prv_dsj_impR elim!: prv_prv_imp_trans[rotated 3])
note 1 = prv_psubst[OF _ _ _ 0, of "[(t,yy)]", simplified]
moreover have "prv (imp (exi x' (eql t (suc (Var x')))) (exi x (eql t (suc (Var x)))))"
by (auto intro!: prv_exi_imp prv_all_gen simp: prv_exi_inst[of x "eql t (suc (Var x))" "Var x'", simplified])
ultimately show ?thesis
by (elim prv_dsj_cases[rotated 3])
(auto intro: prv_dsj_impL prv_dsj_impR elim!: prv_prv_imp_trans[rotated 3])
qed
(* The rulification of the above disjunction amounts to reasoning by zero-suc cases: *)
lemma nprv_zer_suc_casesE:
"nprv (insert (eql t zer) F) \<phi> \<Longrightarrow> nprv (insert (eql t (suc (Var x))) F) \<phi> \<Longrightarrow>
finite F \<Longrightarrow> F \<subseteq> fmla \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> t \<in> atrm \<Longrightarrow>
x \<notin> Fvars \<phi> \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow> x \<notin> \<Union> (Fvars ` F) \<Longrightarrow>
nprv F \<phi>"
by (nprover3 r1: nprv_addDsjLemmaE[OF prv_zer_dsj_suc]
r2: nprv_exiE0[of x "eql t (suc (Var x))"]
r3: nprv_mono[of "insert (eql _ (suc _)) _"])
lemmas nprv_zer_suc_casesE0 = nprv_zer_suc_casesE[OF nprv_hyp _, simped]
lemmas nprv_zer_suc_casesE1 = nprv_zer_suc_casesE[OF _ nprv_hyp, simped]
lemmas nprv_zer_suc_casesE01 = nprv_zer_suc_casesE[OF nprv_hyp nprv_hyp, simped]
(* *)
lemma prv_pls_zer:
assumes [simp]: "t \<in> atrm" shows "prv (eql (pls t zer) t)"
using prv_psubst[OF _ _ _ prv_pls_zer_var, of "[(t,xx)]"]
by simp
lemma prv_pls_suc:
"t \<in> atrm \<Longrightarrow> t' \<in> atrm \<Longrightarrow>
prv (eql (pls t (suc t'))
(suc (pls t t')))"
using prv_psubst[OF _ _ _ prv_pls_suc_var, of "[(t,xx),(t',yy)]"]
by simp
lemma prv_tms_zer:
assumes [simp]: "t \<in> atrm" shows "prv (eql (tms t zer) zer)"
using prv_psubst[OF _ _ _ prv_tms_zer_var, of "[(t,xx)]"]
by simp
lemma prv_tms_suc:
"t \<in> atrm \<Longrightarrow> t' \<in> atrm \<Longrightarrow>
prv (eql (tms t (suc t'))
(pls (tms t t') t))"
using prv_psubst[OF _ _ _ prv_tms_suc_var, of "[(t,xx),(t',yy)]"]
by simp
(* Congruence rules for the operators (follow from substitutivity of equality): *)
lemma prv_suc_imp_cong:
assumes t1[simp]: "t1 \<in> atrm" and t2[simp]: "t2 \<in> atrm"
shows "prv (imp (eql t1 t2)
(eql (suc t1) (suc t2)))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t1,t2] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t1,t2]" "[]"] unfolding z_def[symmetric] by auto
show ?thesis
by (nprover4 r1: nprv_prvI r2: nprv_impI
r3: nprv_eql_substE02[of t1 t2 _ "eql (suc (Var z)) (suc t2)" z]
r4: nprv_eq_eqlI)
qed
(* Rulification: *)
lemmas nprv_suc_congI = nprv_addImpLemmaI[OF prv_suc_imp_cong, simped, rotated 4]
lemmas nprv_suc_congE = nprv_addImpLemmaE[OF prv_suc_imp_cong, simped, rotated 2]
lemmas nprv_suc_congE0 = nprv_suc_congE[OF nprv_hyp _, simped]
lemmas nprv_suc_congE1 = nprv_suc_congE[OF _ nprv_hyp, simped]
lemmas nprv_suc_congE01 = nprv_suc_congE[OF nprv_hyp nprv_hyp, simped]
lemma prv_suc_cong:
assumes t1[simp]: "t1 \<in> atrm" and t2[simp]: "t2 \<in> atrm"
assumes "prv (eql t1 t2)"
shows "prv (eql (suc t1) (suc t2))"
by (meson assms atrm_suc atrm_imp_trm eql prv_imp_mp prv_suc_imp_cong t1 t2)
lemma prv_pls_imp_cong:
assumes t1[simp]: "t1 \<in> atrm" and t1'[simp]: "t1' \<in> atrm"
and t2[simp]: "t2 \<in> atrm" and t2'[simp]: "t2' \<in> atrm"
shows "prv (imp (eql t1 t1')
(imp (eql t2 t2') (eql (pls t1 t2) (pls t1' t2'))))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t1,t1',t2,t2'] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z"
"z \<notin> FvarsT t1" "z \<notin> FvarsT t1'" "z \<notin> FvarsT t2" "z \<notin> FvarsT t2'"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t1,t1',t2,t2']" "[]"] unfolding z_def[symmetric] by auto
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_substE02[of t1 t1' _ "eql (pls (Var z) t2) (pls t1' t2')" z])
apply(nrule r: nprv_eql_substE02[of t2 t2' _ "eql (pls t1' (Var z)) (pls t1' t2')" z])
apply(nrule r: nprv_eq_eqlI) .
qed
(* Rulification: *)
lemmas nprv_pls_congI = nprv_addImp2LemmaI[OF prv_pls_imp_cong, simped, rotated 6]
lemmas nprv_pls_congE = nprv_addImp2LemmaE[OF prv_pls_imp_cong, simped, rotated 4]
lemmas nprv_pls_congE0 = nprv_pls_congE[OF nprv_hyp _ _, simped]
lemmas nprv_pls_congE1 = nprv_pls_congE[OF _ nprv_hyp _, simped]
lemmas nprv_pls_congE2 = nprv_pls_congE[OF _ _ nprv_hyp, simped]
lemmas nprv_pls_congE01 = nprv_pls_congE[OF nprv_hyp nprv_hyp _, simped]
lemmas nprv_pls_congE02 = nprv_pls_congE[OF nprv_hyp _ nprv_hyp, simped]
lemmas nprv_pls_congE12 = nprv_pls_congE[OF _ nprv_hyp nprv_hyp, simped]
lemmas nprv_pls_congE012 = nprv_pls_congE[OF nprv_hyp nprv_hyp nprv_hyp, simped]
lemma prv_pls_cong:
assumes "t1 \<in> atrm" "t1' \<in> atrm" "t2 \<in> atrm" "t2' \<in> atrm"
and "prv (eql t1 t1')" and "prv (eql t2 t2')"
shows "prv (eql (pls t1 t2) (pls t1' t2'))"
by (metis assms atrm_imp_trm cnj eql pls prv_cnjI prv_cnj_imp_monoR2 prv_imp_mp prv_pls_imp_cong)
lemma prv_pls_congL:
"t1 \<in> atrm \<Longrightarrow> t1' \<in> atrm \<Longrightarrow> t2 \<in> atrm \<Longrightarrow>
prv (eql t1 t1') \<Longrightarrow> prv (eql (pls t1 t2) (pls t1' t2))"
by (rule prv_pls_cong[OF _ _ _ _ _ prv_eql_reflT]) auto
lemma prv_pls_congR:
"t1 \<in> atrm \<Longrightarrow> t2 \<in> atrm \<Longrightarrow> t2' \<in> atrm \<Longrightarrow>
prv (eql t2 t2') \<Longrightarrow> prv (eql (pls t1 t2) (pls t1 t2'))"
by (rule prv_pls_cong[OF _ _ _ _ prv_eql_reflT]) auto
lemma nprv_pls_cong:
assumes [simp]: "t1 \<in> atrm" "t1' \<in> atrm" "t2 \<in> atrm" "t2' \<in> atrm"
shows "nprv {eql t1 t1', eql t2 t2'} (eql (pls t1 t2) (pls t1' t2'))"
unfolding nprv_def
by (auto intro!: prv_prv_imp_trans[OF _ _ _ prv_scnj2_imp_cnj] prv_cnj_imp_monoR2 prv_pls_imp_cong)
lemma prv_tms_imp_cong:
assumes t1[simp]: "t1 \<in> atrm" and t1'[simp]: "t1' \<in> atrm"
and t2[simp]: "t2 \<in> atrm" and t2'[simp]: "t2' \<in> atrm"
shows "prv (imp (eql t1 t1')
(imp (eql t2 t2') (eql (tms t1 t2) (tms t1' t2'))))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t1,t1',t2,t2'] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z"
"z \<notin> FvarsT t1" "z \<notin> FvarsT t1'" "z \<notin> FvarsT t2" "z \<notin> FvarsT t2'"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t1,t1',t2,t2']" "[]"] unfolding z_def[symmetric] by auto
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_substE02[of t1 t1' _ "eql (tms (Var z) t2) (tms t1' t2')" z])
apply(nrule r: nprv_eql_substE02[of t2 t2' _ "eql (tms t1' (Var z)) (tms t1' t2')" z])
apply(nrule r: nprv_eq_eqlI) .
qed
(* Rulification: *)
lemmas nprv_tms_congI = nprv_addImp2LemmaI[OF prv_tms_imp_cong, simped, rotated 6]
lemmas nprv_tms_congE = nprv_addImp2LemmaE[OF prv_tms_imp_cong, simped, rotated 4]
lemmas nprv_tms_congE0 = nprv_tms_congE[OF nprv_hyp _ _, simped]
lemmas nprv_tms_congE1 = nprv_tms_congE[OF _ nprv_hyp _, simped]
lemmas nprv_tms_congE2 = nprv_tms_congE[OF _ _ nprv_hyp, simped]
lemmas nprv_tms_congE01 = nprv_tms_congE[OF nprv_hyp nprv_hyp _, simped]
lemmas nprv_tms_congE02 = nprv_tms_congE[OF nprv_hyp _ nprv_hyp, simped]
lemmas nprv_tms_congE12 = nprv_tms_congE[OF _ nprv_hyp nprv_hyp, simped]
lemmas nprv_tms_congE012 = nprv_tms_congE[OF nprv_hyp nprv_hyp nprv_hyp, simped]
lemma prv_tms_cong:
assumes "t1 \<in> atrm" "t1' \<in> atrm" "t2 \<in> atrm" "t2' \<in> atrm"
and "prv (eql t1 t1')" and "prv (eql t2 t2')"
shows "prv (eql (tms t1 t2) (tms t1' t2'))"
by (metis assms atrm_imp_trm cnj eql tms prv_cnjI prv_cnj_imp_monoR2 prv_imp_mp prv_tms_imp_cong)
lemma nprv_tms_cong:
assumes [simp]: "t1 \<in> atrm" "t1' \<in> atrm" "t2 \<in> atrm" "t2' \<in> atrm"
shows "nprv {eql t1 t1', eql t2 t2'} (eql (tms t1 t2) (tms t1' t2'))"
unfolding nprv_def
by (auto intro!: prv_prv_imp_trans[OF _ _ _ prv_scnj2_imp_cnj] prv_cnj_imp_monoR2 prv_tms_imp_cong)
lemma prv_tms_congL:
"t1 \<in> atrm \<Longrightarrow> t1' \<in> atrm \<Longrightarrow> t2 \<in> atrm \<Longrightarrow>
prv (eql t1 t1') \<Longrightarrow> prv (eql (tms t1 t2) (tms t1' t2))"
by (rule prv_tms_cong[OF _ _ _ _ _ prv_eql_reflT]) auto
lemma prv_tms_congR:
"t1 \<in> atrm \<Longrightarrow> t2 \<in> atrm \<Longrightarrow> t2' \<in> atrm \<Longrightarrow>
prv (eql t2 t2') \<Longrightarrow> prv (eql (tms t1 t2) (tms t1 t2'))"
by (rule prv_tms_cong[OF _ _ _ _ prv_eql_reflT]) auto
section \<open>Properties Provable in Q\<close>
subsection \<open>General properties, unconstrained by numerals\<close>
lemma prv_pls_suc_zer:
"t \<in> atrm \<Longrightarrow> prv (eql (pls t (suc zer)) (suc t))"
by (metis (no_types, hide_lams) atrm.atrm_pls atrm_imp_trm
pls prv_eql_trans prv_pls_suc prv_pls_zer prv_suc_cong suc zer_atrm)
lemma prv_LLq_suc_imp:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm"
shows "prv (imp (LLq (suc t1) (suc t2)) (LLq t1 t2))"
proof- define z where "z \<equiv> getFr [xx,yy,zz] [t1,t2] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t1,t2]" "[]"] unfolding z_def[symmetric] by auto
note LLq_pls[of _ _ z,simp]
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_exiE0)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of "Var z" t1]])
apply(nrule r: nprv_clear3_3)
apply(nrule r: nprv_eql_transE01[of "suc t2" "pls (Var z) (suc t1)" _ "suc (pls (Var z) t1)"])
apply(nrule r: nprv_eql_sucE0[of t2 "pls (Var z) t1"])
apply(nrule r: nprv_exiI[of _ _ "Var z" z]) .
qed
(* Rulification: *)
lemmas nprv_LLq_sucI = nprv_addImpLemmaI[OF prv_LLq_suc_imp, simped, rotated 4]
lemmas nprv_LLq_sucE = nprv_addImpLemmaE[OF prv_LLq_suc_imp, simped, rotated 2]
lemmas nprv_LLq_sucE0 = nprv_LLq_sucE[OF nprv_hyp _, simped]
lemmas nprv_LLq_sucE1 = nprv_LLq_sucE[OF _ nprv_hyp, simped]
lemmas nprv_LLq_sucE01 = nprv_LLq_sucE[OF nprv_hyp nprv_hyp, simped]
lemma prv_LLs_imp_LLq:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm"
shows "prv (imp (LLs t1 t2) (LLq t1 t2))"
by (simp add: LLs_LLq prv_imp_cnjL)
lemma prv_LLq_refl:
"prv (LLq zer zer)"
by (auto simp: LLq_pls_zz prv_pls_zer prv_prv_eql_sym intro!: prv_exiI[of zz _ zer])
text \<open>NB: Monotonicity of pls and tms w.r.t. LLq cannot be proved in Q.\<close>
lemma prv_suc_mono_LLq:
assumes "t1 \<in> atrm" "t2 \<in> atrm"
shows "prv (imp (LLq t1 t2) (LLq (suc t1) (suc t2)))"
proof-
have assms1: "t1 \<in> trm" "t2 \<in> trm" using assms by auto
define z where "z \<equiv> getFr [xx,yy,zz] [t1,t2] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t1,t2]" "[]"] using assms1 unfolding z_def[symmetric] by auto
define x where "x \<equiv> getFr [xx,yy,zz,z] [t1,t2] []"
have x_facts[simp]: "x \<in> var" "x \<noteq> xx" "x \<noteq> yy" "x \<noteq> zz" "zz \<noteq> x" "x \<noteq> z" "z \<noteq> x" "x \<notin> FvarsT t1""x \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy,zz,z]" "[t1,t2]" "[]"] using assms1 unfolding x_def[symmetric] by auto
note assms[simp]
note LLq_pls[of _ _ z, simp]
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_exiE0[of z "eql t2 (pls (Var z) t1)"])
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_exiI[of _ _ "Var z"])
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of "Var z" t1]])
apply(nrule r: nprv_eql_substE[of _
"pls (Var z) (suc t1)" "suc (pls (Var z) t1)"
"eql (suc t2) (Var x)" x])
apply(nrule r: nprv_clear2_1)
apply(nrule r: nprv_suc_congI) .
qed
(* Rulification: *)
lemmas nprv_suc_mono_LLqI = nprv_addImpLemmaI[OF prv_suc_mono_LLq, simped, rotated 4]
lemmas nprv_suc_mono_LLqE = nprv_addImpLemmaE[OF prv_suc_mono_LLq, simped, rotated 2]
lemmas nprv_suc_mono_LLqE0 = nprv_suc_mono_LLqE[OF nprv_hyp _, simped]
lemmas nprv_suc_mono_LLqE1 = nprv_suc_mono_LLqE[OF _ nprv_hyp, simped]
lemmas nprv_suc_mono_LLqE01 = nprv_suc_mono_LLqE[OF nprv_hyp nprv_hyp, simped]
subsection \<open>Representability properties\<close>
text \<open>Representability of number inequality\<close>
lemma prv_neg_eql_suc_Num_zer:
"prv (neg (eql (suc (Num n)) zer))"
apply(induct n)
apply (metis Num Num.simps(1) Num_atrm eql fls in_num neg_def prv_eql_sym prv_neg_zer_suc prv_prv_imp_trans suc)
by (metis Num_atrm atrm_imp_trm eql fls neg_def prv_eql_sym prv_neg_zer_suc prv_prv_imp_trans suc zer_atrm)
lemma diff_prv_eql_Num:
assumes "m \<noteq> n"
shows "prv (neg (eql (Num m) (Num n)))"
using assms proof(induct m arbitrary: n)
case 0
then obtain n' where n: "n = Suc n'" by (cases n) auto
thus ?case unfolding n by (simp add: prv_neg_zer_suc)
next
case (Suc m n) note s = Suc
show ?case
proof(cases n)
case 0
thus ?thesis by (simp add: prv_neg_eql_suc_Num_zer)
next
case (Suc n') note n = Suc
thus ?thesis using s
by simp (meson Num Num_atrm eql in_num neg prv_imp_mp prv_imp_neg_rev prv_inj_suc suc)
qed
qed
lemma consistent_prv_eql_Num_equal:
assumes consistent and "prv (eql (Num m) (Num n))"
shows "m = n"
using assms consistent_def3 diff_prv_eql_Num in_num by blast
text \<open>Representability of addition\<close>
lemma prv_pls_zer_zer:
"prv (eql (pls zer zer) zer)"
by (simp add: prv_pls_zer)
lemma prv_eql_pls_plus:
"prv (eql (pls (Num m) (Num n))
(Num (m+n)))"
proof(induct n)
case (Suc n)
note 0 = prv_pls_suc[of "Num m" "Num n", simplified]
show ?case
by (auto intro: prv_eql_trans[OF _ _ _ 0 prv_suc_cong[OF _ _ Suc]])
qed(simp add: prv_pls_zer)
lemma not_plus_prv_neg_eql_pls:
assumes "m + n \<noteq> k"
shows "prv (neg (eql (pls (Num m) (Num n)) (Num k)))"
using assms proof(induction n arbitrary: k)
case 0 hence m: "m \<noteq> k" by simp
note diff_prv_eql_Num[OF m, simp]
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_pls_zer, of "Num m"])
apply(nrule r: nprv_eql_substE
[of _ "pls (Num m) zer" "Num m" "neg (eql (Var xx) (Num k))" xx])
apply(nrule r: prv_nprvI) .
next
case (Suc n)
have 0: "\<And>k'. k = Suc k' \<Longrightarrow>
prv (neg (eql (pls (Num m) (Num n)) (Num k'))) \<and> m + n \<noteq> k'"
using Suc.IH Suc.prems by auto
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc, of "Num m" "Num n"])
apply(nrule r: nprv_eql_substE[of _ "pls (Num m) (suc (Num n))"
"suc (pls (Num m) (Num n))" "neg (eql (Var xx) (Num k))" xx])
apply(nrule r: nprv_clear)
apply(cases k)
subgoal by (nprover2 r1: prv_nprvI r2: prv_neg_suc_zer)
subgoal for k' apply(frule 0)
by (nprover4 r1: nprv_addLemmaE r2: nprv_negI
r3: nprv_negE0 r4: nprv_eql_sucI) .
qed
lemma consistent_prv_eql_pls_plus_rev:
assumes "consistent" "prv (eql (pls (Num m) (Num n)) (Num k))"
shows "k = m + n"
by (metis Num assms consistent_def eql not_plus_prv_neg_eql_pls num pls prv_neg_fls subsetCE)
text \<open>Representability of multiplication\<close>
lemma prv_tms_Num_zer:
"prv (eql (tms (Num n) zer) zer)"
by(auto simp: prv_tms_zer)
lemma prv_eql_tms_times:
"prv (eql (tms (Num m) (Num n)) (Num (m * n)))"
proof(induct n)
case (Suc n)
note 0 = prv_pls_congL[OF _ _ _ Suc, of "Num m", simplified]
thm prv_pls_cong[no_vars]
note add.commute[simp]
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF 0])
apply(nrule r: nprv_addLemmaE[OF prv_tms_suc[of "Num m" "Num n", simplified]])
apply(nrule r: nprv_eql_transE01[of
"tms (Num m) (suc (Num n))"
"pls (tms (Num m) (Num n)) (Num m)" _
"pls (Num (m * n)) (Num m)"])
apply(nrule r: nprv_clear3_2)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_addLemmaE[OF prv_eql_pls_plus[of "m * n" m]])
apply(nrule r: nprv_eql_transE01[of
"tms (Num m) (suc (Num n))"
"pls (Num (m * n)) (Num m)" _
"Num (m * n + m)"]) .
qed(auto simp: prv_tms_zer)
lemma ge_prv_neg_eql_pls_Num_zer:
assumes [simp]: "t \<in> atrm" and m: "m > k"
shows "prv (neg (eql (pls t (Num m)) (Num k)))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] using assms unfolding z_def[symmetric] by auto
show ?thesis using m proof(induction k arbitrary: m)
case (0 m)
show ?case
apply(cases m)
subgoal using 0 by auto
subgoal for n
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_neg_suc_zer[of "pls t (Num n)"]])
apply(nrule r: nprv_negI)
apply(nrule r: nprv_negE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_eql_symE0)
apply(nrule r: nprv_eql_substE[of _ zer "pls t (suc (Num n))" "eql (suc (pls t (Num n))) (Var z)" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_eql_symI)
apply(nrule r: prv_nprvI)
apply(nrule r: prv_pls_suc) . .
next
case (Suc k mm)
then obtain m where mm[simp]: "mm = Suc m" and k: "k < m" by (cases mm) auto
show ?case unfolding mm Num.simps
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF Suc.IH[OF k]])
apply(nrule r: nprv_negI)
apply(nrule r: nprv_negE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_impI_rev)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of t "Num m"]])
apply(nrule r: nprv_eql_substE[of _ "pls t (suc (Num m))" "suc (pls t (Num m))"
"imp (eql (Var z) (suc (Num k))) (eql (pls t (Num m)) (Num k))" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_sucI) .
qed
qed
lemma nprv_pls_Num_injectR:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm"
shows "prv (imp (eql (pls t1 (Num m)) (pls t2 (Num m)))
(eql t1 t2))"
proof-
define z where "z \<equiv> getFr [xx,yy] [t1,t2] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy]" "[t1,t2]" "[]"] unfolding z_def[symmetric] by auto
show ?thesis proof(induction m)
case 0
show ?case unfolding Num.simps
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_pls_zer[of t1]])
apply(nrule r: nprv_eql_substE[of _ "pls t1 zer" "t1" "imp (eql (Var z) (pls t2 zer)) (eql t1 t2)" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_addLemmaE[OF prv_pls_zer[of t2]])
apply(nrule r: nprv_eql_substE[of _ "pls t2 zer" "t2" "imp (eql t1 (Var z)) (eql t1 t2)" z])
apply(nrule r: nprv_impI) .
next
case (Suc m)
note Suc.IH[simp]
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of t1 "Num m"]])
apply(nrule r: nprv_eql_substE[of _ "pls t1 (suc (Num m))" "suc (pls t1 (Num m))"
"imp (eql (Var z) (pls t2 (suc (Num m)))) (eql t1 t2)" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of t2 "Num m"]])
apply(nrule r: nprv_eql_substE[of _ "pls t2 (suc (Num m))" "suc (pls t2 (Num m))"
"imp (eql (suc (pls t1 (Num m))) (Var z)) (eql t1 t2)" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_sucE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: prv_nprv1I) .
qed
qed
(* Rulification: *)
lemmas nprv_pls_Num_injectI = nprv_addImpLemmaI[OF nprv_pls_Num_injectR, simped, rotated 4]
lemmas nprv_pls_Num_injectE = nprv_addImpLemmaE[OF nprv_pls_Num_injectR, simped, rotated 2]
lemmas nprv_pls_Num_injectE0 = nprv_pls_Num_injectE[OF nprv_hyp _, simped]
lemmas nprv_pls_Num_injectE1 = nprv_pls_Num_injectE[OF _ nprv_hyp, simped]
lemmas nprv_pls_Num_injectE01 = nprv_pls_Num_injectE[OF nprv_hyp nprv_hyp, simped]
lemma not_times_prv_neg_eql_tms:
assumes "m * n \<noteq> k"
shows "prv (neg (eql (tms (Num m) (Num n)) (Num k)))"
using assms proof(induction n arbitrary: k)
case 0 hence m: "0 \<noteq> k" by simp have zer: "zer = Num 0" by simp
have [simp]: "prv (neg (eql zer (Num k)))" by (subst zer, rule diff_prv_eql_Num[OF m])
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_tms_zer, of "Num m"])
apply(nrule r: nprv_eql_substE[of _ "tms (Num m) zer" zer "neg (eql (Var xx) (Num k))" xx])
apply(nrule r: prv_nprvI) .
next
case (Suc n)
have [simp]: "nprv {} (neg (eql (pls (tms (Num m) (Num n)) (Num m)) (Num k)))"
proof(cases "k < m")
case [simp]: True
thus ?thesis apply- by (nprover2 r1: prv_nprvI r2: ge_prv_neg_eql_pls_Num_zer)
next
case False
define k' where "k' \<equiv> k - m"
with False have k: "k = k' + m" by auto
hence mm: "m * n \<noteq> k'" using False Suc.prems by auto
note IH = Suc.IH[OF mm]
show ?thesis unfolding k
apply(nrule r: nprv_negI)
apply(nrule r: nprv_addLemmaE[OF prv_prv_eql_sym[OF _ _ prv_eql_pls_plus[of k' m]]])
apply(nrule r: nprv_eql_transE01[of _ "Num (k' + m)"])
apply(nrule r: nprv_clear3_2)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_pls_Num_injectE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_addLemmaE[OF IH])
apply(nrule r: nprv_negE0) .
qed
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_tms_suc, of "Num m" "Num n"])
apply(nrule r: nprv_eql_substE[of _ "tms (Num m) (suc (Num n))" "pls (tms (Num m) (Num n)) (Num m)"
"neg (eql (Var xx) (Num k))" xx])
apply(nrule r: nprv_clear) .
qed
lemma consistent_prv_eql_tms_times_rev:
assumes "consistent" "prv (eql (tms (Num m) (Num n)) (Num k))"
shows "k = m * n"
by (metis Num assms consistent_def eql not_times_prv_neg_eql_tms num tms prv_neg_fls subsetCE)
text \<open>Representability of the order\<close>
lemma leq_prv_LLq_Num:
assumes "m \<le> n"
shows "prv (LLq (Num m) (Num n))"
proof-
obtain i where n: "n = i + m" using assms add.commute le_Suc_ex by blast
note prv_eql_pls_plus[simp]
have "prv (exi zz (eql (Num (i + m)) (pls (Var zz) (Num m))))"
by(nprover2 r1: prv_exiI[of _ _ "Num i"] r2: prv_prv_eql_sym)
thus ?thesis unfolding n by (simp add: LLq_pls_zz)
qed
subsection \<open>The "order-adequacy" properties\<close>
text \<open>These are properties Q1--O9 from
Peter Smith, An Introduction to Gödel's theorems, Second Edition, Page 73.\<close>
lemma prv_LLq_zer: \<comment> \<open>O1\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (LLq zer t)"
proof-
define z where "z \<equiv> getFr [xx,yy] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy]" "[t]" "[]"] unfolding z_def[symmetric] by auto
have "prv (exi z (eql t (pls (Var z) zer)))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_exiI[of _ _ t])
apply(nrule r: nprv_eql_symI)
apply(nrule r: prv_nprvI)
apply(nrule r: prv_pls_zer) .
thus ?thesis by (simp add: LLq_pls[of _ _ z])
qed
lemmas Q1 = prv_LLq_zer
lemma prv_LLq_zer_imp_eql:
assumes [simp]: "t \<in> atrm"
shows "prv (imp (LLq t zer) (eql t zer))"
proof-
define y where "y \<equiv> getFr [] [t] []"
have y_facts[simp]: "y \<in> var" "y \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[]" "[t]" "[]"] unfolding y_def[symmetric] by auto
define z where "z \<equiv> getFr [y] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> y" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[y]" "[t]" "[]"] unfolding z_def[symmetric] by auto
define x where "x \<equiv> getFr [y,z] [t] []"
have x_facts[simp]: "x \<in> var" "x \<noteq> y" "x \<noteq> z" "x \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[y,z]" "[t]" "[]"] unfolding x_def by auto
note LLq_pls[of _ _ z,simp]
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_zer_suc_casesE[of t _ _ y])
apply(nrule r: nprv_exiE0[of z "eql zer (pls (Var z) t)"])
apply(nrule r: nprv_clear3_3)
apply(nrule r: nprv_eql_symE0[of t])
apply(nrule r: nprv_eql_substE01[of "suc (Var y)" t _ "eql zer (pls (Var z) (Var x))" x])
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of "Var z" "Var y",simplified]])
apply(nrule r: nprv_eql_transE01[of zer "pls (Var z) (suc (Var y))" _ "suc (pls (Var z) (Var y))"])
apply(nrule r: nprv_zer_suc_contrE0[of "pls (Var z) (Var y)"]) .
qed
(* Rulification: *)
lemmas nprv_LLq_zer_eqlI = nprv_addImpLemmaI[OF prv_LLq_zer_imp_eql, simped, rotated 3]
lemmas nprv_LLq_zer_eqlE = nprv_addImpLemmaE[OF prv_LLq_zer_imp_eql, simped, rotated 1]
lemmas nprv_LLq_zer_eqlE0 = nprv_LLq_zer_eqlE[OF nprv_hyp _, simped]
lemmas nprv_LLq_zer_eqlE1 = nprv_LLq_zer_eqlE[OF _ nprv_hyp, simped]
lemmas nprv_LLq_zer_eqlE01 = nprv_LLq_zer_eqlE[OF nprv_hyp nprv_hyp, simped]
lemma prv_sdsj_eql_imp_LLq: \<comment> \<open>O2\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (imp (ldsj (map (\<lambda>i. eql t (Num i)) (toN n))) (LLq t (Num n)))"
proof-
define z where "z \<equiv> getFr [xx,yy] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy]" "[t]" "[]"] unfolding z_def[symmetric] by auto
note imp[rule del, intro!] note dsj[intro!]
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_ldsjE0)
subgoal for i
apply(nrule r: nprv_eql_substE[of _ t "Num i" "LLq (Var z) (Num n)" z])
subgoal by (nrule r: nprv_hyp)
subgoal by (nprover3 r1: nprv_addLemmaE[OF leq_prv_LLq_Num])
subgoal by (nrule r: nprv_hyp) . .
qed
(* Rulification: *)
declare subset_eq[simp]
lemmas nprv_sdsj_eql_LLqI = nprv_addImpLemmaI[OF prv_sdsj_eql_imp_LLq, simped, rotated 3]
lemmas nprv_sdsj_eql_LLqE = nprv_addImpLemmaE[OF prv_sdsj_eql_imp_LLq, simped, rotated 1]
declare subset_eq[simp del]
lemmas nprv_sdsj_eql_LLqE0 = nprv_sdsj_eql_LLqE[OF nprv_hyp _, simped]
lemmas nprv_sdsj_eql_LLqE1 = nprv_sdsj_eql_LLqE[OF _ nprv_hyp, simped]
lemmas nprv_sdsj_eql_LLqE01 = nprv_sdsj_eql_LLqE[OF nprv_hyp nprv_hyp, simped]
lemmas O2I = nprv_sdsj_eql_LLqI
lemmas O2E = nprv_sdsj_eql_LLqE
lemmas O2E0 = nprv_sdsj_eql_LLqE0
lemmas O2E1 = nprv_sdsj_eql_LLqE1
lemmas O2E01 = nprv_sdsj_eql_LLqE01
(* *)
lemma prv_LLq_imp_sdsj_eql: \<comment> \<open>O3\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (imp (LLq t (Num n)) (ldsj (map (\<lambda>i. eql t (Num i)) (toN n))))"
using assms proof(induction n arbitrary: t)
case (0 t) note 0[simp]
note prv_LLq_zer_imp_eql[OF 0,simp]
show ?case
by (nprover4 r1: nprv_prvI r2: nprv_impI r3: nprv_ldsjI r4: prv_nprv1I)
next
- case (Suc n) note t[simp] = `t \<in> atrm`
+ case (Suc n) note t[simp] = \<open>t \<in> atrm\<close>
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] unfolding z_def[symmetric] by auto
note subset_eq[simp]
have [simp]: "eql t zer \<in> (\<lambda>x. eql t (Num x)) ` {0..Suc n}" by (force simp: image_def)
have [simp]: "\<And>i. i \<le> n \<Longrightarrow>
eql (suc (Var z)) (suc (Num i)) \<in> (\<lambda>x. eql (suc (Var z)) (Num x)) ` {0..Suc n}"
by (auto simp: image_def intro!: bexI[of _ "Suc _"])
show ?case
apply(nrule r: nprv_prvI)
apply(nrule2 r: nprv_zer_suc_casesE[of t _ _ z])
subgoal by (nprover3 r1: nprv_impI r2: nprv_clear2_1 r3: nprv_ldsjI)
subgoal
apply(nrule r: nprv_eql_substE[of _ t "suc (Var z)"
"imp (LLq (Var xx) (suc (Num n))) (ldsj (map (\<lambda>i. eql (Var xx) (Num i)) (toN (Suc n))))" xx])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_LLq_sucE0)
apply(nrule r: nprv_addImpLemmaE[OF Suc.IH[of "Var z", simplified]])
apply(nrule r: nprv_ldsjE0)
subgoal for i apply(nrule r: nprv_ldsjI[of _ "eql (suc (Var z)) (suc (Num i))"])
apply(nrule r: nprv_suc_congI) . . .
qed
(* Rulification: *)
declare subset_eq[simp]
lemmas prv_LLq_sdsj_eqlI = nprv_addImpLemmaI[OF prv_LLq_imp_sdsj_eql, simped, rotated 3]
lemmas prv_LLq_sdsj_eqlE = nprv_addImpLemmaE[OF prv_LLq_imp_sdsj_eql, simped, rotated 1]
declare subset_eq[simp del]
lemmas prv_LLq_sdsj_eqlE0 = prv_LLq_sdsj_eqlE[OF nprv_hyp _, simped]
lemmas prv_LLq_sdsj_eqlE1 = prv_LLq_sdsj_eqlE[OF _ nprv_hyp, simped]
lemmas prv_LLq_sdsj_eqlE01 = prv_LLq_sdsj_eqlE[OF nprv_hyp nprv_hyp, simped]
lemmas O3I = prv_LLq_sdsj_eqlI
lemmas O3E = prv_LLq_sdsj_eqlE
lemmas O3E0 = prv_LLq_sdsj_eqlE0
lemmas O3E1 = prv_LLq_sdsj_eqlE1
lemmas O3E01 = prv_LLq_sdsj_eqlE01
(* *)
lemma not_leq_prv_neg_LLq_Num:
assumes "\<not> m \<le> n" (* This is just m < n, of course. *)
shows "prv (neg (LLq (Num m) (Num n)))"
proof-
have [simp]: "\<And>i. i \<le> n \<Longrightarrow> prv (imp (eql (Num m) (Num i)) fls)"
unfolding neg_def[symmetric]
using assms by (intro diff_prv_eql_Num) simp
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_negI)
apply(nrule r: O3E0)
apply(nrule r: nprv_ldsjE0)
apply(nrule r: nprv_clear3_2)
apply(nrule r: nprv_clear2_2)
apply(nrule r: prv_nprv1I) .
qed
lemma consistent_prv_LLq_Num_leq:
assumes consistent "prv (LLq (Num m) (Num n))"
shows "m \<le> n"
by (metis Num assms consistent_def LLq not_leq_prv_neg_LLq_Num num prv_neg_fls subsetCE)
(* *)
lemma prv_ball_NumI: \<comment> \<open>O4\<close>
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla"
and [simp]: "\<And> i. i \<le> n \<Longrightarrow> prv (subst \<phi> (Num i) x)"
shows "prv (ball x (Num n) \<phi>)"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_ballI)
apply(nrule r: O3E0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_ldsjE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_eql_substE[of _ "Var x" "Num _" \<phi> x])
apply(nrule r: prv_nprvI) .
lemmas O4 = prv_ball_NumI
lemma prv_bexi_NumI: \<comment> \<open>O5\<close>
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla"
and [simp]: "i \<le> n" "prv (subst \<phi> (Num i) x)"
shows "prv (bexi x (Num n) \<phi>)"
proof-
note leq_prv_LLq_Num[simp]
show ?thesis
by (nprover4 r1: nprv_prvI r2: nprv_bexiI[of _ _ "Num i"] r3: prv_nprvI r4: prv_nprvI)
qed
lemmas O5 = prv_bexi_NumI
lemma prv_LLq_Num_imp_Suc: \<comment> \<open>O6\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (imp (LLq t (Num n)) (LLq t (suc (Num n))))"
proof-
have [simp]: "\<And>i. i \<le> n \<Longrightarrow> prv (LLq (Num i) (suc (Num n)))"
apply(subst Num.simps(2)[symmetric])
by (rule leq_prv_LLq_Num) simp
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: O3E0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_ldsjE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_eql_substE[of _ t "Num _" "LLq (Var xx) (suc (Num n))" xx])
apply(nrule r: prv_nprvI) .
qed
(* Rulification: *)
lemmas nprv_LLq_Num_SucI = nprv_addImpLemmaI[OF prv_LLq_Num_imp_Suc, simped, rotated 3]
lemmas nprv_LLq_Num_SucE = nprv_addImpLemmaE[OF prv_LLq_Num_imp_Suc, simped, rotated 1]
lemmas nprv_LLq_Num_SucE0 = nprv_LLq_Num_SucE[OF nprv_hyp _, simped]
lemmas nprv_LLq_Num_SucE1 = nprv_LLq_Num_SucE[OF _ nprv_hyp, simped]
lemmas nprv_LLq_Num_SucE01 = nprv_LLq_Num_SucE[OF nprv_hyp nprv_hyp, simped]
lemmas O6I = nprv_LLq_Num_SucI
lemmas O6E = nprv_LLq_Num_SucE
lemmas O6E0 = nprv_LLq_Num_SucE0
lemmas O6E1 = nprv_LLq_Num_SucE1
lemmas O6E01 = nprv_LLq_Num_SucE01
text \<open>Crucial for proving O7:\<close>
lemma prv_LLq_suc_Num_pls_Num:
assumes [simp]: "t \<in> atrm"
shows "prv (LLq (suc (Num n)) (pls (suc t) (Num n)))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] unfolding z_def[symmetric] by auto
show ?thesis
proof(induction n)
case 0
have "prv (exi z (eql (pls (suc t) zer) (pls (Var z) (suc zer))))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_exiI[of _ _ t])
apply(nrule r: nprv_addLemmaE[OF prv_pls_zer[of "suc t"]])
apply(nrule r: nprv_eql_substE[of _ "pls (suc t) zer" "suc t" "eql (Var z) (pls t (suc zer))" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_eql_symI)
apply(nrule r: prv_nprvI)
apply(nrule r: prv_pls_suc_zer) .
thus ?case by (simp add: LLq_pls[of _ _ z])
next
case (Suc n)
have nn: "suc (Num n) = suc (Num n)" by simp
note Suc.IH[simp]
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF prv_pls_suc[of "suc t" "Num n"]])
apply(nrule r: nprv_eql_substE[of _ "pls (suc t) (suc (Num n))" "suc (pls (suc t) (Num n))"
"LLq (suc (suc (Num n))) (Var z)" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_suc_mono_LLqI)
apply(nrule r: prv_nprvI) .
qed
qed
lemma prv_Num_LLq_imp_eql_suc: \<comment> \<open>O7\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (imp (LLq (Num n) t)
(dsj (eql (Num n) t)
(LLq (suc (Num n)) t)))"
proof-
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] unfolding z_def[symmetric] by auto
define x where "x \<equiv> getFr [xx,yy,zz,z] [t] []"
have x_facts[simp]: "x \<in> var" "x \<noteq> xx" "x \<noteq> yy" "x \<noteq> zz" "zz \<noteq> x" "x \<noteq> z" "z \<noteq> x" "x \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz,z]" "[t]" "[]"] unfolding x_def[symmetric] by auto
define y where "y \<equiv> getFr [x,z] [t] []"
have y_facts[simp]: "y \<in> var" "y \<notin> FvarsT t" "x \<noteq> y" "y \<noteq> x" "z \<noteq> y" "y \<noteq> z"
using getFr_FvarsT_Fvars[of "[x,z]" "[t]" "[]"] unfolding y_def[symmetric] by auto
have [simp]: "prv (eql (pls zer (Num n)) (Num n))"
by (subst Num.simps(1)[symmetric]) (metis plus_nat.add_0 prv_eql_pls_plus)
have [simp]: "prv (LLq (suc (Num n)) (pls (suc (Var x)) (Num n)))"
by (simp add: prv_LLq_suc_Num_pls_Num)
note LLq_pls[of "Num n" t z, simplified, simp]
show ?thesis
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_exiE0)
apply(nrule r: nprv_clear2_2)
apply(nrule r: nprv_zer_suc_casesE[of "Var z" _ _ x])
subgoal
apply(nrule r: nprv_dsjIL)
apply(nrule r: nprv_impI_rev2[of "{eql (Var z) zer}" "eql t (pls (Var z) (Num n))"])
apply(nrule r: nprv_eql_substE
[of _ "Var z" zer "imp (eql t (pls (Var y) (Num n))) (eql (Num n) t)" y])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_substE[of _ t "pls zer (Num n)" "eql (Num n) (Var y)" y])
apply(nrule r: nprv_clear)
apply(nrule r: prv_nprvI)
apply(nrule r: prv_prv_eql_sym) .
subgoal
apply(nrule r: nprv_dsjIR)
apply(nrule r: nprv_impI_rev2[of "{eql (Var z) (suc (Var x))}" "eql t (pls (Var z) (Num n))"])
apply(nrule r: nprv_eql_substE
[of _ "Var z" "suc (Var x)" "imp (eql t (pls (Var y) (Num n))) (LLq (suc (Num n)) t)" y])
(* *)
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_eql_substE
[of _ t "pls (suc (Var x)) (Num n)" "LLq (suc (Num n)) (Var y)" y])
apply(nrule r: prv_nprvI) . .
qed
(* Rulification (this one is slightly more complex, as it puts together impE with dsjE): *)
lemma prv_Num_LLq_eql_sucE:
"nprv F (LLq (Num n) t) \<Longrightarrow>
nprv (insert (eql (Num n) t) F) \<psi> \<Longrightarrow>
nprv (insert (LLq (suc (Num n)) t) F) \<psi> \<Longrightarrow>
t \<in> atrm \<Longrightarrow> F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
nprv F \<psi>"
apply(nrule r: nprv_addImpLemmaE[OF prv_Num_LLq_imp_eql_suc])
apply(nrule2 r: nprv_dsjE0[of "eql (Num n) t" "LLq (suc (Num n)) t"])
subgoal by (nrule r: nprv_mono[of "insert (eql (Num n) t) F"])
subgoal by (nrule r: nprv_mono[of "insert (LLq (suc (Num n)) t) F"]) .
lemmas prv_Num_LLq_eql_sucE0 = prv_Num_LLq_eql_sucE[OF nprv_hyp _ _, simped]
lemmas prv_Num_LLq_eql_sucE1 = prv_Num_LLq_eql_sucE[OF _ nprv_hyp _, simped]
lemmas prv_Num_LLq_eql_sucE2 = prv_Num_LLq_eql_sucE[OF _ _ nprv_hyp, simped]
lemmas prv_Num_LLq_eql_sucE01 = prv_Num_LLq_eql_sucE[OF nprv_hyp nprv_hyp _, simped]
lemmas prv_Num_LLq_eql_sucE02 = prv_Num_LLq_eql_sucE[OF nprv_hyp _ nprv_hyp, simped]
lemmas prv_Num_LLq_eql_sucE12 = prv_Num_LLq_eql_sucE[OF _ nprv_hyp nprv_hyp, simped]
lemmas prv_Num_LLq_eql_sucE012 = prv_Num_LLq_eql_sucE[OF nprv_hyp nprv_hyp nprv_hyp, simped]
(* *)
lemmas O7E = prv_Num_LLq_eql_sucE
lemmas O7E0 = prv_Num_LLq_eql_sucE0
(**)
(* Although we work in intuitionistic logic,
Q decides equality of arbitrary entities with numerals: *)
lemma prv_dsj_eql_Num_neg:
assumes "t \<in> atrm"
shows "prv (dsj (eql t (Num n)) (neg (eql t (Num n))))"
using assms proof(induction n arbitrary: t)
case [simp]:(0 t)
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] unfolding z_def[symmetric] by auto
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_zer_suc_casesE[of t _ _ z])
subgoal by (nrule r: nprv_dsjIL)
subgoal by (nprover3 r1: nprv_dsjIR r2: nprv_negI r3: nprv_zer_suc_2contrE01) .
next
- case (Suc n) note `t \<in> atrm`[simp]
+ case (Suc n) note \<open>t \<in> atrm\<close>[simp]
define z where "z \<equiv> getFr [xx,yy,zz] [t] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> zz" "zz \<noteq> z" "z \<notin> FvarsT t"
using getFr_FvarsT_Fvars[of "[xx,yy,zz]" "[t]" "[]"] unfolding z_def[symmetric] by auto
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_zer_suc_casesE[of t _ _ z])
subgoal by (nprover3 r1: nprv_dsjIR r2: nprv_negI r3: nprv_zer_suc_2contrE01)
subgoal
apply(nrule r: nprv_eql_substE [of _ t "suc (Var z)"
"dsj (eql (Var z) (suc (Num n))) (neg (eql (Var z) (suc (Num n))))" z])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_addLemmaE[OF Suc.IH[of "Var z"]])
apply(nrule r: nprv_dsjE0)
subgoal by (nprover2 r1: nprv_dsjIL r2: nprv_suc_congI)
subgoal by (nprover4 r1: nprv_dsjIR r2: nprv_negI r3: nprv_negE0 r4: nprv_eql_sucI) . .
qed
(* Rulification: *)
lemmas nprv_eql_Num_casesE = nprv_addDsjLemmaE[OF prv_dsj_eql_Num_neg, simped, rotated]
lemmas nprv_eql_Num_casesE0 = nprv_eql_Num_casesE[OF nprv_hyp _, simped]
lemmas nprv_eql_Num_casesE1 = nprv_eql_Num_casesE[OF _ nprv_hyp, simped]
lemmas nprv_eql_Num_casesE01 = nprv_eql_Num_casesE[OF nprv_hyp nprv_hyp, simped]
(* *)
lemma prv_LLq_Num_dsj: \<comment> \<open>O8\<close>
assumes [simp]: "t \<in> atrm"
shows "prv (dsj (LLq t (Num n)) (LLq (Num n) t))"
proof(induction n)
case 0
note prv_LLq_zer[simp]
show ?case by (nprover3 r1: nprv_prvI r2: nprv_dsjIR r3: prv_nprvI)
next
case (Suc n)
have nn: "suc (Num n) = Num (Suc n)" by simp
have [simp]: "prv (LLq (Num n) (suc (Num n)))"
apply(subst nn) by (rule leq_prv_LLq_Num) simp
show ?case
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_addLemmaE[OF Suc.IH])
apply(nrule r: nprv_dsjE0)
subgoal by (nprover2 r1: nprv_dsjIL r2: O6I)
subgoal
apply(nrule r: nprv_clear2_2)
apply(nrule2 r: nprv_eql_Num_casesE[of t n])
subgoal by (nprover3 r1: nprv_dsjIL
r2: nprv_eql_substE[of _ t "Num n" "LLq (Var xx) (suc (Num n))" xx]
r3: prv_nprvI)
subgoal
apply(nrule r: O7E0[of n t])
subgoal by (nprover2 r1: nprv_eql_symE0 r2: nprv_negE01 )
subgoal by (nrule r: nprv_dsjIR) . . .
qed
(* Rulification: *)
lemma prv_LLq_Num_casesE:
"nprv (insert (LLq t (Num n)) F) \<psi> \<Longrightarrow>
nprv (insert (LLq (Num n) t) F) \<psi> \<Longrightarrow>
t \<in> atrm \<Longrightarrow> F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> \<psi> \<in> fmla \<Longrightarrow>
nprv F \<psi>"
by (rule nprv_addDsjLemmaE[OF prv_LLq_Num_dsj]) auto
lemmas prv_LLq_Num_casesE0 = prv_LLq_Num_casesE[OF nprv_hyp _, simped]
lemmas prv_LLq_Num_casesE1 = prv_LLq_Num_casesE[OF _ nprv_hyp, simped]
lemmas prv_LLq_Num_casesE01 = prv_LLq_Num_casesE[OF nprv_hyp nprv_hyp, simped]
lemmas O8E = prv_LLq_Num_casesE
lemmas O8E0 = prv_LLq_Num_casesE0
lemmas O8E1 = prv_LLq_Num_casesE1
lemmas O8E01 = prv_LLq_Num_casesE01
(* *)
lemma prv_imp_LLq_neg_Num_suc:
assumes [simp]: "t \<in> atrm"
shows "prv (imp (LLq t (suc (Num n)))
(imp ((neg (eql t (suc (Num n)))))
(LLq t (Num n))))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_impI)
apply(nrule r: O3E0[of t "Suc n"])
apply(nrule r: nprv_clear3_3)
apply(nrule r: nprv_ldsjE0)
subgoal for i
apply(nrule r: nprv_clear3_2)
apply(nrule r: nprv_impI_rev2[of "{eql t (Num i)}" "neg (eql t (suc (Num n)))"])
apply(nrule r: nprv_eql_substE[of _ t "Num i"
"imp (neg (eql (Var xx) (suc (Num n)))) (LLq (Var xx) (Num n))" xx])
apply(nrule r: nprv_clear)
apply(nrule r: nprv_impI)
apply(cases "i = Suc n")
subgoal by (nprover2 r1: nprv_negE0 r2: nprv_eql_reflI)
subgoal by (nprover2 r1: prv_nprvI r2: leq_prv_LLq_Num) . .
(* Rulification *)
lemmas nprv_LLq_neg_Num_sucI = nprv_addImp2LemmaI[OF prv_imp_LLq_neg_Num_suc, simped, rotated 3]
lemmas nprv_LLq_neg_Num_sucE = nprv_addImp2LemmaE[OF prv_imp_LLq_neg_Num_suc, simped, rotated 1]
lemmas nprv_LLq_neg_Num_sucE0 = nprv_LLq_neg_Num_sucE[OF nprv_hyp _ _, simped]
lemmas nprv_LLq_neg_Num_sucE1 = nprv_LLq_neg_Num_sucE[OF _ nprv_hyp _, simped]
lemmas nprv_LLq_neg_Num_sucE2 = nprv_LLq_neg_Num_sucE[OF _ _ nprv_hyp, simped]
lemmas nprv_LLq_neg_Num_sucE01 = nprv_LLq_neg_Num_sucE[OF nprv_hyp nprv_hyp _, simped]
lemmas nprv_LLq_neg_Num_sucE02 = nprv_LLq_neg_Num_sucE[OF nprv_hyp _ nprv_hyp, simped]
lemmas nprv_LLq_neg_Num_sucE12 = nprv_LLq_neg_Num_sucE[OF _ nprv_hyp nprv_hyp, simped]
lemmas nprv_LLq_neg_Num_sucE012 = nprv_LLq_neg_Num_sucE[OF nprv_hyp nprv_hyp nprv_hyp, simped]
(* *)
lemma prv_ball_Num_imp_ball_suc: \<comment> \<open>O9\<close>
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla"
shows "prv (imp (ball x (Num n) \<phi>)
(ball x (suc (Num n)) (imp (neg (eql (Var x) (suc (Num n)))) \<phi>)))"
apply(nrule r: nprv_prvI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_ballI)
apply(nrule r: nprv_impI)
apply(nrule r: nprv_LLq_neg_Num_sucE01)
apply(nrule r: nprv_clear4_2)
apply(nrule r: nprv_clear3_2)
apply(nrule r: nprv_ballE0[of x "Num n" \<phi> _ "Var x"]) .
(* Rulification: *)
lemmas prv_ball_Num_ball_sucI = nprv_addImpLemmaI[OF prv_ball_Num_imp_ball_suc, simped, rotated 4]
lemmas prv_ball_Num_ball_sucE = nprv_addImpLemmaE[OF prv_ball_Num_imp_ball_suc, simped, rotated 2]
lemmas prv_ball_Num_ball_sucE0 = prv_ball_Num_ball_sucE[OF nprv_hyp _, simped]
lemmas prv_ball_Num_ball_sucE1 = prv_ball_Num_ball_sucE[OF _ nprv_hyp, simped]
lemmas prv_ball_Num_ball_sucE01 = prv_ball_Num_ball_sucE[OF nprv_hyp nprv_hyp, simped]
lemmas O9I = prv_ball_Num_ball_sucI
lemmas O9E = prv_ball_Num_ball_sucE
lemmas O9E0 = prv_ball_Num_ball_sucE0
lemmas O9E1 = prv_ball_Num_ball_sucE1
lemmas O9E01 = prv_ball_Num_ball_sucE01
subsection \<open>Verifying the abstract ordering assumptions\<close>
lemma LLq_num:
assumes \<phi>[simp]: "\<phi> \<in> fmla" "Fvars \<phi> = {zz}" and q: "q \<in> num" and p: "\<forall> p \<in> num. prv (subst \<phi> p zz)"
shows "prv (all zz (imp (LLq (Var zz) q) \<phi>))"
proof-
obtain n where q: "q = Num n" using q num_Num by auto
\<comment> \<open>NB: We did not need the whole strength of the assumption p -- we only needed that to hold for
numerals smaller than n. However, the abstract framework allowed us to make this strong assumption,
and did not need to even assume an order on the numerals.\<close>
show ?thesis unfolding q ball_def[symmetric] using p p num_Num by (intro O4) auto
qed
lemma LLq_num2:
assumes "p \<in> num"
shows "\<exists>P\<subseteq>num. finite P \<and> prv (dsj (sdsj {eql (Var yy) r |r. r \<in> P}) (LLq p (Var yy)))"
proof-
obtain n where q[simp]: "p = Num n" using assms num_Num by auto
have [simp]: "{eql (Var yy) r |r. \<exists>i. r = Num i \<and> i \<le> n} \<subseteq> fmla" by auto
show ?thesis
apply(nrule r: exI[of _ "{Num i | i . i \<le> n}"])
apply(nrule r: nprv_prvI)
apply(nrule r: O8E[of "Var yy" n])
subgoal
apply(nrule r: nprv_dsjIL)
apply(nrule r: O3E0)
apply(nrule r: nprv_ldsjE0)
apply(nrule r: nprv_sdsjI[of _ "eql (Var yy) (Num _)"])
apply(nrule r: nprv_hyp) .
subgoal by (nrule r: nprv_dsjIR) .
qed
end \<comment> \<open>context @{locale Deduct_Q}\<close>
sublocale Deduct_Q < lab: Deduct_with_PseudoOrder where Lq = "LLq (Var zz) (Var yy)"
apply standard apply auto[] using Fvars_Lq apply auto[]
using LLq_num LLq_num2 apply auto
done
(*<*)
end
(*>*)
diff --git a/thys/Syntax_Independent_Logic/Syntax.thy b/thys/Syntax_Independent_Logic/Syntax.thy
--- a/thys/Syntax_Independent_Logic/Syntax.thy
+++ b/thys/Syntax_Independent_Logic/Syntax.thy
@@ -1,2374 +1,2374 @@
chapter \<open>Syntax\<close>
(*<*)
theory Syntax
imports Prelim
begin
(*>*)
section \<open>Generic Syntax\<close>
text \<open>We develop some generic (meta-)axioms for syntax and substitution.
We only assume that the syntax of our logic has notions of variable, term and formula,
which \emph{include} subsets of "numeric" variables, terms and formulas,
the latter being endowed with notions of free variables and substitution subject to
some natural properties.\<close>
locale Generic_Syntax =
fixes
var :: "'var set" \<comment> \<open>numeric variables (i.e., variables ranging over numbers)\<close>
and trm :: "'trm set" \<comment> \<open>numeric trms, which include the numeric variables\<close>
and fmla :: "'fmla set" \<comment> \<open>numeric formulas\<close>
and Var :: "'var \<Rightarrow> 'trm" \<comment> \<open>trms include at least the variables\<close>
and FvarsT :: "'trm \<Rightarrow> 'var set" \<comment> \<open>free variables for trms\<close>
and substT :: "'trm \<Rightarrow> 'trm \<Rightarrow> 'var \<Rightarrow> 'trm" \<comment> \<open>substitution for trms\<close>
and Fvars :: "'fmla \<Rightarrow> 'var set" \<comment> \<open>free variables for formulas\<close>
and subst :: "'fmla \<Rightarrow> 'trm \<Rightarrow> 'var \<Rightarrow> 'fmla" \<comment> \<open>substitution for formulas\<close>
assumes
infinite_var: "infinite var" \<comment> \<open>the variables are assumed infinite\<close>
and \<comment> \<open>Assumptions about the infrastructure (free vars, substitution and the embedding of variables into trms.
NB: We need fewer assumptions for trm substitution than for formula substitution!\<close>
Var[simp,intro!]: "\<And>x. x \<in> var \<Longrightarrow> Var x \<in> trm"
and
inj_Var[simp]: "\<And> x y. x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> (Var x = Var y \<longleftrightarrow> x = y)"
and
finite_FvarsT: "\<And> t. t \<in> trm \<Longrightarrow> finite (FvarsT t)"
and
FvarsT: "\<And>t. t \<in> trm \<Longrightarrow> FvarsT t \<subseteq> var"
and
substT[simp,intro]: "\<And>t1 t x. t1 \<in> trm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> substT t1 t x \<in> trm"
and
FvarsT_Var[simp]: "\<And> x. x \<in> var \<Longrightarrow> FvarsT (Var x) = {x}"
and
substT_Var[simp]: "\<And> x t y. x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> t \<in> trm \<Longrightarrow>
substT (Var x) t y = (if x = y then t else Var x)"
and
substT_notIn[simp]:
"\<And>t1 t2 x. x \<in> var \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x \<notin> FvarsT t1 \<Longrightarrow> substT t1 t2 x = t1"
and
\<comment> \<open>Assumptions about the infrastructure (free vars and substitution) on formulas\<close>
finite_Fvars: "\<And> \<phi>. \<phi> \<in> fmla \<Longrightarrow> finite (Fvars \<phi>)"
and
Fvars: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars \<phi> \<subseteq> var"
and
subst[simp,intro]: "\<And>\<phi> t x. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> subst \<phi> t x \<in> fmla"
and
Fvars_subst_in:
"\<And> \<phi> t x. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> x \<in> Fvars \<phi> \<Longrightarrow>
Fvars (subst \<phi> t x) = Fvars \<phi> - {x} \<union> FvarsT t"
and
subst_compose_eq_or:
"\<And> \<phi> t1 t2 x1 x2. \<phi> \<in> fmla \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x1 \<in> var \<Longrightarrow> x2 \<in> var \<Longrightarrow>
x1 = x2 \<or> x2 \<notin> Fvars \<phi> \<Longrightarrow> subst (subst \<phi> t1 x1) t2 x2 = subst \<phi> (substT t1 t2 x2) x1"
and
subst_compose_diff:
"\<And> \<phi> t1 t2 x1 x2. \<phi> \<in> fmla \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x1 \<in> var \<Longrightarrow> x2 \<in> var \<Longrightarrow>
x1 \<noteq> x2 \<Longrightarrow> x1 \<notin> FvarsT t2 \<Longrightarrow>
subst (subst \<phi> t1 x1) t2 x2 = subst (subst \<phi> t2 x2) (substT t1 t2 x2) x1"
and
subst_same_Var[simp]:
"\<And> \<phi> x. \<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> subst \<phi> (Var x) x = \<phi>"
and
subst_notIn[simp]:
"\<And> x \<phi> t. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> x \<notin> Fvars \<phi> \<Longrightarrow> subst \<phi> t x = \<phi>"
begin
lemma var_NE: "var \<noteq> {}"
using infinite_var by auto
lemma Var_injD: "Var x = Var y \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> x = y"
by auto
lemma FvarsT_VarD: "x \<in> FvarsT (Var y) \<Longrightarrow> y \<in> var \<Longrightarrow> x = y"
by auto
lemma FvarsT': "t \<in> trm \<Longrightarrow> x \<in> FvarsT t \<Longrightarrow> x \<in> var"
using FvarsT by auto
lemma Fvars': "\<phi> \<in> fmla \<Longrightarrow> x \<in> Fvars \<phi> \<Longrightarrow> x \<in> var"
using Fvars by auto
lemma Fvars_subst[simp]:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
Fvars (subst \<phi> t x) = (Fvars \<phi> - {x}) \<union> (if x \<in> Fvars \<phi> then FvarsT t else {})"
by (simp add: Fvars_subst_in)
lemma in_Fvars_substD:
"y \<in> Fvars (subst \<phi> t x) \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var
\<Longrightarrow> y \<in> (Fvars \<phi> - {x}) \<union> (if x \<in> Fvars \<phi> then FvarsT t else {})"
using Fvars_subst by auto
lemma inj_on_Var: "inj_on Var var"
using inj_Var unfolding inj_on_def by auto
lemma subst_compose_same:
"\<And> \<phi> t1 t2 x. \<phi> \<in> fmla \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (subst \<phi> t1 x) t2 x = subst \<phi> (substT t1 t2 x) x"
using subst_compose_eq_or by blast
lemma subst_subst[simp]:
assumes \<phi>[simp]: "\<phi> \<in> fmla" and t[simp]:"t \<in> trm" and x[simp]:"x \<in> var" and y[simp]:"y \<in> var"
assumes yy: "x \<noteq> y" "y \<notin> Fvars \<phi>"
shows "subst (subst \<phi> (Var y) x) t y = subst \<phi> t x"
using subst_compose_eq_or[OF \<phi> _ t x y, of "Var y"] using subst_notIn yy by simp
lemma subst_comp:
"\<And> x y \<phi> t. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow>
x \<noteq> y \<Longrightarrow> y \<notin> FvarsT t \<Longrightarrow>
subst (subst \<phi> (Var x) y) t x = subst (subst \<phi> t x) t y"
by (simp add: subst_compose_diff)
lemma exists_nat_var:
"\<exists> f::nat\<Rightarrow>'var. inj f \<and> range f \<subseteq> var"
by (simp add: infinite_countable_subset infinite_var)
definition Variable :: "nat \<Rightarrow> 'var" where
"Variable = (SOME f. inj f \<and> range f \<subseteq> var)"
lemma Variable_inj_var:
"inj Variable \<and> range Variable \<subseteq> var"
unfolding Variable_def using someI_ex[OF exists_nat_var] .
lemma inj_Variable[simp]: "\<And> i j. Variable i = Variable j \<longleftrightarrow> i = j"
and Variable[simp,intro!]: "\<And>i. Variable i \<in> var"
using Variable_inj_var image_def unfolding inj_def by auto
text \<open>Convenient notations for some variables
We reserve the first 10 indexes for any special variables we
may wish to consider later.\<close>
abbreviation xx where "xx \<equiv> Variable 10"
abbreviation yy where "yy \<equiv> Variable 11"
abbreviation zz where "zz \<equiv> Variable 12"
abbreviation xx' where "xx' \<equiv> Variable 13"
abbreviation yy' where "yy' \<equiv> Variable 14"
abbreviation zz' where "zz' \<equiv> Variable 15"
lemma xx: "xx \<in> var"
and yy: "yy \<in> var"
and zz: "zz \<in> var"
and xx': "xx' \<in> var"
and yy': "yy' \<in> var"
and zz': "zz' \<in> var"
by auto
lemma vars_distinct[simp]:
"xx \<noteq> yy" "yy \<noteq> xx" "xx \<noteq> zz" "zz \<noteq> xx" "xx \<noteq> xx'" "xx' \<noteq> xx" "xx \<noteq> yy'" "yy' \<noteq> xx" "xx \<noteq> zz'" "zz' \<noteq> xx"
"yy \<noteq> zz" "zz \<noteq> yy" "yy \<noteq> xx'" "xx' \<noteq> yy" "yy \<noteq> yy'" "yy' \<noteq> yy" "yy \<noteq> zz'" "zz' \<noteq> yy"
"zz \<noteq> xx'" "xx' \<noteq> zz" "zz \<noteq> yy'" "yy' \<noteq> zz" "zz \<noteq> zz'" "zz' \<noteq> zz"
"xx' \<noteq> yy'" "yy' \<noteq> xx'" "xx' \<noteq> zz'" "zz' \<noteq> xx'"
"yy' \<noteq> zz'" "zz' \<noteq> yy'"
by auto
subsection \<open>Instance Operator\<close>
definition inst :: "'fmla \<Rightarrow> 'trm \<Rightarrow> 'fmla" where
"inst \<phi> t = subst \<phi> t xx"
lemma inst[simp]: "\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> inst \<phi> t \<in> fmla"
unfolding inst_def by auto
definition getFresh :: "'var set \<Rightarrow> 'var" where
"getFresh V = (SOME x. x \<in> var \<and> x \<notin> V)"
lemma getFresh: "finite V \<Longrightarrow> getFresh V \<in> var \<and> getFresh V \<notin> V"
by (metis (no_types, lifting) finite_subset getFresh_def infinite_var someI_ex subsetI)
definition getFr :: "'var list \<Rightarrow> 'trm list \<Rightarrow> 'fmla list \<Rightarrow> 'var" where
"getFr xs ts \<phi>s =
getFresh (set xs \<union> (\<Union>(FvarsT ` set ts)) \<union> (\<Union>(Fvars ` set \<phi>s)))"
lemma getFr_FvarsT_Fvars:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "getFr xs ts \<phi>s \<in> var \<and>
getFr xs ts \<phi>s \<notin> set xs \<and>
(t \<in> set ts \<longrightarrow> getFr xs ts \<phi>s \<notin> FvarsT t) \<and>
(\<phi> \<in> set \<phi>s \<longrightarrow> getFr xs ts \<phi>s \<notin> Fvars \<phi>)"
proof-
have "finite (set xs \<union> (\<Union>(FvarsT ` set ts)) \<union> (\<Union>(Fvars ` set \<phi>s)))"
using assms finite_FvarsT finite_Fvars by auto
from getFresh[OF this] show ?thesis using assms unfolding getFr_def by auto
qed
lemma getFr[simp,intro]:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "getFr xs ts \<phi>s \<in> var"
using assms getFr_FvarsT_Fvars by auto
lemma getFr_var:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "t \<in> set ts"
shows "getFr xs ts \<phi>s \<notin> set xs"
using assms getFr_FvarsT_Fvars by auto
lemma getFr_FvarsT:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "t \<in> set ts"
shows "getFr xs ts \<phi>s \<notin> FvarsT t"
using assms getFr_FvarsT_Fvars by auto
lemma getFr_Fvars:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "\<phi> \<in> set \<phi>s"
shows "getFr xs ts \<phi>s \<notin> Fvars \<phi>"
using assms getFr_FvarsT_Fvars by auto
subsection \<open>Fresh Variables\<close>
fun getFreshN :: "'var set \<Rightarrow> nat \<Rightarrow> 'var list" where
"getFreshN V 0 = []"
| "getFreshN V (Suc n) = (let u = getFresh V in u # getFreshN (insert u V) n)"
lemma getFreshN: "finite V \<Longrightarrow>
set (getFreshN V n) \<subseteq> var \<and> set (getFreshN V n) \<inter> V = {} \<and> length (getFreshN V n) = n \<and> distinct (getFreshN V n)"
by (induct n arbitrary: V) (auto simp: getFresh Let_def)
definition getFrN :: "'var list \<Rightarrow> 'trm list \<Rightarrow> 'fmla list \<Rightarrow> nat \<Rightarrow> 'var list" where
"getFrN xs ts \<phi>s n =
getFreshN (set xs \<union> (\<Union>(FvarsT ` set ts)) \<union> (\<Union>(Fvars ` set \<phi>s))) n"
lemma getFrN_FvarsT_Fvars:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "set (getFrN xs ts \<phi>s n) \<subseteq> var \<and>
set (getFrN xs ts \<phi>s n) \<inter> set xs = {} \<and>
(t \<in> set ts \<longrightarrow> set (getFrN xs ts \<phi>s n) \<inter> FvarsT t = {}) \<and>
(\<phi> \<in> set \<phi>s \<longrightarrow> set (getFrN xs ts \<phi>s n) \<inter> Fvars \<phi> = {}) \<and>
length (getFrN xs ts \<phi>s n) = n \<and>
distinct (getFrN xs ts \<phi>s n)"
proof-
have "finite (set xs \<union> (\<Union>(FvarsT ` set ts)) \<union> (\<Union>(Fvars ` set \<phi>s)))"
using assms finite_FvarsT finite_Fvars by auto
from getFreshN[OF this] show ?thesis using assms unfolding getFrN_def by auto
qed
lemma getFrN[simp,intro]:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "set (getFrN xs ts \<phi>s n) \<subseteq> var"
using assms getFrN_FvarsT_Fvars by auto
lemma getFrN_var:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "t \<in> set ts"
shows "set (getFrN xs ts \<phi>s n) \<inter> set xs = {}"
using assms getFrN_FvarsT_Fvars by auto
lemma getFrN_FvarsT:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "t \<in> set ts"
shows "set (getFrN xs ts \<phi>s n) \<inter> FvarsT t = {}"
using assms getFrN_FvarsT_Fvars by auto
lemma getFrN_Fvars:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla" and "\<phi> \<in> set \<phi>s"
shows "set (getFrN xs ts \<phi>s n) \<inter> Fvars \<phi> = {}"
using assms getFrN_FvarsT_Fvars by auto
lemma getFrN_length:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "length (getFrN xs ts \<phi>s n) = n"
using assms getFrN_FvarsT_Fvars by auto
lemma getFrN_distinct[simp,intro]:
assumes "set xs \<subseteq> var" "set ts \<subseteq> trm" and "set \<phi>s \<subseteq> fmla"
shows "distinct (getFrN xs ts \<phi>s n)"
using assms getFrN_FvarsT_Fvars by auto
subsection \<open>Parallel Term Substitution\<close>
fun rawpsubstT :: "'trm \<Rightarrow> ('trm \<times> 'var) list \<Rightarrow> 'trm" where
"rawpsubstT t [] = t"
| "rawpsubstT t ((t1,x1) # txs) = rawpsubstT (substT t t1 x1) txs"
lemma rawpsubstT[simp]:
assumes "t \<in> trm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "rawpsubstT t txs \<in> trm"
using assms by (induct txs arbitrary: t) fastforce+
definition psubstT :: "'trm \<Rightarrow> ('trm \<times> 'var) list \<Rightarrow> 'trm" where
"psubstT t txs =
(let xs = map snd txs; ts = map fst txs; us = getFrN xs (t # ts) [] (length xs) in
rawpsubstT (rawpsubstT t (zip (map Var us) xs)) (zip ts us))"
text \<open>The psubstT versions of the subst axioms.\<close>
lemma psubstT[simp,intro]:
assumes "t \<in> trm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "psubstT t txs \<in> trm"
proof-
define us where us: "us = getFrN (map snd txs) (t # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT t = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "t # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "t # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "t # map fst txs" "[]" "length txs"]
getFrN_distinct[of "map snd txs" "t # map fst txs" "[]" "length txs"]
by auto (metis (no_types, hide_lams) IntI empty_iff image_iff old.prod.inject surjective_pairing)
show ?thesis using assms us_facts unfolding psubstT_def
by (force simp: Let_def us[symmetric] dest: set_zip_leftD set_zip_rightD intro!: rawpsubstT)
qed
lemma rawpsubstT_Var_not[simp]:
assumes "x \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "x \<notin> snd ` (set txs)"
shows "rawpsubstT (Var x) txs = Var x"
using assms by (induct txs) auto
lemma psubstT_Var_not[simp]:
assumes "x \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "x \<notin> snd ` (set txs)"
shows "psubstT (Var x) txs = Var x"
proof-
define us where us: "us = getFrN (map snd txs) (Var x # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"x \<notin> set us"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "Var x # map fst txs" "[]" "Var x" "length txs"]
getFrN_FvarsT[of "map snd txs" "Var x # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "Var x # map fst txs" "[]" "Var x" "length txs"]
getFrN_length[of "map snd txs" "Var x # map fst txs" "[]" "length txs"]
by (auto simp: set_eq_iff)
have [simp]: "rawpsubstT (Var x) (zip (map Var us) (map snd txs)) = Var x"
using assms us_facts
by(intro rawpsubstT_Var_not) (force dest: set_zip_rightD set_zip_leftD)+
have [simp]: "rawpsubstT (Var x) (zip (map fst txs) us) = Var x"
using assms us_facts
by(intro rawpsubstT_Var_not) (force dest: set_zip_rightD set_zip_leftD)+
show ?thesis using assms us_facts unfolding psubstT_def
by (auto simp: Let_def us[symmetric])
qed
lemma rawpsubstT_notIn[simp]:
assumes "x \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "t \<in> trm"
and "FvarsT t \<inter> snd ` (set txs) = {}"
shows "rawpsubstT t txs = t"
using assms by (induct txs) auto
lemma psubstT_notIn[simp]:
assumes "x \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "t \<in> trm"
and "FvarsT t \<inter> snd ` (set txs) = {}"
shows "psubstT t txs = t"
proof-
define us where us: "us = getFrN (map snd txs) (t # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT t = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "t # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "t # map fst txs" "[]" t "length txs"]
getFrN_length[of "map snd txs" "t # map fst txs" "[]" "length txs"]
by (auto simp: set_eq_iff)
have [simp]: "rawpsubstT t (zip (map Var us) (map snd txs)) = t"
using assms us_facts
by(intro rawpsubstT_notIn) (auto 0 3 dest: set_zip_rightD set_zip_leftD)
have [simp]: "rawpsubstT t (zip (map fst txs) us) = t"
using assms us_facts
by(intro rawpsubstT_notIn) (auto 0 3 dest: set_zip_rightD set_zip_leftD)
show ?thesis using assms us_facts unfolding psubstT_def
by (auto simp: Let_def us[symmetric])
qed
subsection \<open>Parallel Formula Substitution\<close>
fun rawpsubst :: "'fmla \<Rightarrow> ('trm \<times> 'var) list \<Rightarrow> 'fmla" where
"rawpsubst \<phi> [] = \<phi>"
| "rawpsubst \<phi> ((t1,x1) # txs) = rawpsubst (subst \<phi> t1 x1) txs"
lemma rawpsubst[simp]:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst \<phi> txs \<in> fmla"
using assms by (induct txs arbitrary: \<phi>) fastforce+
definition psubst :: "'fmla \<Rightarrow> ('trm \<times> 'var) list \<Rightarrow> 'fmla" where
"psubst \<phi> txs =
(let xs = map snd txs; ts = map fst txs; us = getFrN xs ts [\<phi>] (length xs) in
rawpsubst (rawpsubst \<phi> (zip (map Var us) xs)) (zip ts us))"
text \<open>The psubst versions of the subst axioms.\<close>
lemma psubst[simp,intro]:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "psubst \<phi> txs \<in> fmla"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" \<phi> "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
getFrN_distinct[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
by (auto 8 0 simp: set_eq_iff image_iff Bex_def Ball_def)
show ?thesis using assms us_facts unfolding psubst_def
by (auto 0 3 simp: Let_def us[symmetric] dest: set_zip_rightD set_zip_leftD intro!: rawpsubst)
qed
lemma Fvars_rawpsubst_su:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "Fvars (rawpsubst \<phi> txs) \<subseteq>
(Fvars \<phi> - snd ` (set txs)) \<union> (\<Union> {FvarsT t | t x . (t,x) \<in> set txs})"
using assms proof(induction txs arbitrary: \<phi>)
case (Cons tx txs \<phi>)
then obtain t x where tx: "tx = (t,x)" by force
have t: "t \<in> trm" and x: "x \<in> var" using Cons.prems unfolding tx by auto
define \<chi> where "\<chi> = subst \<phi> t x"
have 0: "Fvars \<chi> = Fvars \<phi> - {x} \<union> (if x \<in> Fvars \<phi> then FvarsT t else {})"
using Cons.prems unfolding \<chi>_def by (auto simp: tx t)
have \<chi>: "\<chi> \<in> fmla" unfolding \<chi>_def using Cons.prems t x by auto
have "Fvars (rawpsubst \<chi> txs) \<subseteq>
(Fvars \<chi> - snd ` (set txs)) \<union>
(\<Union> {FvarsT t | t x . (t,x) \<in> set txs})"
using Cons.prems \<chi> by (intro Cons.IH) auto
also have "\<dots> \<subseteq> Fvars \<phi> - insert x (snd ` set txs) \<union> \<Union>{FvarsT ta |ta. \<exists>xa. ta = t \<and> xa = x \<or> (ta, xa) \<in> set txs}"
(is "_ \<subseteq> ?R") by(auto simp: 0 tx Cons.prems)
finally have 1: "Fvars (rawpsubst \<chi> txs) \<subseteq> ?R" .
have 2: "Fvars \<chi> = Fvars \<phi> - {x} \<union> (if x \<in> Fvars \<phi> then FvarsT t else {})"
using Cons.prems t x unfolding \<chi>_def using Fvars_subst by auto
show ?case using 1 by (simp add: tx \<chi>_def[symmetric] 2)
qed auto
lemma in_Fvars_rawpsubst_imp:
assumes "y \<in> Fvars (rawpsubst \<phi> txs)"
and "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "(y \<in> Fvars \<phi> - snd ` (set txs)) \<or>
(y \<in> \<Union> { FvarsT t | t x . (t,x) \<in> set txs})"
using Fvars_rawpsubst_su[OF assms(2-4)]
using assms(1) by blast
lemma Fvars_rawpsubst:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)" and "\<forall> x \<in> snd ` (set txs). \<forall> t \<in> fst ` (set txs). x \<notin> FvarsT t"
shows "Fvars (rawpsubst \<phi> txs) =
(Fvars \<phi> - snd ` (set txs)) \<union>
(\<Union> {if x \<in> Fvars \<phi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using assms proof(induction txs arbitrary: \<phi>)
case (Cons a txs \<phi>)
then obtain t x where a: "a = (t,x)" by force
have t: "t \<in> trm" and x: "x \<in> var" using Cons.prems unfolding a by auto
- have x_txs: "\<And>ta xa. (ta, xa) \<in> set txs \<Longrightarrow> x \<noteq> xa" using `distinct (map snd (a # txs))`
+ have x_txs: "\<And>ta xa. (ta, xa) \<in> set txs \<Longrightarrow> x \<noteq> xa" using \<open>distinct (map snd (a # txs))\<close>
unfolding a by (auto simp: rev_image_eqI)
have xt: "x \<notin> FvarsT t \<and> snd ` set txs \<inter> FvarsT t = {}" using Cons.prems unfolding a by auto
hence 0: "Fvars \<phi> - {x} \<union> FvarsT t - snd ` set txs =
Fvars \<phi> - insert x (snd ` set txs) \<union> FvarsT t"
by auto
define \<chi> where \<chi>_def: "\<chi> = subst \<phi> t x"
have \<chi>: "\<chi> \<in> fmla" unfolding \<chi>_def using Cons.prems t x by auto
have 1: "Fvars (rawpsubst \<chi> txs) =
(Fvars \<chi> - snd ` (set txs)) \<union>
(\<Union> {if x \<in> Fvars \<chi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using Cons.prems \<chi> by (intro Cons.IH) auto
have 2: "Fvars \<chi> = Fvars \<phi> - {x} \<union> (if x \<in> Fvars \<phi> then FvarsT t else {})"
using Cons.prems t x unfolding \<chi>_def using Fvars_subst by auto
define f where "f \<equiv> \<lambda>ta xa. if xa \<in> Fvars \<phi> then FvarsT ta else {}"
have 3: "\<Union> {f ta xa |ta xa. (ta, xa) \<in> set ((t, x) # txs)} =
f t x \<union> (\<Union> {f ta xa |ta xa. (ta, xa) \<in> set txs})" by auto
have 4: "snd ` set ((t, x) # txs) = {x} \<union> snd ` set txs" by auto
have 5: "f t x \<inter> snd ` set txs = {}" unfolding f_def using xt by auto
have 6: "\<Union> {if xa \<in> Fvars \<phi> - {x} \<union> f t x then FvarsT ta else {} | ta xa. (ta, xa) \<in> set txs}
= (\<Union> {f ta xa | ta xa. (ta, xa) \<in> set txs})"
unfolding f_def using xt x_txs by (fastforce split: if_splits)
have "Fvars \<phi> - {x} \<union> f t x - snd ` set txs \<union>
\<Union> {if xa \<in> Fvars \<phi> - {x} \<union> f t x then FvarsT ta else {}
| ta xa. (ta, xa) \<in> set txs} =
Fvars \<phi> - snd ` set ((t, x) # txs) \<union>
\<Union> {f ta xa |ta xa. (ta, xa) \<in> set ((t, x) # txs)}"
unfolding 3 4 6 unfolding Un_Diff2[OF 5] Un_assoc unfolding Diff_Diff_Un ..
thus ?case unfolding a rawpsubst.simps 1 2 \<chi>_def[symmetric] f_def by simp
qed auto
lemma in_Fvars_rawpsubstD:
assumes "y \<in> Fvars (rawpsubst \<phi> txs)"
and "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)" and "\<forall> x \<in> snd ` (set txs). \<forall> t \<in> fst ` (set txs). x \<notin> FvarsT t"
shows "(y \<in> Fvars \<phi> - snd ` (set txs)) \<or>
(y \<in> \<Union> {if x \<in> Fvars \<phi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using Fvars_rawpsubst assms by auto
lemma Fvars_psubst:
assumes "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "Fvars (psubst \<phi> txs) =
(Fvars \<phi> - snd ` (set txs)) \<union>
(\<Union> {if x \<in> Fvars \<phi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
by (auto 9 0 simp: set_eq_iff image_iff)
define \<chi> where \<chi>_def: "\<chi> = rawpsubst \<phi> (zip (map Var us) (map snd txs))"
have \<chi>: "\<chi> \<in> fmla" unfolding \<chi>_def using assms us_facts
by (intro rawpsubst) (auto dest!: set_zip_D)
have set_us: "set us = snd ` (set (zip (map fst txs) us))"
using us_facts by (intro snd_set_zip[symmetric]) auto
have set_txs: "snd ` set txs = snd ` (set (zip (map Var us) (map snd txs)))"
using us_facts by (intro snd_set_zip_map_snd[symmetric]) auto
have "\<And> t x. (t, x) \<in> set (zip (map Var us) (map snd txs)) \<Longrightarrow> \<exists> u. t = Var u"
using us_facts set_zip_leftD by fastforce
hence 00: "\<And> t x. (t, x) \<in> set (zip (map Var us) (map snd txs))
\<longleftrightarrow> (\<exists> u \<in> var. t = Var u \<and> (Var u, x) \<in> set (zip (map Var us) (map snd txs)))"
using us_facts set_zip_leftD by fastforce
have "Fvars \<chi> =
Fvars \<phi> - snd ` set txs \<union>
\<Union>{if x \<in> Fvars \<phi> then FvarsT t else {} |t x.
(t, x) \<in> set (zip (map Var us) (map snd txs))}"
unfolding \<chi>_def set_txs using assms us_facts
apply(intro Fvars_rawpsubst)
subgoal by auto
subgoal by (auto dest!: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by auto
subgoal by (auto 0 6 simp: set_txs[symmetric] set_eq_iff subset_eq image_iff in_set_zip
dest: spec[where P="\<lambda>x. x \<in> set us \<longrightarrow> (\<forall>y \<in> set txs. x \<noteq> snd y)", THEN mp[OF _ nth_mem]]) .
also have "\<dots> =
Fvars \<phi> - snd ` set txs \<union>
\<Union>{if x \<in> Fvars \<phi> then {u} else {} |u x. u \<in> var \<and> (Var u, x) \<in> set (zip (map Var us) (map snd txs))}"
(is "\<dots> = ?R")
using FvarsT_Var by (metis (no_types, hide_lams) 00)
finally have 0: "Fvars \<chi> = ?R" .
have 1: "Fvars (rawpsubst \<chi> (zip (map fst txs) us)) =
(Fvars \<chi> - set us) \<union>
(\<Union> {if u \<in> Fvars \<chi> then FvarsT t else {} | t u . (t,u) \<in> set (zip (map fst txs) us)})"
unfolding set_us using us_facts assms \<chi>
apply (intro Fvars_rawpsubst)
subgoal by (auto dest: set_zip_rightD)
subgoal by (auto dest: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by (metis IntI Union_iff empty_iff fst_set_zip_map_fst image_eqI set_us) .
have 2: "Fvars \<chi> - set us = Fvars \<phi> - snd ` set txs"
unfolding 0 using us_facts(1,2)
by (fastforce dest!: set_zip_leftD split: if_splits)
have 3:
"(\<Union> {if u \<in> Fvars \<chi> then FvarsT t else {} | t u . (t,u) \<in> set (zip (map fst txs) us)}) =
(\<Union> {if x \<in> Fvars \<phi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
proof safe
fix xx tt y
assume xx: "xx \<in> (if y \<in> Fvars \<chi> then FvarsT tt else {})"
and ty: "(tt, y) \<in> set (zip (map fst txs) us)"
have ttin: "tt \<in> fst ` set txs" using ty using set_zip_leftD by fastforce
have yin: "y \<in> set us" using ty by (meson set_zip_D)
have yvar: "y \<in> var" using us_facts yin by auto
have ynotin: "y \<notin> snd ` set txs" "y \<notin> Fvars \<phi>" using yin us_facts by auto
show "xx \<in> \<Union>{if x \<in> Fvars \<phi> then FvarsT t else {} |t x. (t, x) \<in> set txs}"
proof(cases "y \<in> Fvars \<chi>")
case True note y = True
hence xx: "xx \<in> FvarsT tt" using xx by simp
obtain x where x\<phi>: "x \<in> Fvars \<phi>"
and yx: "(Var y, x) \<in> set (zip (map Var us) (map snd txs))"
using y ynotin unfolding 0 by auto (metis empty_iff insert_iff)
have yx: "(y, x) \<in> set (zip us (map snd txs))"
using yvar us_facts by (intro inj_on_set_zip_map[OF inj_on_Var yx]) auto
have "(tt, x) \<in> set txs" apply(rule set_zip_map_fst_snd[OF yx ty])
- using `distinct (map snd txs)` us_facts by auto
+ using \<open>distinct (map snd txs)\<close> us_facts by auto
thus ?thesis using xx x\<phi> by auto
qed(insert xx, auto)
next
fix y tt xx
assume y: "y \<in> (if xx \<in> Fvars \<phi> then FvarsT tt else {})"
and tx: "(tt, xx) \<in> set txs"
hence xxsnd: "xx \<in> snd ` set txs" by force
obtain u where uin: "u \<in> set us" and uxx: "(u, xx) \<in> set (zip us (map snd txs))"
by (metis xxsnd in_set_impl_in_set_zip2 length_map set_map set_zip_leftD us_facts(5))
hence uvar: "u \<in> var" using us_facts by auto
show "y \<in> \<Union>{if u \<in> Fvars \<chi> then FvarsT t else {} |t u. (t, u) \<in> set (zip (map fst txs) us)}"
proof(cases "xx \<in> Fvars \<phi>")
case True note xx = True
hence y: "y \<in> FvarsT tt" using y by auto
have "(Var u, xx) \<in> set (zip (map Var us) (map snd txs))"
using us_facts by (intro set_zip_length_map[OF uxx]) auto
hence u\<chi>: "u \<in> Fvars \<chi>" using uin xx uvar unfolding 0 by auto
have ttu: "(tt, u) \<in> set (zip (map fst txs) us)"
using assms us_facts by (intro set_zip_map_fst_snd2[OF uxx tx]) auto
show ?thesis using u\<chi> ttu y by auto
qed(insert y, auto)
qed
show ?thesis
by (simp add: psubst_def Let_def us[symmetric] \<chi>_def[symmetric] 1 2 3)
qed
lemma in_Fvars_psubstD:
assumes "y \<in> Fvars (psubst \<phi> txs)"
and "\<phi> \<in> fmla" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "y \<in> (Fvars \<phi> - snd ` (set txs)) \<union>
(\<Union> {if x \<in> Fvars \<phi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using assms Fvars_psubst by auto
lemma subst2_fresh_switch:
assumes "\<phi> \<in> fmla" "t \<in> trm" "s \<in> trm" "x \<in> var" "y \<in> var"
and "x \<noteq> y" "x \<notin> FvarsT s" "y \<notin> FvarsT t"
shows "subst (subst \<phi> s y) t x = subst (subst \<phi> t x) s y" (is "?L = ?R")
using assms by (simp add: subst_compose_diff[of \<phi> s t y x])
lemma rawpsubst2_fresh_switch:
assumes "\<phi> \<in> fmla" "t \<in> trm" "s \<in> trm" "x \<in> var" "y \<in> var"
and "x \<noteq> y" "x \<notin> FvarsT s" "y \<notin> FvarsT t"
shows "rawpsubst \<phi> ([(s,y),(t,x)]) = rawpsubst \<phi> ([(t,x),(s,y)])"
using assms by (simp add: subst2_fresh_switch)
lemma rawpsubst_compose:
assumes "\<phi> \<in> fmla" and "snd ` (set txs1) \<subseteq> var" and "fst ` (set txs1) \<subseteq> trm"
and "snd ` (set txs2) \<subseteq> var" and "fst ` (set txs2) \<subseteq> trm"
shows "rawpsubst (rawpsubst \<phi> txs1) txs2 = rawpsubst \<phi> (txs1 @ txs2)"
using assms by (induct txs1 arbitrary: txs2 \<phi>) auto
lemma rawpsubst_subst_fresh_switch:
assumes "\<phi> \<in> fmla" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "\<forall> x \<in> snd ` (set txs). x \<notin> FvarsT s"
and "\<forall> t \<in> fst ` (set txs). y \<notin> FvarsT t"
and "distinct (map snd txs)"
and "s \<in> trm" and "y \<in> var" "y \<notin> snd ` (set txs)"
shows "rawpsubst (subst \<phi> s y) txs = rawpsubst \<phi> (txs @ [(s,y)])"
using assms proof(induction txs arbitrary: \<phi> s y)
case (Cons tx txs)
obtain t x where tx[simp]: "tx = (t,x)" by force
have x: "x \<in> var" and t: "t \<in> trm" using Cons unfolding tx by auto
have "rawpsubst \<phi> ((s, y) # (t, x) # txs) = rawpsubst \<phi> ([(s, y),(t, x)] @ txs)" by simp
also have "\<dots> = rawpsubst (rawpsubst \<phi> [(s, y),(t, x)]) txs"
using Cons by auto
also have "rawpsubst \<phi> [(s, y),(t, x)] = rawpsubst \<phi> [(t, x),(s, y)]"
using Cons by (intro rawpsubst2_fresh_switch) auto
also have "rawpsubst (rawpsubst \<phi> [(t, x),(s, y)]) txs = rawpsubst \<phi> ([(t, x),(s, y)] @ txs)"
using Cons by auto
also have "\<dots> = rawpsubst (subst \<phi> t x) (txs @ [(s,y)])" using Cons by auto
also have "\<dots> = rawpsubst \<phi> (((t, x) # txs) @ [(s, y)])" by simp
finally show ?case unfolding tx by auto
qed auto
lemma subst_rawpsubst_fresh_switch:
assumes "\<phi> \<in> fmla" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "\<forall> x \<in> snd ` (set txs). x \<notin> FvarsT s"
and "\<forall> t \<in> fst ` (set txs). y \<notin> FvarsT t"
and "distinct (map snd txs)"
and "s \<in> trm" and "y \<in> var" "y \<notin> snd ` (set txs)"
shows "subst (rawpsubst \<phi> txs) s y = rawpsubst \<phi> ((s,y) # txs)"
using assms proof(induction txs arbitrary: \<phi> s y)
case (Cons tx txs)
obtain t x where tx[simp]: "tx = (t,x)" by force
have x: "x \<in> var" and t: "t \<in> trm" using Cons unfolding tx by auto
have "subst (rawpsubst (subst \<phi> t x) txs) s y = rawpsubst (subst \<phi> t x) ((s,y) # txs)"
using Cons.prems by (intro Cons.IH) auto
also have "\<dots> = rawpsubst (rawpsubst \<phi> [(t,x)]) ((s,y) # txs)" by simp
also have "\<dots> = rawpsubst \<phi> ([(t,x)] @ ((s,y) # txs))"
using Cons.prems by auto
also have "\<dots> = rawpsubst \<phi> ([(t,x),(s,y)] @ txs)" by simp
also have "\<dots> = rawpsubst (rawpsubst \<phi> [(t,x),(s,y)]) txs"
using Cons.prems by auto
also have "rawpsubst \<phi> [(t,x),(s,y)] = rawpsubst \<phi> [(s,y),(t,x)]"
using Cons.prems by (intro rawpsubst2_fresh_switch) auto
also have "rawpsubst (rawpsubst \<phi> [(s,y),(t,x)]) txs = rawpsubst \<phi> ([(s,y),(t,x)] @ txs)"
using Cons.prems by auto
finally show ?case by simp
qed auto
lemma rawpsubst_compose_freshVar:
assumes "\<phi> \<in> fmla" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
and "\<And> i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> snd (txs!j) \<notin> FvarsT (fst (txs!i))"
and us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
shows "rawpsubst (rawpsubst \<phi> (zip (map Var us) (map snd txs))) (zip (map fst txs) us) = rawpsubst \<phi> txs"
using assms proof(induction txs arbitrary: us \<phi>)
case (Cons tx txs uus \<phi>)
obtain t x where tx[simp]: "tx = (t,x)" by force
obtain u us where uus[simp]: "uus = u # us" using Cons by (cases uus) auto
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us" and u_facts: "u \<in> var" "u \<notin> Fvars \<phi>"
"u \<notin> \<Union> (FvarsT ` (fst ` (set txs)))"
"u \<notin> snd ` (set txs)" "u \<notin> set us"
using Cons by auto
let ?uxs = "zip (map Var us) (map snd txs)"
have 1: "rawpsubst (subst \<phi> (Var u) x) ?uxs = rawpsubst \<phi> (?uxs @ [(Var u,x)])"
using u_facts Cons.prems
by (intro rawpsubst_subst_fresh_switch) (auto simp: subsetD dest!: set_zip_D)
let ?uuxs = "zip (map Var uus) (map snd (tx # txs))"
let ?tus = "zip (map fst txs) us" let ?ttxs = "zip (map fst (tx # txs)) uus"
have 2: "u \<in> Fvars (rawpsubst \<phi> (zip (map Var us) (map snd txs))) \<Longrightarrow> False"
using Cons.prems apply- apply(drule in_Fvars_rawpsubstD)
subgoal by auto
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by auto
subgoal premises prems using us_facts(1,4,5)
by (auto 0 3 simp: in_set_zip subset_eq set_eq_iff image_iff
dest: spec[where P="\<lambda>x. x \<in> set us \<longrightarrow> (\<forall>y \<in> set txs. x \<noteq> snd y)",
THEN mp[OF _ nth_mem], THEN bspec[OF _ nth_mem]])
subgoal
by (auto simp: in_set_zip subset_eq split: if_splits) .
have 3: "\<And> xx tt. xx \<in> FvarsT t \<Longrightarrow> (tt, xx) \<notin> set txs"
using Cons.prems(4,5) tx unfolding set_conv_nth
by simp (metis One_nat_def Suc_leI diff_Suc_1 fst_conv le_imp_less_Suc
nth_Cons_0 snd_conv zero_less_Suc)
have 00: "rawpsubst (rawpsubst \<phi> ?uuxs) ?ttxs = rawpsubst (subst (rawpsubst \<phi> (?uxs @ [(Var u, x)])) t u) ?tus"
by (simp add: 1)
have "rawpsubst \<phi> (?uxs @ [(Var u, x)]) = rawpsubst (rawpsubst \<phi> ?uxs) [(Var u, x)]"
using Cons.prems by (intro rawpsubst_compose[symmetric]) (auto intro!: rawpsubst dest!: set_zip_D)
also have "rawpsubst (rawpsubst \<phi> ?uxs) [(Var u, x)] = subst (rawpsubst \<phi> ?uxs) (Var u) x" by simp
finally have "subst (rawpsubst \<phi> (?uxs @ [(Var u, x)])) t u =
subst (subst (rawpsubst \<phi> ?uxs) (Var u) x) t u" by simp
also have "\<dots> = subst (rawpsubst \<phi> ?uxs) t x"
using Cons 2 by (intro subst_subst) (auto intro!: rawpsubst dest!: set_zip_D)
also have "\<dots> = rawpsubst \<phi> ((t,x) # ?uxs)"
using Cons.prems 3 apply(intro subst_rawpsubst_fresh_switch)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest!: set_zip_D)
by (auto dest!: set_zip_D)
also have "\<dots> = rawpsubst \<phi> ([(t,x)] @ ?uxs)" by simp
also have "\<dots> = rawpsubst (rawpsubst \<phi> [(t,x)]) ?uxs"
using Cons.prems by (intro rawpsubst_compose[symmetric]) (auto dest!: set_zip_D)
finally have "rawpsubst (subst (rawpsubst \<phi> (?uxs @ [(Var u, x)])) t u) ?tus =
rawpsubst (rawpsubst (rawpsubst \<phi> [(t,x)]) ?uxs) ?tus" by auto
hence "rawpsubst (rawpsubst \<phi> ?uuxs) ?ttxs = rawpsubst (rawpsubst (rawpsubst \<phi> [(t,x)]) ?uxs) ?tus"
using 00 by auto
also have "\<dots> = rawpsubst (rawpsubst \<phi> [(t,x)]) txs"
using Cons.prems apply (intro Cons.IH rawpsubst)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (auto dest!: set_zip_D in_Fvars_substD)
subgoal by (metis Suc_mono diff_Suc_1 length_Cons nat.simps(3) nth_Cons')
by (auto dest!: set_zip_D in_Fvars_substD)
also have "\<dots> = rawpsubst \<phi> ([(t,x)] @ txs)"
using Cons.prems by (intro rawpsubst_compose) (auto dest!: set_zip_D)
finally show ?case by simp
qed auto
lemma rawpsubst_compose_freshVar2_aux:
assumes \<phi>[simp]: "\<phi> \<in> fmla"
and ts: "set ts \<subseteq> trm"
and xs: "set xs \<subseteq> var" "distinct xs"
and us_facts: "set us \<subseteq> var" "distinct us"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set us \<inter> set xs = {}"
and vs_facts: "set vs \<subseteq> var" "distinct vs"
"set vs \<inter> Fvars \<phi> = {}"
"set vs \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set vs \<inter> set xs = {}"
and l: "length us = length xs" "length vs = length xs" "length ts = length xs"
and (* Extra hypothesis, only to get induction through: *) d: "set us \<inter> set vs = {}"
shows "rawpsubst (rawpsubst \<phi> (zip (map Var us) xs)) (zip ts us) =
rawpsubst (rawpsubst \<phi> (zip (map Var vs) xs)) (zip ts vs)"
using assms proof(induction xs arbitrary: \<phi> ts us vs)
case (Cons x xs \<phi> tts uus vvs)
obtain t ts u us v vs where tts[simp]: "tts = t # ts" and lts[simp]: "length ts = length xs"
and uus[simp]: "uus = u # us" and lus[simp]: "length us = length xs"
and vvs[simp]: "vvs = v # vs" and lvs[simp]: "length vs = length xs"
- using `length uus = length (x # xs)` `length vvs = length (x # xs)` `length tts = length (x # xs)`
+ using \<open>length uus = length (x # xs)\<close> \<open>length vvs = length (x # xs)\<close> \<open>length tts = length (x # xs)\<close>
apply(cases tts)
subgoal by auto
subgoal apply(cases uus)
subgoal by auto
subgoal by (cases vvs) auto . .
let ?\<phi>ux = "subst \<phi> (Var u) x" let ?\<phi>vx = "subst \<phi> (Var v) x"
have 0: "rawpsubst (rawpsubst ?\<phi>ux (zip (map Var us) xs)) (zip ts us) =
rawpsubst (rawpsubst ?\<phi>ux (zip (map Var vs) xs)) (zip ts vs)"
apply(rule Cons.IH) using Cons.prems by (auto intro!: rawpsubst dest!: set_zip_D)
have 1: "rawpsubst ?\<phi>ux (zip (map Var vs) xs) =
subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var u) x"
using Cons.prems
by (intro subst_rawpsubst_fresh_switch[simplified,symmetric])
(force intro!: rawpsubst dest!: set_zip_D simp: subset_eq)+
have 11: "rawpsubst ?\<phi>vx (zip (map Var vs) xs) =
subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var v) x"
using Cons.prems
by (intro subst_rawpsubst_fresh_switch[simplified,symmetric])
(auto intro!: rawpsubst dest!: set_zip_D simp: subset_eq)
have "subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var u) x) t u =
subst (rawpsubst \<phi> (zip (map Var vs) xs)) t x"
using Cons.prems
by (intro subst_subst) (force intro!: rawpsubst dest!: set_zip_D in_Fvars_rawpsubst_imp simp: Fvars_rawpsubst)+
also have "\<dots> = subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var v) x) t v"
using Cons.prems
by (intro subst_subst[symmetric])
(force intro!: rawpsubst dest!: set_zip_D in_Fvars_rawpsubst_imp simp: Fvars_rawpsubst)+
finally have
2: "subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var u) x) t u =
subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var v) x) t v" .
have "rawpsubst (subst (rawpsubst ?\<phi>ux (zip (map Var us) xs)) t u) (zip ts us) =
subst (rawpsubst (rawpsubst ?\<phi>ux (zip (map Var us) xs)) (zip ts us)) t u"
using Cons.prems
by (intro subst_rawpsubst_fresh_switch[simplified,symmetric]) (auto intro!: rawpsubst dest!: set_zip_D)
also have "\<dots> = subst (rawpsubst (rawpsubst ?\<phi>ux (zip (map Var vs) xs)) (zip ts vs)) t u"
unfolding 0 ..
also have "\<dots> = rawpsubst (subst (rawpsubst ?\<phi>ux (zip (map Var vs) xs)) t u) (zip ts vs)"
using Cons.prems
by (intro subst_rawpsubst_fresh_switch[simplified]) (auto intro!: rawpsubst dest!: set_zip_D)
also have "\<dots> = rawpsubst (subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var u) x) t u) (zip ts vs)"
unfolding 1 ..
also have "\<dots> = rawpsubst (subst (subst (rawpsubst \<phi> (zip (map Var vs) xs)) (Var v) x) t v) (zip ts vs)"
unfolding 2 ..
also have "\<dots> = rawpsubst (subst (rawpsubst ?\<phi>vx (zip (map Var vs) xs)) t v) (zip ts vs)"
unfolding 11 ..
finally have "rawpsubst (subst (rawpsubst ?\<phi>ux (zip (map Var us) xs)) t u) (zip ts us) =
rawpsubst (subst (rawpsubst ?\<phi>vx (zip (map Var vs) xs)) t v) (zip ts vs)" .
thus ?case by simp
qed auto
text \<open>... now getting rid of the disjointness hypothesis:\<close>
lemma rawpsubst_compose_freshVar2:
assumes \<phi>[simp]: "\<phi> \<in> fmla"
and ts: "set ts \<subseteq> trm"
and xs: "set xs \<subseteq> var" "distinct xs"
and us_facts: "set us \<subseteq> var" "distinct us"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set us \<inter> set xs = {}"
and vs_facts: "set vs \<subseteq> var" "distinct vs"
"set vs \<inter> Fvars \<phi> = {}"
"set vs \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set vs \<inter> set xs = {}"
and l: "length us = length xs" "length vs = length xs" "length ts = length xs"
shows "rawpsubst (rawpsubst \<phi> (zip (map Var us) xs)) (zip ts us) =
rawpsubst (rawpsubst \<phi> (zip (map Var vs) xs)) (zip ts vs)" (is "?L = ?R")
proof-
define ws where "ws = getFrN (xs @ us @ vs) ts [\<phi>] (length xs)"
note fv = getFrN_Fvars[of "xs @ us @ vs" "ts" "[\<phi>]" _ "length xs"]
and fvt = getFrN_FvarsT[of "xs @ us @ vs" "ts" "[\<phi>]" _ "length xs"]
and var = getFrN_var[of "xs @ us @ vs" "ts" "[\<phi>]" _ "length xs"]
and l = getFrN_length[of "xs @ us @ vs" "ts" "[\<phi>]" "length xs"]
have ws_facts: "set ws \<subseteq> var" "distinct ws"
"set ws \<inter> Fvars \<phi> = {}"
"set ws \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set ws \<inter> set xs = {}" "set ws \<inter> set us = {}" "set ws \<inter> set vs = {}"
"length ws = length xs" using assms unfolding ws_def
apply -
subgoal by auto
subgoal by auto
subgoal using fv by auto
subgoal using fvt IntI empty_iff by fastforce
subgoal using var IntI empty_iff by fastforce
subgoal using var IntI empty_iff by fastforce
subgoal using var IntI empty_iff by fastforce
subgoal using l by auto .
have "?L = rawpsubst (rawpsubst \<phi> (zip (map Var ws) xs)) (zip ts ws)"
apply(rule rawpsubst_compose_freshVar2_aux) using assms ws_facts by auto
also have "\<dots> = ?R"
apply(rule rawpsubst_compose_freshVar2_aux) using assms ws_facts by auto
finally show ?thesis .
qed
lemma psubst_subst_fresh_switch:
assumes "\<phi> \<in> fmla" "snd ` set txs \<subseteq> var" "fst ` set txs \<subseteq> trm"
and "\<forall>x\<in>snd ` set txs. x \<notin> FvarsT s" "\<forall>t\<in>fst ` set txs. y \<notin> FvarsT t"
and "distinct (map snd txs)"
and "s \<in> trm" "y \<in> var" "y \<notin> snd ` set txs"
shows "psubst (subst \<phi> s y) txs = subst (psubst \<phi> txs) s y"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
note fvt = getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and fv = getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and var = getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and l = getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us apply -
subgoal by auto
subgoal using fv by (cases txs, auto)
subgoal using fvt by (cases txs, auto)
subgoal using var by (cases txs, auto)
subgoal using l by auto
subgoal by auto .
define vs where vs: "vs = getFrN (map snd txs) (map fst txs) [subst \<phi> s y] (length txs)"
note fvt = getFrN_FvarsT[of "map snd txs" "map fst txs" "[subst \<phi> s y]" _ "length txs"]
and fv = getFrN_Fvars[of "map snd txs" "map fst txs" "[subst \<phi> s y]" _ "length txs"]
and var = getFrN_var[of "map snd txs" "map fst txs" "[subst \<phi> s y]" _ "length txs"]
and l = getFrN_length[of "map snd txs" "map fst txs" "[subst \<phi> s y]" "length txs"]
have vs_facts: "set vs \<subseteq> var"
"set vs \<inter> Fvars (subst \<phi> s y) = {}"
"set vs \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs \<inter> snd ` (set txs) = {}"
"length vs = length txs"
"distinct vs"
using assms unfolding vs apply -
subgoal by auto
subgoal using fv by (cases txs, auto)
subgoal using fvt by (cases txs, auto)
subgoal using var by (cases txs, auto)
subgoal using l by auto
subgoal by auto .
define ws where ws: "ws = getFrN (y # map snd txs) (s # map fst txs) [\<phi>] (length txs)"
note fvt = getFrN_FvarsT[of "y # map snd txs" "s # map fst txs" "[\<phi>]" _ "length txs"]
and fv = getFrN_Fvars[of "y # map snd txs" "s # map fst txs" "[\<phi>]" _ "length txs"]
and var = getFrN_var[of "y # map snd txs" "s # map fst txs" "[\<phi>]" _ "length txs"]
and l = getFrN_length[of "y # map snd txs" "s # map fst txs" "[\<phi>]" "length txs"]
have ws_facts: "set ws \<subseteq> var"
"set ws \<inter> Fvars \<phi> = {}" "y \<notin> set ws" "set ws \<inter> FvarsT s = {}"
"set ws \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set ws \<inter> snd ` (set txs) = {}"
"length ws = length txs"
"distinct ws"
using assms unfolding ws apply -
subgoal by auto
subgoal using fv by (cases txs, auto)
subgoal using var by (cases txs, auto)
subgoal using fvt by (cases txs, auto)
subgoal using fvt by (cases txs, auto)
subgoal using var by (cases txs, auto)
subgoal using l by (cases txs, auto)
by auto
let ?vxs = "zip (map Var vs) (map snd txs)"
let ?tvs = "(zip (map fst txs) vs)"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tus = "(zip (map fst txs) us)"
let ?wxs = "zip (map Var ws) (map snd txs)"
let ?tws = "(zip (map fst txs) ws)"
have 0: "rawpsubst (subst \<phi> s y) ?wxs = subst (rawpsubst \<phi> ?wxs) s y"
apply(subst rawpsubst_compose[of \<phi> ?wxs "[(s,y)]",simplified])
using assms ws_facts apply -
subgoal by auto
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by auto
subgoal by auto
subgoal apply(subst rawpsubst_subst_fresh_switch)
by (auto dest!: set_zip_D simp: subset_eq rawpsubst_subst_fresh_switch) .
have 1: "rawpsubst (rawpsubst \<phi> ?wxs) ?tws = rawpsubst (rawpsubst \<phi> ?uxs) ?tus"
using assms ws_facts us_facts by (intro rawpsubst_compose_freshVar2) (auto simp: subset_eq)
have "rawpsubst (rawpsubst (subst \<phi> s y) ?vxs) ?tvs =
rawpsubst (rawpsubst (subst \<phi> s y) ?wxs) ?tws"
using assms ws_facts vs_facts
by (intro rawpsubst_compose_freshVar2) (auto simp: subset_eq)
also have "\<dots> = rawpsubst (subst (rawpsubst \<phi> ?wxs) s y) ?tws" unfolding 0 ..
also have "\<dots> = subst (rawpsubst (rawpsubst \<phi> ?wxs) ?tws) s y"
apply(subst rawpsubst_compose[of "rawpsubst \<phi> ?wxs" ?tws "[(s,y)]",simplified])
using assms ws_facts apply -
subgoal by (auto dest!: set_zip_D simp: subset_eq intro!: rawpsubst)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by auto
subgoal by auto
subgoal by (subst rawpsubst_subst_fresh_switch)
(auto dest!: set_zip_D simp: subset_eq rawpsubst_subst_fresh_switch
intro!: rawpsubst) .
also have "\<dots> = subst (rawpsubst (rawpsubst \<phi> ?uxs) ?tus) s y" unfolding 1 ..
finally show ?thesis unfolding psubst_def by (simp add: Let_def vs[symmetric] us[symmetric])
qed
text \<open>For many cases, the simpler rawpsubst can replace psubst:\<close>
lemma psubst_eq_rawpsubst:
assumes "\<phi> \<in> fmla" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
(* ... namely, when the substituted variables do not belong to trms substituted for previous variables: *)
and "\<And> i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> snd (txs!j) \<notin> FvarsT (fst (txs!i))"
shows "psubst \<phi> txs = rawpsubst \<phi> txs"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
note fvt = getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and fv = getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and var = getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
and l = getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
apply -
subgoal by auto
subgoal using fv by auto
subgoal using fvt by force
subgoal using var by (force simp: image_iff)
using l by auto
show ?thesis
using rawpsubst_compose_freshVar assms us_facts
by (simp add: psubst_def Let_def us[symmetric])
qed
text \<open>Some particular cases:\<close>
lemma psubst_eq_subst:
assumes "\<phi> \<in> fmla" "x \<in> var" and "t \<in> trm"
shows "psubst \<phi> [(t,x)] = subst \<phi> t x"
proof-
have "psubst \<phi> [(t,x)] = rawpsubst \<phi> [(t,x)]" apply(rule psubst_eq_rawpsubst)
using assms by auto
thus ?thesis by auto
qed
lemma psubst_eq_rawpsubst2:
assumes "\<phi> \<in> fmla" "x1 \<in> var" "x2 \<in> var" "t1 \<in> trm" "t2 \<in> trm"
and "x1 \<noteq> x2" "x2 \<notin> FvarsT t1"
shows "psubst \<phi> [(t1,x1),(t2,x2)] = rawpsubst \<phi> [(t1,x1),(t2,x2)]"
apply(rule psubst_eq_rawpsubst)
using assms using less_SucE by force+
lemma psubst_eq_rawpsubst3:
assumes "\<phi> \<in> fmla" "x1 \<in> var" "x2 \<in> var" "x3 \<in> var" "t1 \<in> trm" "t2 \<in> trm" "t3 \<in> trm"
and "x1 \<noteq> x2" "x1 \<noteq> x3" "x2 \<noteq> x3"
"x2 \<notin> FvarsT t1" "x3 \<notin> FvarsT t1" "x3 \<notin> FvarsT t2"
shows "psubst \<phi> [(t1,x1),(t2,x2),(t3,x3)] = rawpsubst \<phi> [(t1,x1),(t2,x2),(t3,x3)]"
using assms using less_SucE apply(intro psubst_eq_rawpsubst)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal for i j
apply(cases j)
subgoal by auto
subgoal by (simp add: nth_Cons') . .
lemma psubst_eq_rawpsubst4:
assumes "\<phi> \<in> fmla" "x1 \<in> var" "x2 \<in> var" "x3 \<in> var" "x4 \<in> var"
"t1 \<in> trm" "t2 \<in> trm" "t3 \<in> trm" "t4 \<in> trm"
and "x1 \<noteq> x2" "x1 \<noteq> x3" "x2 \<noteq> x3" "x1 \<noteq> x4" "x2 \<noteq> x4" "x3 \<noteq> x4"
"x2 \<notin> FvarsT t1" "x3 \<notin> FvarsT t1" "x3 \<notin> FvarsT t2" "x4 \<notin> FvarsT t1" "x4 \<notin> FvarsT t2" "x4 \<notin> FvarsT t3"
shows "psubst \<phi> [(t1,x1),(t2,x2),(t3,x3),(t4,x4)] = rawpsubst \<phi> [(t1,x1),(t2,x2),(t3,x3),(t4,x4)]"
using assms using less_SucE apply(intro psubst_eq_rawpsubst)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal for i j
apply(cases j)
subgoal by auto
subgoal by (simp add: nth_Cons') . .
lemma rawpsubst_same_Var[simp]:
assumes "\<phi> \<in> fmla" "set xs \<subseteq> var"
shows "rawpsubst \<phi> (map (\<lambda>x. (Var x,x)) xs) = \<phi>"
using assms by (induct xs) auto
lemma psubst_same_Var[simp]:
assumes "\<phi> \<in> fmla" "set xs \<subseteq> var" and "distinct xs"
shows "psubst \<phi> (map (\<lambda>x. (Var x,x)) xs) = \<phi>"
proof-
have "psubst \<phi> (map (\<lambda>x. (Var x,x)) xs) = rawpsubst \<phi> (map (\<lambda>x. (Var x,x)) xs)"
using assms by (intro psubst_eq_rawpsubst) (auto simp: nth_eq_iff_index_eq subsetD)
thus ?thesis using assms by auto
qed
lemma rawpsubst_notIn[simp]:
assumes "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "\<phi> \<in> fmla"
and "Fvars \<phi> \<inter> snd ` (set txs) = {}"
shows "rawpsubst \<phi> txs = \<phi>"
using assms by (induct txs) auto
lemma psubst_notIn[simp]:
assumes "x \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "\<phi> \<in> fmla"
and "Fvars \<phi> \<inter> snd ` (set txs) = {}"
shows "psubst \<phi> txs = \<phi>"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi> = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
using assms unfolding us apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by (fastforce simp: image_iff)
subgoal by auto .
(* *)
have [simp]: "rawpsubst \<phi> (zip (map Var us) (map snd txs)) = \<phi>"
using assms us_facts apply(intro rawpsubst_notIn)
subgoal by (auto dest!: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by auto
subgoal by (auto dest!: set_zip_rightD) .
have [simp]: "rawpsubst \<phi> (zip (map fst txs) us) = \<phi>"
using assms us_facts apply(intro rawpsubst_notIn)
subgoal by (auto dest!: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by auto
subgoal by (auto dest!: set_zip_rightD) .
show ?thesis using assms us_facts unfolding psubst_def
by(auto simp: Let_def us[symmetric])
qed
end \<comment> \<open>context @{locale Generic_Syntax}\<close>
section \<open>Adding Numerals to the Generic Syntax\<close>
locale Syntax_with_Numerals =
Generic_Syntax var trm fmla Var FvarsT substT Fvars subst
for var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
+
fixes
\<comment> \<open>Abstract notion of numerals, as a subset of the ground terms:\<close>
num :: "'trm set"
assumes
numNE: "num \<noteq> {}"
and
num: "num \<subseteq> trm"
and
FvarsT_num[simp, intro!]: "\<And>n. n \<in> num \<Longrightarrow> FvarsT n = {}"
begin
lemma substT_num1[simp]: "t \<in> trm \<Longrightarrow> y \<in> var \<Longrightarrow> n \<in> num \<Longrightarrow> substT n t y = n"
using num by auto
lemma in_num[simp]: "n \<in> num \<Longrightarrow> n \<in> trm" using num by auto
lemma subst_comp_num:
assumes "\<phi> \<in> fmla" "x \<in> var" "y \<in> var" "n \<in> num"
shows "x \<noteq> y \<Longrightarrow> subst (subst \<phi> (Var x) y) n x = subst (subst \<phi> n x) n y"
using assms by (simp add: subst_comp)
lemma rawpsubstT_num:
assumes "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "n \<in> num"
shows "rawpsubstT n txs = n"
using assms by (induct txs) auto
lemma psubstT_num[simp]:
assumes "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm" "n \<in> num"
shows "psubstT n txs = n"
proof-
define us where us: "us = getFrN (map snd txs) (n # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT n = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "n # map fst txs" "[]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "n # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "n # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "n # map fst txs" "[]" "length txs"]
by (auto 7 0 simp: set_eq_iff image_iff)
let ?t = "rawpsubstT n (zip (map Var us) (map snd txs))"
have t: "?t = n"
using assms us_facts apply(intro rawpsubstT_num)
subgoal by (auto dest!: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by auto .
have "rawpsubstT ?t (zip (map fst txs) us) = n"
unfolding t using assms us_facts apply(intro rawpsubstT_num)
subgoal by (auto dest!: set_zip_rightD)
subgoal by (auto dest!: set_zip_leftD)
subgoal by auto .
thus ?thesis unfolding psubstT_def by(simp add: Let_def us[symmetric])
qed
end \<comment> \<open>context @{locale Syntax_with_Numerals}\<close>
section \<open>Adding Connectives and Quantifiers\<close>
locale Syntax_with_Connectives =
Generic_Syntax var trm fmla Var FvarsT substT Fvars subst
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
+
fixes
\<comment> \<open>Logical connectives\<close>
eql :: "'trm \<Rightarrow> 'trm \<Rightarrow> 'fmla"
and
cnj :: "'fmla \<Rightarrow> 'fmla \<Rightarrow> 'fmla"
and
imp :: "'fmla \<Rightarrow> 'fmla \<Rightarrow> 'fmla"
and
all :: "'var \<Rightarrow> 'fmla \<Rightarrow> 'fmla"
and
exi :: "'var \<Rightarrow> 'fmla \<Rightarrow> 'fmla"
assumes
eql[simp,intro]: "\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> eql t1 t2 \<in> fmla"
and
cnj[simp,intro]: "\<And> \<phi>1 \<phi>2. \<phi>1 \<in> fmla \<Longrightarrow> \<phi>2 \<in> fmla \<Longrightarrow> cnj \<phi>1 \<phi>2 \<in> fmla"
and
imp[simp,intro]: "\<And> \<phi>1 \<phi>2. \<phi>1 \<in> fmla \<Longrightarrow> \<phi>2 \<in> fmla \<Longrightarrow> imp \<phi>1 \<phi>2 \<in> fmla"
and
all[simp,intro]: "\<And> x \<phi>. x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> all x \<phi> \<in> fmla"
and
exi[simp,intro]: "\<And> x \<phi>. x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> exi x \<phi> \<in> fmla"
and
Fvars_eql[simp]:
"\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> Fvars (eql t1 t2) = FvarsT t1 \<union> FvarsT t2"
and
Fvars_cnj[simp]:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> Fvars (cnj \<phi> \<chi>) = Fvars \<phi> \<union> Fvars \<chi>"
and
Fvars_imp[simp]:
"\<And> \<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> Fvars (imp \<phi> \<chi>) = Fvars \<phi> \<union> Fvars \<chi>"
and
Fvars_all[simp]:
"\<And> x \<phi>. x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars (all x \<phi>) = Fvars \<phi> - {x}"
and
Fvars_exi[simp]:
"\<And> x \<phi>. x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars (exi x \<phi>) = Fvars \<phi> - {x}"
and
\<comment> \<open>Assumed properties of substitution\<close>
subst_cnj[simp]:
"\<And> x \<phi> \<chi> t. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (cnj \<phi> \<chi>) t x = cnj (subst \<phi> t x) (subst \<chi> t x)"
and
subst_imp[simp]:
"\<And> x \<phi> \<chi> t. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (imp \<phi> \<chi>) t x = imp (subst \<phi> t x) (subst \<chi> t x)"
and
subst_all[simp]:
"\<And> x y \<phi> t. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow>
x \<noteq> y \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow> subst (all x \<phi>) t y = all x (subst \<phi> t y)"
and
subst_exi[simp]:
"\<And> x y \<phi> t. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow>
x \<noteq> y \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow> subst (exi x \<phi>) t y = exi x (subst \<phi> t y)"
and
subst_eql[simp]:
"\<And> t1 t2 t x. t \<in> trm \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (eql t1 t2) t x = eql (substT t1 t x) (substT t2 t x)"
begin
(*
(* "is the unique": isTU t x \<phi> is the formula "t is the unique x such that phi(x)"
*)
definition isTU :: "'trm \<Rightarrow> 'var \<Rightarrow> 'fmla \<Rightarrow> 'fmla" where
"isTU t x \<phi> \<equiv>
cnj (subst \<phi> t x)
(all x (imp \<phi> (subst \<phi> t x)))"
(* TODO: properties: works well when x is not free in t,
in particular, when t is num n *)
*)
text \<open>Formula equivalence, $\longleftrightarrow$, a derived connective\<close>
definition eqv :: "'fmla \<Rightarrow> 'fmla \<Rightarrow> 'fmla" where
"eqv \<phi> \<chi> = cnj (imp \<phi> \<chi>) (imp \<chi> \<phi>)"
lemma
eqv[simp]: "\<And>\<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> eqv \<phi> \<chi> \<in> fmla"
and
Fvars_eqv[simp]: "\<And>\<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
Fvars (eqv \<phi> \<chi>) = Fvars \<phi> \<union> Fvars \<chi>"
and
subst_eqv[simp]:
"\<And>\<phi> \<chi> t x. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (eqv \<phi> \<chi>) t x = eqv (subst \<phi> t x) (subst \<chi> t x)"
unfolding eqv_def by auto
lemma subst_all_idle[simp]:
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla" "t \<in> trm"
shows "subst (all x \<phi>) t x = all x \<phi>"
by (intro subst_notIn) auto
lemma subst_exi_idle[simp]:
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla" "t \<in> trm"
shows "subst (exi x \<phi>) t x = exi x \<phi>"
by (rule subst_notIn) auto
text \<open>Parallel substitution versus connectives and quantifiers.\<close>
lemma rawpsubst_cnj:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst (cnj \<phi>1 \<phi>2) txs = cnj (rawpsubst \<phi>1 txs) (rawpsubst \<phi>2 txs)"
using assms by (induct txs arbitrary: \<phi>1 \<phi>2) auto
lemma psubst_cnj[simp]:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst (cnj \<phi>1 \<phi>2) txs = cnj (psubst \<phi>1 txs) (psubst \<phi>2 txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [cnj \<phi>1 \<phi>2] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi>1 = {}"
"set us \<inter> Fvars \<phi>2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[cnj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[cnj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[cnj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[cnj \<phi>1 \<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by auto
subgoal by auto .
define vs1 where vs1: "vs1 = getFrN (map snd txs) (map fst txs) [\<phi>1] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> Fvars \<phi>1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms unfolding vs1
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>1]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
subgoal by auto
subgoal by auto .
define vs2 where vs2: "vs2 = getFrN (map snd txs) (map fst txs) [\<phi>2] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> Fvars \<phi>2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms unfolding vs2
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
subgoal by auto
subgoal by auto .
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs1 = "zip (map fst txs) vs1"
let ?vxs1 = "zip (map Var vs1) (map snd txs)"
let ?tvs2 = "zip (map fst txs) vs2"
let ?vxs2 = "zip (map Var vs2) (map snd txs)"
let ?c = "rawpsubst (cnj \<phi>1 \<phi>2) ?uxs"
have c: "?c = cnj (rawpsubst \<phi>1 ?uxs) (rawpsubst \<phi>2 ?uxs)"
using assms us_facts
by (intro rawpsubst_cnj) (auto intro!: rawpsubstT dest!: set_zip_D)
have 0: "rawpsubst ?c ?tus =
cnj (rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus) (rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus)"
unfolding c using assms us_facts
by (intro rawpsubst_cnj) (auto dest!: set_zip_D intro!: rawpsubst)
have 1: "rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>1 ?vxs1) ?tvs1"
using assms vs1_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
have 2: "rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>2 ?vxs2) ?tvs2"
using assms vs2_facts us_facts
by (intro rawpsubst_compose_freshVar2)(auto intro!: rawpsubst)
show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 0 1 2)
qed
lemma rawpsubst_imp:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst (imp \<phi>1 \<phi>2) txs = imp (rawpsubst \<phi>1 txs) (rawpsubst \<phi>2 txs)"
using assms apply (induct txs arbitrary: \<phi>1 \<phi>2)
subgoal by auto
subgoal for tx txs \<phi>1 \<phi>2 by (cases tx) auto .
lemma psubst_imp[simp]:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst (imp \<phi>1 \<phi>2) txs = imp (psubst \<phi>1 txs) (psubst \<phi>2 txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [imp \<phi>1 \<phi>2] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi>1 = {}"
"set us \<inter> Fvars \<phi>2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[imp \<phi>1 \<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[imp \<phi>1 \<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[imp \<phi>1 \<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[imp \<phi>1 \<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define vs1 where vs1: "vs1 = getFrN (map snd txs) (map fst txs) [\<phi>1] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> Fvars \<phi>1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms unfolding vs1
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>1]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
by auto
define vs2 where vs2: "vs2 = getFrN (map snd txs) (map fst txs) [\<phi>2] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> Fvars \<phi>2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms unfolding vs2
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
by auto
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs1 = "zip (map fst txs) vs1"
let ?vxs1 = "zip (map Var vs1) (map snd txs)"
let ?tvs2 = "zip (map fst txs) vs2"
let ?vxs2 = "zip (map Var vs2) (map snd txs)"
let ?c = "rawpsubst (imp \<phi>1 \<phi>2) ?uxs"
have c: "?c = imp (rawpsubst \<phi>1 ?uxs) (rawpsubst \<phi>2 ?uxs)"
apply(rule rawpsubst_imp) using assms us_facts apply (auto intro!: rawpsubstT)
apply(drule set_zip_rightD) apply simp apply blast
apply(drule set_zip_leftD) apply simp apply blast .
have 0: "rawpsubst ?c ?tus =
imp (rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus) (rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus)"
unfolding c
using assms us_facts
by (intro rawpsubst_imp) (auto intro!: rawpsubst dest!: set_zip_D)
have 1: "rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>1 ?vxs1) ?tvs1"
using assms vs1_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
have 2: "rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>2 ?vxs2) ?tvs2"
using assms vs2_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 0 1 2)
qed
lemma rawpsubst_eqv:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst (eqv \<phi>1 \<phi>2) txs = eqv (rawpsubst \<phi>1 txs) (rawpsubst \<phi>2 txs)"
using assms apply (induct txs arbitrary: \<phi>1 \<phi>2)
subgoal by auto
subgoal for tx txs \<phi>1 \<phi>2 by (cases tx) auto .
lemma psubst_eqv[simp]:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst (eqv \<phi>1 \<phi>2) txs = eqv (psubst \<phi>1 txs) (psubst \<phi>2 txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [eqv \<phi>1 \<phi>2] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi>1 = {}"
"set us \<inter> Fvars \<phi>2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[eqv \<phi>1 \<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[eqv \<phi>1 \<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[eqv \<phi>1 \<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[eqv \<phi>1 \<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define vs1 where vs1: "vs1 = getFrN (map snd txs) (map fst txs) [\<phi>1] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> Fvars \<phi>1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms unfolding vs1
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>1]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
by auto
define vs2 where vs2: "vs2 = getFrN (map snd txs) (map fst txs) [\<phi>2] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> Fvars \<phi>2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms unfolding vs2
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by force
by auto
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs1 = "zip (map fst txs) vs1"
let ?vxs1 = "zip (map Var vs1) (map snd txs)"
let ?tvs2 = "zip (map fst txs) vs2"
let ?vxs2 = "zip (map Var vs2) (map snd txs)"
let ?c = "rawpsubst (eqv \<phi>1 \<phi>2) ?uxs"
have c: "?c = eqv (rawpsubst \<phi>1 ?uxs) (rawpsubst \<phi>2 ?uxs)"
using assms us_facts
by (intro rawpsubst_eqv) (auto intro!: rawpsubstT dest!: set_zip_D)
have 0: "rawpsubst ?c ?tus =
eqv (rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus) (rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus)"
unfolding c using assms us_facts
by (intro rawpsubst_eqv) (auto intro!: rawpsubst dest!: set_zip_D)
have 1: "rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>1 ?vxs1) ?tvs1"
using assms vs1_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
have 2: "rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>2 ?vxs2) ?tvs2"
using assms vs2_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 0 1 2)
qed
lemma rawpsubst_all:
assumes "\<phi> \<in> fmla" "y \<in> var"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "y \<notin> snd ` (set txs)" "y \<notin> \<Union> (FvarsT ` fst ` (set txs))"
shows "rawpsubst (all y \<phi>) txs = all y (rawpsubst \<phi> txs)"
using assms apply (induct txs arbitrary: \<phi>)
subgoal by auto
subgoal for tx txs \<phi> by (cases tx) auto .
lemma psubst_all[simp]:
assumes "\<phi> \<in> fmla" "y \<in> var"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "y \<notin> snd ` (set txs)" "y \<notin> \<Union> (FvarsT ` fst ` (set txs))"
and "distinct (map snd txs)"
shows "psubst (all y \<phi>) txs = all y (psubst \<phi> txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [all y \<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> (Fvars \<phi> - {y}) = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[all y \<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[all y \<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[all y \<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[all y \<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define vs where vs: "vs = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have vs_facts: "set vs \<subseteq> var"
"set vs \<inter> Fvars \<phi> = {}"
"set vs \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs \<inter> snd ` (set txs) = {}"
"length vs = length txs"
"distinct vs"
using assms unfolding vs
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define ws where ws: "ws = getFrN (y # map snd txs) (map fst txs) [\<phi>] (length txs)"
have ws_facts: "set ws \<subseteq> var"
"set ws \<inter> Fvars \<phi> = {}" "y \<notin> set ws"
"set ws \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set ws \<inter> snd ` (set txs) = {}"
"length ws = length txs"
"distinct ws"
using assms unfolding ws
using getFrN_Fvars[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "y # map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
have 0: "rawpsubst (all y \<phi>) (zip (map Var ws) (map snd txs)) =
all y (rawpsubst \<phi> (zip (map Var ws) (map snd txs)))"
using assms ws_facts apply(intro rawpsubst_all)
subgoal by auto
subgoal by auto
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest: set_zip_D) .
have 1: "rawpsubst ((rawpsubst \<phi> (zip (map Var ws) (map snd txs)))) (zip (map fst txs) ws) =
rawpsubst ((rawpsubst \<phi> (zip (map Var vs) (map snd txs)))) (zip (map fst txs) vs)"
apply(rule rawpsubst_compose_freshVar2)
using assms ws_facts vs_facts by (auto intro!: rawpsubst)
have "rawpsubst (rawpsubst (all y \<phi>) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubst (rawpsubst (all y \<phi>) (zip (map Var ws) (map snd txs))) (zip (map fst txs) ws)"
using assms ws_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
also have
"\<dots> = all y (rawpsubst ((rawpsubst \<phi> (zip (map Var ws) (map snd txs)))) (zip (map fst txs) ws))"
unfolding 0 using assms ws_facts
by (intro rawpsubst_all) (auto dest!: set_zip_D intro!: rawpsubst)
also have
"\<dots> = all y (rawpsubst (rawpsubst \<phi> (zip (map Var vs) (map snd txs))) (zip (map fst txs) vs))"
unfolding 1 ..
finally show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs[symmetric])
qed
lemma rawpsubst_exi:
assumes "\<phi> \<in> fmla" "y \<in> var"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "y \<notin> snd ` (set txs)" "y \<notin> \<Union> (FvarsT ` fst ` (set txs))"
shows "rawpsubst (exi y \<phi>) txs = exi y (rawpsubst \<phi> txs)"
using assms apply (induct txs arbitrary: \<phi>)
subgoal by auto
subgoal for tx txs \<phi> by (cases tx) auto .
lemma psubst_exi[simp]:
assumes "\<phi> \<in> fmla" "y \<in> var"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "y \<notin> snd ` (set txs)" "y \<notin> \<Union> (FvarsT ` fst ` (set txs))"
and "distinct (map snd txs)"
shows "psubst (exi y \<phi>) txs = exi y (psubst \<phi> txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [exi y \<phi>] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> (Fvars \<phi> - {y}) = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[exi y \<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[exi y \<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[exi y \<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[exi y \<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
define vs where vs: "vs = getFrN (map snd txs) (map fst txs) [\<phi>] (length txs)"
have vs_facts: "set vs \<subseteq> var"
"set vs \<inter> Fvars \<phi> = {}"
"set vs \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs \<inter> snd ` (set txs) = {}"
"length vs = length txs"
"distinct vs"
using assms unfolding vs
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
define ws where ws: "ws = getFrN (y # map snd txs) (map fst txs) [\<phi>] (length txs)"
have ws_facts: "set ws \<subseteq> var"
"set ws \<inter> Fvars \<phi> = {}" "y \<notin> set ws"
"set ws \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set ws \<inter> snd ` (set txs) = {}"
"length ws = length txs"
"distinct ws"
using assms unfolding ws
using getFrN_Fvars[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_FvarsT[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_var[of "y # map snd txs" "map fst txs" "[\<phi>]" _ "length txs"]
getFrN_length[of "y # map snd txs" "map fst txs" "[\<phi>]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
have 0: "rawpsubst (exi y \<phi>) (zip (map Var ws) (map snd txs)) =
exi y (rawpsubst \<phi> (zip (map Var ws) (map snd txs)))"
using assms ws_facts apply(intro rawpsubst_exi)
subgoal by auto
subgoal by auto
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest: set_zip_D) .
have 1: "rawpsubst ((rawpsubst \<phi> (zip (map Var ws) (map snd txs)))) (zip (map fst txs) ws) =
rawpsubst ((rawpsubst \<phi> (zip (map Var vs) (map snd txs)))) (zip (map fst txs) vs)"
using assms ws_facts vs_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
have "rawpsubst (rawpsubst (exi y \<phi>) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubst (rawpsubst (exi y \<phi>) (zip (map Var ws) (map snd txs))) (zip (map fst txs) ws)"
using assms ws_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
also have
"\<dots> = exi y (rawpsubst ((rawpsubst \<phi> (zip (map Var ws) (map snd txs)))) (zip (map fst txs) ws))"
using assms ws_facts unfolding 0
by (intro rawpsubst_exi) (auto dest!: set_zip_D intro!: rawpsubst)
also have
"\<dots> = exi y (rawpsubst (rawpsubst \<phi> (zip (map Var vs) (map snd txs))) (zip (map fst txs) vs))"
unfolding 1 ..
finally show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs[symmetric])
qed
end \<comment> \<open>context @{locale Syntax_with_Connectives}\<close>
locale Syntax_with_Numerals_and_Connectives =
Syntax_with_Numerals
var trm fmla Var FvarsT substT Fvars subst
num
+
Syntax_with_Connectives
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and num
and eql cnj imp all exi
begin
lemma subst_all_num[simp]:
assumes "\<phi> \<in> fmla" "x \<in> var" "y \<in> var" "n \<in> num"
shows "x \<noteq> y \<Longrightarrow> subst (all x \<phi>) n y = all x (subst \<phi> n y)"
using assms by simp
lemma subst_exi_num[simp]:
assumes "\<phi> \<in> fmla" "x \<in> var" "y \<in> var" "n \<in> num"
shows "x \<noteq> y \<Longrightarrow> subst (exi x \<phi>) n y = exi x (subst \<phi> n y)"
using assms by simp
text \<open>The "soft substitution" function:\<close>
definition softSubst :: "'fmla \<Rightarrow> 'trm \<Rightarrow> 'var \<Rightarrow> 'fmla" where
"softSubst \<phi> t x = exi x (cnj (eql (Var x) t) \<phi>)"
lemma softSubst[simp,intro]: "\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> softSubst \<phi> t x \<in> fmla"
unfolding softSubst_def by auto
lemma Fvars_softSubst[simp]:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
Fvars (softSubst \<phi> t x) = (Fvars \<phi> \<union> FvarsT t - {x})"
unfolding softSubst_def by auto
lemma Fvars_softSubst_subst_in:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow> x \<in> Fvars \<phi> \<Longrightarrow>
Fvars (softSubst \<phi> t x) = Fvars (subst \<phi> t x)"
by auto
lemma Fvars_softSubst_subst_notIn:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> x \<notin> FvarsT t \<Longrightarrow> x \<notin> Fvars \<phi> \<Longrightarrow>
Fvars (softSubst \<phi> t x) = Fvars (subst \<phi> t x) \<union> FvarsT t"
by auto
end \<comment> \<open>context @{locale Syntax_with_Connectives}\<close>
text \<open>The addition of False among logical connectives\<close>
locale Syntax_with_Connectives_False =
Syntax_with_Connectives
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
+
fixes fls::'fmla
assumes
fls[simp,intro!]: "fls \<in> fmla"
and
Fvars_fls[simp,intro!]: "Fvars fls = {}"
and
subst_fls[simp]:
"\<And>t x. t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> subst fls t x = fls"
begin
text \<open>Negation as a derrived connective:\<close>
definition neg :: "'fmla \<Rightarrow> 'fmla" where
"neg \<phi> = imp \<phi> fls"
lemma
neg[simp]: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> neg \<phi> \<in> fmla"
and
Fvars_neg[simp]: "\<And>\<phi>. \<phi> \<in> fmla \<Longrightarrow> Fvars (neg \<phi>) = Fvars \<phi>"
and
subst_neg[simp]:
"\<And>\<phi> t x. \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (neg \<phi>) t x = neg (subst \<phi> t x)"
unfolding neg_def by auto
text \<open>True as a derived connective:\<close>
definition tru where "tru = neg fls"
lemma
tru[simp,intro!]: "tru \<in> fmla"
and
Fvars_tru[simp]: "Fvars tru = {}"
and
subst_tru[simp]: "\<And> t x. t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> subst tru t x = tru"
unfolding tru_def by auto
subsection \<open>Iterated conjunction\<close>
text \<open>First we define list-based conjunction:\<close>
fun lcnj :: "'fmla list \<Rightarrow> 'fmla" where
"lcnj [] = tru"
| "lcnj (\<phi> # \<phi>s) = cnj \<phi> (lcnj \<phi>s)"
lemma lcnj[simp,intro!]: "set \<phi>s \<subseteq> fmla \<Longrightarrow> lcnj \<phi>s \<in> fmla"
by (induct \<phi>s) auto
lemma Fvars_lcnj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> Fvars (lcnj \<phi>s) = \<Union> (set (map Fvars \<phi>s))"
by(induct \<phi>s) auto
lemma subst_lcnj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (lcnj \<phi>s) t x = lcnj (map (\<lambda>\<phi>. subst \<phi> t x) \<phi>s)"
by(induct \<phi>s) auto
text \<open>Then we define (finite-)set-based conjunction:\<close>
definition scnj :: "'fmla set \<Rightarrow> 'fmla" where
"scnj F = lcnj (asList F)"
lemma scnj[simp,intro!]: "F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> scnj F \<in> fmla"
unfolding scnj_def by auto
lemma Fvars_scnj[simp]:
"F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow>Fvars (scnj F) = \<Union> (Fvars ` F)"
unfolding scnj_def by auto
subsection \<open>Parallel substitution versus the new connectives\<close>
lemma rawpsubst_fls:
"snd ` (set txs) \<subseteq> var \<Longrightarrow> fst ` (set txs) \<subseteq> trm \<Longrightarrow> rawpsubst fls txs = fls"
by (induct txs) auto
lemma psubst_fls[simp]:
assumes "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "psubst fls txs = fls"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [fls] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[fls]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[fls]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[fls]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[fls]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
have [simp]: "rawpsubst fls (zip (map Var us) (map snd txs)) = fls"
using us_facts assms by (intro rawpsubst_fls) (auto dest!: set_zip_D)
show ?thesis using assms us_facts
unfolding psubst_def by (auto simp add: Let_def us[symmetric] intro!: rawpsubst_fls dest!: set_zip_D)
qed
lemma psubst_neg[simp]:
assumes "\<phi> \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst (neg \<phi>) txs = neg (psubst \<phi> txs)"
unfolding neg_def using assms psubst_imp psubst_fls by auto
lemma psubst_tru[simp]:
assumes "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst tru txs = tru"
unfolding tru_def using assms psubst_neg[of fls txs] psubst_fls by auto
lemma psubst_lcnj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> snd ` (set txs) \<subseteq> var \<Longrightarrow> fst ` (set txs) \<subseteq> trm \<Longrightarrow>
distinct (map snd txs) \<Longrightarrow>
psubst (lcnj \<phi>s) txs = lcnj (map (\<lambda>\<phi>. psubst \<phi> txs) \<phi>s)"
by (induct \<phi>s) auto
end \<comment> \<open>context @{locale Syntax_with_Connectives_False}\<close>
section \<open>Adding Disjunction\<close>
text \<open>NB: In intuitionistic logic, disjunction is not definable from the other connectives.\<close>
locale Syntax_with_Connectives_False_Disj =
Syntax_with_Connectives_False
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
+
fixes dsj :: "'fmla \<Rightarrow> 'fmla \<Rightarrow> 'fmla"
assumes
dsj[simp]: "\<And>\<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> dsj \<phi> \<chi> \<in> fmla"
and
Fvars_dsj[simp]: "\<And>\<phi> \<chi>. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow>
Fvars (dsj \<phi> \<chi>) = Fvars \<phi> \<union> Fvars \<chi>"
and
subst_dsj[simp]:
"\<And> x \<phi> \<chi> t. \<phi> \<in> fmla \<Longrightarrow> \<chi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (dsj \<phi> \<chi>) t x = dsj (subst \<phi> t x) (subst \<chi> t x)"
begin
subsection \<open>Iterated disjunction\<close>
text \<open>First we define list-based disjunction:\<close>
fun ldsj :: "'fmla list \<Rightarrow> 'fmla" where
"ldsj [] = fls"
| "ldsj (\<phi> # \<phi>s) = dsj \<phi> (ldsj \<phi>s)"
lemma ldsj[simp,intro!]: "set \<phi>s \<subseteq> fmla \<Longrightarrow> ldsj \<phi>s \<in> fmla"
by (induct \<phi>s) auto
lemma Fvars_ldsj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> Fvars (ldsj \<phi>s) = \<Union> (set (map Fvars \<phi>s))"
by(induct \<phi>s) auto
lemma subst_ldsj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
subst (ldsj \<phi>s) t x = ldsj (map (\<lambda>\<phi>. subst \<phi> t x) \<phi>s)"
by(induct \<phi>s) auto
text \<open>Then we define (finite-)set-based disjunction:\<close>
definition sdsj :: "'fmla set \<Rightarrow> 'fmla" where
"sdsj F = ldsj (asList F)"
lemma sdsj[simp,intro!]: "F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> sdsj F \<in> fmla"
unfolding sdsj_def by auto
lemma Fvars_sdsj[simp]:
"F \<subseteq> fmla \<Longrightarrow> finite F \<Longrightarrow> Fvars (sdsj F) = \<Union> (Fvars ` F)"
unfolding sdsj_def by auto
subsection \<open>Parallel substitution versus the new connectives\<close>
lemma rawpsubst_dsj:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst (dsj \<phi>1 \<phi>2) txs = dsj (rawpsubst \<phi>1 txs) (rawpsubst \<phi>2 txs)"
using assms apply (induct txs arbitrary: \<phi>1 \<phi>2)
subgoal by auto
subgoal for tx txs \<phi>1 \<phi>2 apply (cases tx) by auto .
lemma psubst_dsj[simp]:
assumes "\<phi>1 \<in> fmla" "\<phi>2 \<in> fmla"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)"
shows "psubst (dsj \<phi>1 \<phi>2) txs = dsj (psubst \<phi>1 txs) (psubst \<phi>2 txs)"
proof-
define us where us: "us = getFrN (map snd txs) (map fst txs) [dsj \<phi>1 \<phi>2] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> Fvars \<phi>1 = {}"
"set us \<inter> Fvars \<phi>2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[dsj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[dsj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[dsj \<phi>1 \<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[dsj \<phi>1 \<phi>2]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
subgoal by (fastforce simp: image_iff)
by auto
define vs1 where vs1: "vs1 = getFrN (map snd txs) (map fst txs) [\<phi>1] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> Fvars \<phi>1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms unfolding vs1
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>1]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>1]" "length txs"]
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define vs2 where vs2: "vs2 = getFrN (map snd txs) (map fst txs) [\<phi>2] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> Fvars \<phi>2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms unfolding vs2
using getFrN_Fvars[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[\<phi>2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[\<phi>2]" "length txs"]
apply -
apply -
subgoal by auto
subgoal by fastforce
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?tvs1 = "zip (map fst txs) vs1"
let ?vxs1 = "zip (map Var vs1) (map snd txs)"
let ?tvs2 = "zip (map fst txs) vs2"
let ?vxs2 = "zip (map Var vs2) (map snd txs)"
let ?c = "rawpsubst (dsj \<phi>1 \<phi>2) ?uxs"
have c: "?c = dsj (rawpsubst \<phi>1 ?uxs) (rawpsubst \<phi>2 ?uxs)"
apply(rule rawpsubst_dsj) using assms us_facts apply (auto intro!: rawpsubstT)
apply(drule set_zip_rightD) apply simp apply blast
apply(drule set_zip_leftD) apply simp apply blast .
have 0: "rawpsubst ?c ?tus =
dsj (rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus) (rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus)"
unfolding c using assms us_facts
by (intro rawpsubst_dsj) (auto intro!: rawpsubst dest!: set_zip_D)
have 1: "rawpsubst (rawpsubst \<phi>1 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>1 ?vxs1) ?tvs1"
using assms vs1_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
have 2: "rawpsubst (rawpsubst \<phi>2 ?uxs) ?tus = rawpsubst (rawpsubst \<phi>2 ?vxs2) ?tvs2"
using assms vs2_facts us_facts
by (intro rawpsubst_compose_freshVar2) (auto intro!: rawpsubst)
show ?thesis unfolding psubst_def by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 0 1 2)
qed
lemma psubst_ldsj[simp]:
"set \<phi>s \<subseteq> fmla \<Longrightarrow> snd ` (set txs) \<subseteq> var \<Longrightarrow> fst ` (set txs) \<subseteq> trm \<Longrightarrow>
distinct (map snd txs) \<Longrightarrow>
psubst (ldsj \<phi>s) txs = ldsj (map (\<lambda>\<phi>. psubst \<phi> txs) \<phi>s)"
by (induct \<phi>s) auto
end \<comment> \<open>context @{locale Syntax_with_Connectives_False_Disj}\<close>
section \<open>Adding an Ordering-Like Formula\<close>
locale Syntax_with_Numerals_and_Connectives_False_Disj =
Syntax_with_Connectives_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
+
Syntax_with_Numerals_and_Connectives
var trm fmla Var FvarsT substT Fvars subst
num
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
text \<open>... and in addition a formula expressing order (think: less than or equal to)\<close>
locale Syntax_PseudoOrder =
Syntax_with_Numerals_and_Connectives_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
+
fixes
\<comment> \<open>Lq is a formula with free variables xx yy:\<close>
Lq :: 'fmla
assumes
Lq[simp,intro!]: "Lq \<in> fmla"
and
Fvars_Lq[simp]: "Fvars Lq = {zz,yy}"
begin
definition LLq where "LLq t1 t2 = psubst Lq [(t1,zz), (t2,yy)]"
lemma LLq_def2: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> yy \<notin> FvarsT t1 \<Longrightarrow>
LLq t1 t2 = subst (subst Lq t1 zz) t2 yy"
unfolding LLq_def by (rule psubst_eq_rawpsubst2[simplified]) auto
lemma LLq[simp,intro]:
assumes "t1 \<in> trm" "t2 \<in> trm"
shows "LLq t1 t2 \<in> fmla"
using assms unfolding LLq_def by auto
lemma LLq2[simp,intro!]:
"n \<in> num \<Longrightarrow> LLq n (Var yy') \<in> fmla"
by auto
lemma Fvars_LLq[simp]: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> yy \<notin> FvarsT t1 \<Longrightarrow>
Fvars (LLq t1 t2) = FvarsT t1 \<union> FvarsT t2"
by (auto simp add: LLq_def2 subst2_fresh_switch)
lemma LLq_simps[simp]:
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (LLq m (Var yy)) n yy = LLq m n"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (LLq m (Var yy')) n yy = LLq m (Var yy')"
"m \<in> num \<Longrightarrow> subst (LLq m (Var yy')) (Var yy) yy' = LLq m (Var yy)"
"n \<in> num \<Longrightarrow> subst (LLq (Var xx) (Var yy)) n xx = LLq n (Var yy)"
"n \<in> num \<Longrightarrow> subst (LLq (Var zz) (Var yy)) n yy = LLq (Var zz) n"
"m \<in> num \<Longrightarrow> subst (LLq (Var zz) (Var yy)) m zz = LLq m (Var yy)"
"m \<in> num \<Longrightarrow> n \<in> num \<Longrightarrow> subst (LLq (Var zz) n) m xx = LLq (Var zz) n"
by (auto simp: LLq_def2 subst2_fresh_switch)
end \<comment> \<open>context @{locale Syntax_PseudoOrder}\<close>
section \<open>Allowing the Renaming of Quantified Variables\<close>
text \<open>So far, we did not need any renaming axiom for the quantifiers. However,
our axioms for substitution implicitly assume the irrelevance of the bound names;
in other words, their usual instances would have this property; and since this assumption
greatly simplifies the formal development, we make it at this point.\<close>
locale Syntax_with_Connectives_Rename =
Syntax_with_Connectives
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
+
assumes all_rename:
"\<And>\<phi> x y. \<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> y \<notin> Fvars \<phi> \<Longrightarrow>
all x \<phi> = all y (subst \<phi> (Var y) x)"
and exi_rename:
"\<And>\<phi> x y. \<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> y \<notin> Fvars \<phi> \<Longrightarrow>
exi x \<phi> = exi y (subst \<phi> (Var y) x)"
begin
lemma all_rename2:
"\<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> (y = x \<or> y \<notin> Fvars \<phi>) \<Longrightarrow>
all x \<phi> = all y (subst \<phi> (Var y) x)"
using all_rename by (cases "y = x") (auto simp del: Fvars_subst)
lemma exi_rename2:
"\<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> (y = x \<or> y \<notin> Fvars \<phi>) \<Longrightarrow>
exi x \<phi> = exi y (subst \<phi> (Var y) x)"
using exi_rename by (cases "y = x") (auto simp del: Fvars_subst)
section \<open>The Exists-Unique Quantifier\<close>
text \<open>It is phrased in such a way as to avoid substitution:\<close>
definition exu :: "'var \<Rightarrow> 'fmla \<Rightarrow> 'fmla" where
"exu x \<phi> \<equiv> let y = getFr [x] [] [\<phi>] in
cnj (exi x \<phi>) (exi y (all x (imp \<phi> (eql (Var x) (Var y)))))"
lemma exu[simp,intro]:
"x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> exu x \<phi> \<in> fmla"
unfolding exu_def by (simp add: Let_def)
lemma Fvars_exu[simp]:
"x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> Fvars (exu x \<phi>) = Fvars \<phi> - {x}"
unfolding exu_def by (auto simp: Let_def getFr_Fvars)
lemma exu_def_var:
assumes [simp]: "x \<in> var" "y \<in> var" "y \<noteq> x" "y \<notin> Fvars \<phi>" "\<phi> \<in> fmla"
shows
"exu x \<phi> = cnj (exi x \<phi>) (exi y (all x (imp \<phi> (eql (Var x) (Var y)))))"
proof-
have [simp]: "x \<noteq> y" using assms by blast
define z where z: "z \<equiv> getFr [x] [] [\<phi>]"
have z_facts[simp]: "z \<in> var" "z \<noteq> x" "x \<noteq> z" "z \<notin> Fvars \<phi>"
unfolding z using getFr_FvarsT_Fvars[of "[x]" "[]" "[\<phi>]"] by auto
define u where u: "u \<equiv> getFr [x,y,z] [] [\<phi>]"
have u_facts[simp]: "u \<in> var" "u \<noteq> x" "u \<noteq> z" "y \<noteq> u" "u \<noteq> y" "x \<noteq> u" "z \<noteq> u" "u \<notin> Fvars \<phi>"
unfolding u using getFr_FvarsT_Fvars[of "[x,y,z]" "[]" "[\<phi>]"] by auto
have "exu x \<phi> = cnj (exi x \<phi>) (exi u (all x (imp \<phi> (eql (Var x) (Var u)))))"
by (auto simp: exu_def Let_def z[symmetric] exi_rename[of "all x (imp \<phi> (eql (Var x) (Var z)))" z u])
also have "\<dots> = cnj (exi x \<phi>) (exi y (all x (imp \<phi> (eql (Var x) (Var y)))))"
by (auto simp: exi_rename[of "all x (imp \<phi> (eql (Var x) (Var u)))" u y]
split: if_splits)
finally show ?thesis .
qed
lemma subst_exu[simp]:
assumes [simp]: "\<phi> \<in> fmla" "t \<in> trm" "x \<in> var" "y \<in> var" "x \<noteq> y" "x \<notin> FvarsT t"
shows "subst (exu x \<phi>) t y = exu x (subst \<phi> t y)"
proof-
define u where u: "u \<equiv> getFr [x,y] [t] [\<phi>]"
have u_facts[simp]: "u \<in> var" "u \<noteq> x" "u \<noteq> y" "y \<noteq> u" "x \<noteq> u"
"u \<notin> FvarsT t" "u \<notin> Fvars \<phi>"
unfolding u using getFr_FvarsT_Fvars[of "[x,y]" "[t]" "[\<phi>]"] by auto
show ?thesis
by (auto simp: Let_def exu_def_var[of _ u] subst_compose_diff)
qed
lemma subst_exu_idle[simp]:
assumes [simp]: "x \<in> var" "\<phi> \<in> fmla" "t \<in> trm"
shows "subst (exu x \<phi>) t x = exu x \<phi>"
by (intro subst_notIn) auto
lemma exu_rename:
assumes [simp]: "\<phi> \<in> fmla" "x \<in> var" "y \<in> var" "y \<notin> Fvars \<phi>"
shows "exu x \<phi> = exu y (subst \<phi> (Var y) x)"
proof(cases "y = x")
case [simp]: False
define z where z: "z = getFr [x] [] [\<phi>]"
have z_facts[simp]: "z \<in> var" "z \<noteq> x" "x \<noteq> z" "z \<notin> Fvars \<phi>"
unfolding z using getFr_FvarsT_Fvars[of "[x]" "[]" "[\<phi>]"] by auto
define u where u: "u \<equiv> getFr [x,y,z] [] [\<phi>]"
have u_facts[simp]: "u \<in> var" "u \<noteq> x" "x \<noteq> u" "u \<noteq> y" "y \<noteq> u" "u \<noteq> z" "z \<noteq> u"
"u \<notin> Fvars \<phi>"
unfolding u using getFr_FvarsT_Fvars[of "[x,y,z]" "[]" "[\<phi>]"] by auto
show ?thesis
by (auto simp: exu_def_var[of _ u] exi_rename[of _ _ y] all_rename[of _ _ y])
qed auto
lemma exu_rename2:
"\<phi> \<in> fmla \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> (y = x \<or> y \<notin> Fvars \<phi>) \<Longrightarrow>
exu x \<phi> = exu y (subst \<phi> (Var y) x)"
using exu_rename by (cases "y = x") (auto simp del: Fvars_subst)
end \<comment> \<open>context @{locale Syntax_with_Connectives_Rename}\<close>
(*<*)
end
(*>*)
\ No newline at end of file
diff --git a/thys/Syntax_Independent_Logic/Syntax_Arith.thy b/thys/Syntax_Independent_Logic/Syntax_Arith.thy
--- a/thys/Syntax_Independent_Logic/Syntax_Arith.thy
+++ b/thys/Syntax_Independent_Logic/Syntax_Arith.thy
@@ -1,1669 +1,1669 @@
chapter \<open>Arithmetic Constructs\<close>
text \<open>Less genereric syntax, more committed towards embedding arithmetics\<close>
(*<*)
theory Syntax_Arith imports Syntax
begin
(*>*)
text \<open>(An embedding of) the syntax of arithmetic, obtained by adding plus and times\<close>
locale Syntax_Arith_aux =
Syntax_with_Connectives_Rename
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
+
Syntax_with_Numerals_and_Connectives_False_Disj
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
+
fixes
zer :: "'trm"
and
suc :: "'trm \<Rightarrow> 'trm"
and
pls :: "'trm \<Rightarrow> 'trm \<Rightarrow> 'trm"
and
tms :: "'trm \<Rightarrow> 'trm \<Rightarrow> 'trm"
assumes
Fvars_zero[simp,intro!]: "FvarsT zer = {}"
and
substT_zer[simp]: "\<And> t x. t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
substT zer t x = zer"
and
suc[simp]: "\<And>t. t \<in> trm \<Longrightarrow> suc t \<in> trm"
and
FvarsT_suc[simp]: "\<And> t. t \<in> trm \<Longrightarrow>
FvarsT (suc t) = FvarsT t"
and
substT_suc[simp]: "\<And> t1 t x. t1 \<in> trm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
substT (suc t1) t x = suc (substT t1 t x)"
and
pls[simp]: "\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> pls t1 t2 \<in> trm"
and
Fvars_pls[simp]: "\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow>
FvarsT (pls t1 t2) = FvarsT t1 \<union> FvarsT t2"
and
substT_pls[simp]: "\<And> t1 t2 t x. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
substT (pls t1 t2) t x = pls (substT t1 t x) (substT t2 t x)"
and
tms[simp]: "\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> tms t1 t2 \<in> trm"
and
Fvars_tms[simp]: "\<And> t1 t2. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow>
FvarsT (tms t1 t2) = FvarsT t1 \<union> FvarsT t2"
and
substT_tms[simp]: "\<And> t1 t2 t x. t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
substT (tms t1 t2) t x = tms (substT t1 t x) (substT t2 t x)"
begin
text \<open>The embedding of numbers into our abstract notion of numerals
(not required to be surjective)\<close>
fun Num :: "nat \<Rightarrow> 'trm" where
"Num 0 = zer"
|"Num (Suc n) = suc (Num n)"
end \<comment> \<open>context @{locale Syntax_Arith_aux}\<close>
locale Syntax_Arith =
Syntax_Arith_aux
var trm fmla Var FvarsT substT Fvars subst
eql cnj imp all exi
fls
dsj
num
zer suc pls tms
for
var :: "'var set" and trm :: "'trm set" and fmla :: "'fmla set"
and Var FvarsT substT Fvars subst
and eql cnj imp all exi
and fls
and dsj
and num
zer suc pls tms
+
assumes
\<comment> \<open>We assume that numbers are the only numerals:\<close>
num_Num: "num = range Num"
begin
lemma Num[simp,intro!]: "Num n \<in> num"
using num_Num by auto
lemma FvarsT_Num[simp]: "FvarsT (Num n) = {}"
by auto
lemma substT_Num[simp]: "x \<in> var \<Longrightarrow> t \<in> trm \<Longrightarrow> substT (Num n) t x = Num n"
by auto
lemma zer[simp,intro!]: "zer \<in> num"
and suc_num[simp]: "\<And>n. n \<in> num \<Longrightarrow> suc n \<in> num"
by (metis Num Num.simps(1), metis Num Num.simps(2) imageE num_Num)
section \<open>Arithmetic Terms\<close>
text \<open>Arithmetic terms are inductively defined to contain the numerals and the variables
and be closed under the arithmetic operators:\<close>
inductive_set atrm :: "'trm set" where
atrm_num[simp]: "n \<in> num \<Longrightarrow> n \<in> atrm"
|atrm_Var[simp,intro]: "x \<in> var \<Longrightarrow> Var x \<in> atrm"
|atrm_suc[simp,intro]: "t \<in> atrm \<Longrightarrow> suc t \<in> atrm"
|atrm_pls[simp,intro]: "t \<in> atrm \<Longrightarrow> t' \<in> atrm \<Longrightarrow> pls t t' \<in> atrm"
|atrm_tms[simp,intro]: "t \<in> atrm \<Longrightarrow> t' \<in> atrm \<Longrightarrow> tms t t' \<in> atrm"
lemma atrm_imp_trm[simp]: assumes "t \<in> atrm" shows "t \<in> trm"
using assms by induct auto
lemma atrm_trm: "atrm \<subseteq> trm"
using atrm_imp_trm by auto
lemma zer_atrm[simp]: "zer \<in> atrm" by auto
lemma Num_atrm[simp]: "Num n \<in> atrm"
by auto
lemma substT_atrm[simp]:
assumes "r \<in> atrm" and "x \<in> var" and "t \<in> atrm"
shows "substT r t x \<in> atrm"
using assms by (induct) auto
text \<open>Whereas we did not assume the rich set of formula-substitution properties to hold
for all terms, we can prove that these properties hold for arithmetic terms.\<close>
text \<open>Properties for arithmetic terms corresponding to the axioms for formulas:\<close>
lemma FvarsT_substT:
assumes "s \<in> atrm" "t \<in> trm" "x \<in> var"
shows "FvarsT (substT s t x) = (FvarsT s - {x}) \<union> (if x \<in> FvarsT s then FvarsT t else {})"
using assms by induct auto
lemma substT_compose_eq_or:
assumes "s \<in> atrm" "t1 \<in> trm" "t2 \<in> trm" "x1 \<in> var" "x2 \<in> var"
and "x1 = x2 \<or> x2 \<notin> FvarsT s"
shows "substT (substT s t1 x1) t2 x2 = substT s (substT t1 t2 x2) x1"
using assms apply induct
subgoal by auto
subgoal by auto
subgoal by (metis FvarsT_suc atrm_imp_trm substT substT_suc)
subgoal by (metis Fvars_pls UnCI atrm_imp_trm substT substT_pls)
subgoal by (metis Fvars_tms UnCI atrm_imp_trm substT substT_tms) .
lemma substT_compose_diff:
assumes "s \<in> atrm" "t1 \<in> trm" "t2 \<in> trm" "x1 \<in> var" "x2 \<in> var"
and "x1 \<noteq> x2" "x1 \<notin> FvarsT t2"
shows "substT (substT s t1 x1) t2 x2 = substT (substT s t2 x2) (substT t1 t2 x2) x1"
using assms apply induct
subgoal by auto
subgoal by auto
subgoal by (metis atrm_imp_trm substT substT_suc)
subgoal by (metis atrm_imp_trm substT substT_pls)
subgoal by (metis atrm_imp_trm substT substT_tms) .
lemma substT_same_Var[simp]:
assumes "s \<in> atrm" "x \<in> var"
shows "substT s (Var x) x = s"
using assms by induct auto
text \<open>... and corresponding to some corollaries we proved for formulas
(with essentially the same proofs):\<close>
lemma in_FvarsT_substTD:
"y \<in> FvarsT (substT r t x) \<Longrightarrow> r \<in> atrm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var
\<Longrightarrow> y \<in> (FvarsT r - {x}) \<union> (if x \<in> FvarsT r then FvarsT t else {})"
using FvarsT_substT by auto
lemma substT_compose_same:
"\<And> s t1 t2 x. s \<in> atrm \<Longrightarrow> t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow>
substT (substT s t1 x) t2 x = substT s (substT t1 t2 x) x"
using substT_compose_eq_or by blast
lemma substT_substT[simp]:
assumes s[simp]: "s \<in> atrm" and t[simp]:"t \<in> trm" and x[simp]:"x \<in> var" and y[simp]:"y \<in> var"
assumes yy: "x \<noteq> y" "y \<notin> FvarsT s"
shows "substT (substT s (Var y) x) t y = substT s t x"
using substT_compose_eq_or[OF s _ t x y, of "Var y"] using subst_notIn yy by simp
lemma substT_comp:
"\<And> x y s t. s \<in> atrm \<Longrightarrow> t \<in> trm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow>
x \<noteq> y \<Longrightarrow> y \<notin> FvarsT t \<Longrightarrow>
substT (substT s (Var x) y) t x = substT (substT s t x) t y"
by (simp add: substT_compose_diff)
text \<open>Now the corresponding development of parallel substitution for arithmetic terms:\<close>
lemma rawpsubstT_atrm[simp,intro]:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
shows "rawpsubstT r txs \<in> atrm"
using assms by (induct txs arbitrary: r) auto
lemma psubstT_atrm[simp,intro]:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
shows "psubstT r txs \<in> atrm"
proof-
have txs_trm: "fst ` (set txs) \<subseteq> trm" using assms atrm_trm by auto
define us where us: "us \<equiv> getFrN (map snd txs) (r # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(1,2) txs_trm unfolding us
using getFrN_FvarsT[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
getFrN_distinct[of "map snd txs" "r # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
by auto
(* *)
show ?thesis using assms us_facts unfolding psubstT_def
by (force simp: Let_def us[symmetric]
intro!: rawpsubstT_atrm[of _ "zip (map fst txs) us"] dest!: set_zip_D)
qed
lemma Fvars_rawpsubst_su:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
shows "FvarsT (rawpsubstT r txs) \<subseteq>
(FvarsT r - snd ` (set txs)) \<union> (\<Union> {FvarsT t | t x . (t,x) \<in> set txs})"
using assms proof(induction txs arbitrary: r)
case (Cons tx txs r)
then obtain t x where tx: "tx = (t,x)" by force
have t: "t \<in> trm" and x: "x \<in> var" using Cons.prems unfolding tx by auto
define \<chi> where "\<chi> \<equiv> substT r t x"
have 0: "FvarsT \<chi> = FvarsT r - {x} \<union> (if x \<in> FvarsT r then FvarsT t else {})"
using Cons.prems unfolding \<chi>_def by (auto simp: tx t FvarsT_substT)
have \<chi>: "\<chi> \<in> trm" "\<chi> \<in> atrm" unfolding \<chi>_def using Cons.prems t x by (auto simp add: tx)
have "FvarsT (rawpsubstT \<chi> txs) \<subseteq>
(FvarsT \<chi> - snd ` (set txs)) \<union>
(\<Union> {FvarsT t | t x . (t,x) \<in> set txs})"
using Cons.prems \<chi> by (intro Cons.IH) auto
also have "\<dots> \<subseteq> FvarsT r - insert x (snd ` set txs) \<union> \<Union>{FvarsT ta |ta. \<exists>xa. ta = t \<and> xa = x \<or> (ta, xa) \<in> set txs}"
(is "_ \<subseteq> ?R") by(auto simp: 0 tx Cons.prems)
finally have 1: "FvarsT (rawpsubstT \<chi> txs) \<subseteq> ?R" .
have 2: "FvarsT \<chi> = FvarsT r - {x} \<union> (if x \<in> FvarsT r then FvarsT t else {})"
using Cons.prems t x unfolding \<chi>_def using FvarsT_substT by auto
show ?case using 1 by (simp add: tx \<chi>_def[symmetric] 2)
qed auto
lemma in_FvarsT_rawpsubstT_imp:
assumes "y \<in> FvarsT (rawpsubstT r txs)"
and "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
shows "(y \<in> FvarsT r - snd ` (set txs)) \<or>
(y \<in> \<Union> { FvarsT t | t x . (t,x) \<in> set txs})"
using Fvars_rawpsubst_su[OF assms(2-4)]
using assms(1) by blast
lemma FvarsT_rawpsubstT:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)" and "\<forall> x \<in> snd ` (set txs). \<forall> t \<in> fst ` (set txs). x \<notin> FvarsT t"
shows "FvarsT (rawpsubstT r txs) =
(FvarsT r - snd ` (set txs)) \<union>
(\<Union> {if x \<in> FvarsT r then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using assms proof(induction txs arbitrary: r)
case (Cons a txs r)
then obtain t x where a: "a = (t,x)" by force
have t: "t \<in> trm" and x: "x \<in> var" using Cons.prems unfolding a by auto
have xt: "x \<notin> FvarsT t \<and> snd ` set txs \<inter> FvarsT t = {}" using Cons.prems unfolding a by auto
hence 0: "FvarsT r - {x} \<union> FvarsT t - snd ` set txs = FvarsT r - insert x (snd ` set txs) \<union> FvarsT t"
by auto
- have x_txs: "\<And>ta xa. (ta, xa) \<in> set txs \<Longrightarrow> x \<noteq> xa" using `distinct (map snd (a # txs))`
+ have x_txs: "\<And>ta xa. (ta, xa) \<in> set txs \<Longrightarrow> x \<noteq> xa" using \<open>distinct (map snd (a # txs))\<close>
unfolding a by (auto simp: rev_image_eqI)
define \<chi> where \<chi>_def: "\<chi> \<equiv> substT r t x"
have \<chi>: "\<chi> \<in> trm" "\<chi> \<in> atrm" unfolding \<chi>_def using Cons.prems t x by (auto simp: a)
have 1: "FvarsT (rawpsubstT \<chi> txs) =
(FvarsT \<chi> - snd ` (set txs)) \<union>
(\<Union> {if x \<in> FvarsT \<chi> then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using Cons.prems \<chi> by (intro Cons.IH) auto
have 2: "FvarsT \<chi> = FvarsT r - {x} \<union> (if x \<in> FvarsT r then FvarsT t else {})"
using Cons.prems t x unfolding \<chi>_def using FvarsT_substT by auto
define f where "f \<equiv> \<lambda>ta xa. if xa \<in> FvarsT r then FvarsT ta else {}"
have 3: "\<Union> {f ta xa |ta xa. (ta, xa) \<in> set ((t, x) # txs)} =
f t x \<union> (\<Union> {f ta xa |ta xa. (ta, xa) \<in> set txs})" by auto
have 4: "snd ` set ((t, x) # txs) = {x} \<union> snd ` set txs" by auto
have 5: "f t x \<inter> snd ` set txs = {}" unfolding f_def using xt by auto
have 6: "\<Union> {if xa \<in> FvarsT r - {x} \<union> f t x then FvarsT ta else {} | ta xa. (ta, xa) \<in> set txs}
= (\<Union> {f ta xa | ta xa. (ta, xa) \<in> set txs})"
unfolding f_def using xt x_txs by (fastforce split: if_splits)
have "FvarsT r - {x} \<union> f t x - snd ` set txs \<union>
\<Union> {if xa \<in> FvarsT r - {x} \<union> f t x then FvarsT ta else {}
| ta xa. (ta, xa) \<in> set txs} =
FvarsT r - snd ` set ((t, x) # txs) \<union>
\<Union> {f ta xa |ta xa. (ta, xa) \<in> set ((t, x) # txs)}"
unfolding 3 4 6 unfolding Un_Diff2[OF 5] Un_assoc unfolding Diff_Diff_Un ..
thus ?case unfolding a rawpsubstT.simps 1 2 \<chi>_def[symmetric] f_def by simp
qed auto
lemma in_FvarsT_rawpsubstTD:
assumes "y \<in> FvarsT (rawpsubstT r txs)"
and "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)" and "\<forall> x \<in> snd ` (set txs). \<forall> t \<in> fst ` (set txs). x \<notin> FvarsT t"
shows "(y \<in> FvarsT r - snd ` (set txs)) \<or>
(y \<in> \<Union> {if x \<in> FvarsT r then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using FvarsT_rawpsubstT assms by auto
lemma FvarsT_psubstT:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "FvarsT (psubstT r txs) =
(FvarsT r - snd ` (set txs)) \<union>
(\<Union> {if x \<in> FvarsT r then FvarsT t else {} | t x . (t,x) \<in> set txs})"
proof-
have txs_trm: "fst ` (set txs) \<subseteq> trm" using assms by auto
define us where us: "us \<equiv> getFrN (map snd txs) (r # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(1,2) txs_trm unfolding us
using getFrN_FvarsT[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
by auto
have [simp]: "\<And> aa b. b \<in> set (map snd txs) \<Longrightarrow>
aa \<in> set (map Var us) \<Longrightarrow> b \<notin> FvarsT aa"
using us_facts by (fastforce simp: image_def Int_def)
have [simp]:
"\<And>b ac bc. b \<in> set us \<Longrightarrow> b \<in> FvarsT ac \<Longrightarrow> (ac, bc) \<notin> set txs"
using us_facts(3) by (fastforce simp: image_def Int_def)
define \<chi> where \<chi>_def: "\<chi> \<equiv> rawpsubstT r (zip (map Var us) (map snd txs))"
have \<chi>: "\<chi> \<in> atrm" unfolding \<chi>_def
using assms using us_facts by (intro rawpsubstT_atrm) (force dest!: set_zip_D)+
hence "\<chi> \<in> trm" by auto
note \<chi> = \<chi> this
have set_us: "set us = snd ` (set (zip (map fst txs) us))"
using us_facts by (intro snd_set_zip[symmetric]) auto
have set_txs: "snd ` set txs = snd ` (set (zip (map Var us) (map snd txs)))"
using us_facts by (intro snd_set_zip_map_snd[symmetric]) auto
have "\<And> t x. (t, x) \<in> set (zip (map Var us) (map snd txs)) \<Longrightarrow> \<exists> u. t = Var u"
using us_facts set_zip_leftD by fastforce
hence 00: "\<And> t x. (t, x) \<in> set (zip (map Var us) (map snd txs))
\<longleftrightarrow> (\<exists> u \<in> var. t = Var u \<and> (Var u, x) \<in> set (zip (map Var us) (map snd txs)))"
using us_facts set_zip_leftD by fastforce
have "FvarsT \<chi> =
FvarsT r - snd ` set txs \<union>
\<Union>{if x \<in> FvarsT r then FvarsT t else {} |t x.
(t, x) \<in> set (zip (map Var us) (map snd txs))}"
unfolding \<chi>_def set_txs using assms us_facts set_txs
by (intro FvarsT_rawpsubstT) (force dest!: set_zip_D)+
also have "\<dots> =
FvarsT r - snd ` set txs \<union>
\<Union>{if x \<in> FvarsT r then {u} else {} |u x. u \<in> var \<and> (Var u, x) \<in> set (zip (map Var us) (map snd txs))}"
(is "\<dots> = ?R")
apply(subst 00)
by (metis (no_types, hide_lams) FvarsT_Var)
finally have 0: "FvarsT \<chi> = ?R" .
have 1: "FvarsT (rawpsubstT \<chi> (zip (map fst txs) us)) =
(FvarsT \<chi> - set us) \<union>
(\<Union> {if u \<in> FvarsT \<chi> then FvarsT t else {} | t u . (t,u) \<in> set (zip (map fst txs) us)})"
unfolding us_facts set_us using assms \<chi> apply (intro FvarsT_rawpsubstT)
subgoal by auto
subgoal using us_facts by (auto dest!: set_zip_D)
subgoal using us_facts by (auto dest!: set_zip_D)
subgoal using us_facts by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D) .
have 2: "FvarsT \<chi> - set us = FvarsT r - snd ` set txs"
unfolding 0 apply auto
using set_zip_leftD us_facts(1) apply fastforce
using set_zip_leftD us_facts(1) apply fastforce
using us_facts(2) by auto
have 3:
"(\<Union> {if u \<in> FvarsT \<chi> then FvarsT t else {} | t u . (t,u) \<in> set (zip (map fst txs) us)}) =
(\<Union> {if x \<in> FvarsT r then FvarsT t else {} | t x . (t,x) \<in> set txs})"
proof safe
fix xx tt y
assume xx: "xx \<in> (if y \<in> FvarsT \<chi> then FvarsT tt else {})"
and ty: "(tt, y) \<in> set (zip (map fst txs) us)"
have ttin: "tt \<in> fst ` set txs" using ty using set_zip_leftD by fastforce
have yin: "y \<in> set us" using ty by (meson set_zip_D)
have yvar: "y \<in> var" using us_facts yin by auto
have ynotin: "y \<notin> snd ` set txs" "y \<notin> FvarsT r" using yin us_facts by auto
show "xx \<in> \<Union>{if x \<in> FvarsT r then FvarsT t else {} |t x. (t, x) \<in> set txs}"
proof(cases "y \<in> FvarsT \<chi>")
case True note y = True
hence xx: "xx \<in> FvarsT tt" using xx by simp
obtain x where xr: "x \<in> FvarsT r"
and yx: "(Var y, x) \<in> set (zip (map Var us) (map snd txs))"
using y ynotin unfolding 0 by (auto split: if_splits)
have yx: "(y, x) \<in> set (zip us (map snd txs))"
using yvar us_facts by (intro inj_on_set_zip_map[OF inj_on_Var yx]) auto
have "(tt, x) \<in> set txs" apply(rule set_zip_map_fst_snd[OF yx ty])
- using `distinct (map snd txs)` us_facts by auto
+ using \<open>distinct (map snd txs)\<close> us_facts by auto
thus ?thesis using xx xr by auto
qed(insert xx, auto)
next
fix y tt xx
assume y: "y \<in> (if xx \<in> FvarsT r then FvarsT tt else {})"
and tx: "(tt, xx) \<in> set txs"
hence xxsnd: "xx \<in> snd ` set txs" by force
obtain u where uin: "u \<in> set us" and uxx: "(u, xx) \<in> set (zip us (map snd txs))"
by (metis xxsnd in_set_impl_in_set_zip2 length_map set_map set_zip_leftD us_facts(5))
hence uvar: "u \<in> var" using us_facts by auto
show "y \<in> \<Union>{if u \<in> FvarsT \<chi> then FvarsT t else {} |t u. (t, u) \<in> set (zip (map fst txs) us)}"
proof(cases "xx \<in> FvarsT r")
case True note xx = True
hence y: "y \<in> FvarsT tt" using y by auto
have "(Var u, xx) \<in> set (zip (map Var us) (map snd txs))"
apply(rule set_zip_length_map[OF uxx]) using us_facts by auto
hence u\<chi>: "u \<in> FvarsT \<chi>" using uin xx uvar unfolding 0 by auto
have ttu: "(tt, u) \<in> set (zip (map fst txs) us)"
apply(rule set_zip_map_fst_snd2[OF uxx tx]) using assms us_facts by auto
show ?thesis using u\<chi> ttu y by auto
qed(insert y, auto)
qed
show ?thesis
by (simp add: psubstT_def Let_def us[symmetric] \<chi>_def[symmetric] 1 2 3)
qed
lemma in_FvarsT_psubstTD:
assumes "y \<in> FvarsT (psubstT r txs)"
and "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "y \<in> (FvarsT r - snd ` (set txs)) \<union>
(\<Union> {if x \<in> FvarsT r then FvarsT t else {} | t x . (t,x) \<in> set txs})"
using assms FvarsT_psubstT by auto
lemma substT2_fresh_switch:
assumes "r \<in> atrm" "t \<in> trm" "s \<in> trm" "x \<in> var" "y \<in> var"
and "x \<noteq> y" "x \<notin> FvarsT s" "y \<notin> FvarsT t"
shows "substT (substT r s y) t x = substT (substT r t x) s y" (is "?L = ?R")
using assms by (simp add: substT_compose_diff[of r s t y x])
lemma rawpsubst2_fresh_switch:
assumes "r \<in> atrm" "t \<in> trm" "s \<in> trm" "x \<in> var" "y \<in> var"
and "x \<noteq> y" "x \<notin> FvarsT s" "y \<notin> FvarsT t"
shows "rawpsubstT r ([(s,y),(t,x)]) = rawpsubstT r ([(t,x),(s,y)])"
using assms by (simp add: substT2_fresh_switch)
(* this actually works for any trms, does not need atrms: *)
lemma rawpsubstT_compose:
assumes "t \<in> trm" and "snd ` (set txs1) \<subseteq> var" and "fst ` (set txs1) \<subseteq> atrm"
and "snd ` (set txs2) \<subseteq> var" and "fst ` (set txs2) \<subseteq> atrm"
shows "rawpsubstT (rawpsubstT t txs1) txs2 = rawpsubstT t (txs1 @ txs2)"
using assms apply (induct txs1 arbitrary: txs2 t)
subgoal by simp
subgoal for tx1 txs1 txs2 t apply (cases tx1) by auto .
lemma rawpsubstT_subst_fresh_switch:
assumes "r \<in> atrm" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "\<forall> x \<in> snd ` (set txs). x \<notin> FvarsT s"
and "\<forall> t \<in> fst ` (set txs). y \<notin> FvarsT t"
and "distinct (map snd txs)"
and "s \<in> atrm" and "y \<in> var" "y \<notin> snd ` (set txs)"
shows "rawpsubstT (substT r s y) txs = rawpsubstT r (txs @ [(s,y)])"
using assms proof(induction txs arbitrary: r s y)
case (Cons tx txs)
obtain t x where tx[simp]: "tx = (t,x)" by force
have x: "x \<in> var" and t: "t \<in> trm" using Cons unfolding tx by auto
have "rawpsubstT r ((s, y) # (t, x) # txs) = rawpsubstT r ([(s, y),(t, x)] @ txs)" by simp
also have "\<dots> = rawpsubstT (rawpsubstT r [(s, y),(t, x)]) txs"
using Cons by auto
also have "rawpsubstT r [(s, y),(t, x)] = rawpsubstT r [(t, x),(s, y)]"
using Cons by (intro rawpsubst2_fresh_switch) auto
also have "rawpsubstT (rawpsubstT r [(t, x),(s, y)]) txs = rawpsubstT r ([(t, x),(s, y)] @ txs)"
using Cons by (intro rawpsubstT_compose) auto
also have "\<dots> = rawpsubstT (substT r t x) (txs @ [(s,y)])" using Cons by auto
also have "\<dots> = rawpsubstT r (((t, x) # txs) @ [(s, y)])" by simp
finally show ?case unfolding tx by auto
qed auto
lemma substT_rawpsubstT_fresh_switch:
assumes "r \<in> atrm" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "\<forall> x \<in> snd ` (set txs). x \<notin> FvarsT s"
and "\<forall> t \<in> fst ` (set txs). y \<notin> FvarsT t"
and "distinct (map snd txs)"
and "s \<in> atrm" and "y \<in> var" "y \<notin> snd ` (set txs)"
shows "substT (rawpsubstT r txs) s y = rawpsubstT r ((s,y) # txs)"
using assms proof(induction txs arbitrary: r s y)
case (Cons tx txs)
obtain t x where tx[simp]: "tx = (t,x)" by force
have x: "x \<in> var" and t: "t \<in> trm" using Cons unfolding tx by auto
have "substT (rawpsubstT (substT r t x) txs) s y = rawpsubstT (substT r t x) ((s,y) # txs)"
using Cons.prems by (intro Cons.IH) auto
also have "\<dots> = rawpsubstT (rawpsubstT r [(t,x)]) ((s,y) # txs)" by simp
also have "\<dots> = rawpsubstT r ([(t,x)] @ ((s,y) # txs))"
using Cons.prems by (intro rawpsubstT_compose) auto
also have "\<dots> = rawpsubstT r ([(t,x),(s,y)] @ txs)" by simp
also have "\<dots> = rawpsubstT (rawpsubstT r [(t,x),(s,y)]) txs"
using Cons.prems by (intro rawpsubstT_compose[symmetric]) auto
also have "rawpsubstT r [(t,x),(s,y)] = rawpsubstT r [(s,y),(t,x)]"
using Cons.prems by (intro rawpsubst2_fresh_switch) auto
also have "rawpsubstT (rawpsubstT r [(s,y),(t,x)]) txs = rawpsubstT r ([(s,y),(t,x)] @ txs)"
using Cons.prems by (intro rawpsubstT_compose) auto
finally show ?case by simp
qed auto
lemma rawpsubstT_compose_freshVar:
assumes "r \<in> atrm" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
and "\<And> i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> snd (txs!j) \<notin> FvarsT (fst (txs!i))"
and us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
shows "rawpsubstT (rawpsubstT r (zip (map Var us) (map snd txs))) (zip (map fst txs) us) = rawpsubstT r txs"
using assms proof(induction txs arbitrary: us r)
case (Cons tx txs uus r)
obtain t x where tx[simp]: "tx = (t,x)" by force
obtain u us where uus[simp]: "uus = u # us" using Cons by (cases uus) auto
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us" and u_facts: "u \<in> var" "u \<notin> FvarsT r"
"u \<notin> \<Union> (FvarsT ` (fst ` (set txs)))"
"u \<notin> snd ` (set txs)" "u \<notin> set us"
using Cons by auto
have [simp]: "\<And> bb xaa ab. bb \<in> FvarsT (Var xaa) \<Longrightarrow>
(ab, bb) \<in> set txs \<Longrightarrow> xaa \<notin> set us"
using us_facts(1,4) by force
let ?uxs = "zip (map Var us) (map snd txs)"
have 1: "rawpsubstT (substT r (Var u) x) ?uxs = rawpsubstT r (?uxs @ [(Var u,x)])"
using Cons.prems u_facts apply(intro rawpsubstT_subst_fresh_switch)
subgoal by auto
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest!: set_zip_D)
by (auto dest!: set_zip_D)
let ?uuxs = "zip (map Var uus) (map snd (tx # txs))"
let ?tus = "zip (map fst txs) us" let ?ttxs = "zip (map fst (tx # txs)) uus"
have 2: "u \<in> FvarsT (rawpsubstT r (zip (map Var us) (map snd txs))) \<Longrightarrow> False"
apply(drule in_FvarsT_rawpsubstTD) apply-
subgoal using Cons.prems by auto
subgoal using Cons.prems by (auto dest!: set_zip_D)
subgoal using Cons.prems by (force dest!: set_zip_D)
subgoal using Cons.prems by (auto dest!: set_zip_D)
subgoal by (auto dest!: set_zip_D)
subgoal using us_facts(1,4,5) Cons.prems(7)
by(fastforce dest!: set_zip_D split: if_splits simp: u_facts(5)) .
have 3: "(tt, xx) \<notin> set txs" if "xx \<in> FvarsT t" for tt xx
unfolding set_conv_nth mem_Collect_eq
proof safe
fix i
assume "(tt, xx) = txs ! i" "i < length txs"
then show False
using that Cons.prems(4) Cons.prems(5)[of 0 "Suc i"] tx
by (auto simp: nth_Cons' split: if_splits dest: sym)
qed
have 00: "rawpsubstT (rawpsubstT r ?uuxs) ?ttxs = rawpsubstT (substT (rawpsubstT r (?uxs @ [(Var u, x)])) t u) ?tus"
by (simp add: 1)
have "rawpsubstT r (?uxs @ [(Var u, x)]) = rawpsubstT (rawpsubstT r ?uxs) [(Var u, x)]"
using Cons.prems
by (intro rawpsubstT_compose[symmetric]) (auto 0 3 dest!: set_zip_D)
also have "rawpsubstT (rawpsubstT r ?uxs) [(Var u, x)] = substT (rawpsubstT r ?uxs) (Var u) x" by simp
finally have "substT (rawpsubstT r (?uxs @ [(Var u, x)])) t u =
substT (substT (rawpsubstT r ?uxs) (Var u) x) t u" by simp
also have "\<dots> = substT (rawpsubstT r ?uxs) t x"
using Cons 2 by (intro substT_substT) (auto 0 3 intro!: rawpsubstT_atrm[of r] dest!: set_zip_D)
also have "\<dots> = rawpsubstT r ((t,x) # ?uxs)"
using Cons.prems 3
by (intro substT_rawpsubstT_fresh_switch) (auto 0 3 dest!: set_zip_D FvarsT_VarD)
also have "\<dots> = rawpsubstT r ([(t,x)] @ ?uxs)" by simp
also have "\<dots> = rawpsubstT (rawpsubstT r [(t,x)]) ?uxs"
using Cons.prems by (intro rawpsubstT_compose[symmetric]) (auto 0 3 dest!: set_zip_D)
finally have "rawpsubstT (substT (rawpsubstT r (?uxs @ [(Var u, x)])) t u) ?tus =
rawpsubstT (rawpsubstT (rawpsubstT r [(t,x)]) ?uxs) ?tus" by auto
hence "rawpsubstT (rawpsubstT r ?uuxs) ?ttxs = rawpsubstT (rawpsubstT (rawpsubstT r [(t,x)]) ?uxs) ?tus"
using 00 by auto
also have "\<dots> = rawpsubstT (rawpsubstT r [(t,x)]) txs"
using Cons.prems apply(intro Cons.IH)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by (metis Suc_leI le_imp_less_Suc length_Cons nth_Cons_Suc)
subgoal by auto
subgoal by (auto intro!: rawpsubstT dest!: set_zip_D in_FvarsT_substTD
split: if_splits)
by auto
finally show ?case by simp
qed auto
lemma rawpsubstT_compose_freshVar2_aux:
assumes r[simp]: "r \<in> atrm"
and ts: "set ts \<subseteq> atrm"
and xs: "set xs \<subseteq> var" "distinct xs"
and us_facts: "set us \<subseteq> var" "distinct us"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set us \<inter> set xs = {}"
and vs_facts: "set vs \<subseteq> var" "distinct vs"
"set vs \<inter> FvarsT r = {}"
"set vs \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set vs \<inter> set xs = {}"
and l: "length us = length xs" "length vs = length xs" "length ts = length xs"
and (* Extra hypothesis, only to get induction through: *) d: "set us \<inter> set vs = {}"
shows "rawpsubstT (rawpsubstT r (zip (map Var us) xs)) (zip ts us) =
rawpsubstT (rawpsubstT r (zip (map Var vs) xs)) (zip ts vs)"
using assms proof(induction xs arbitrary: r ts us vs)
case (Cons x xs r tts uus vvs)
obtain t ts u us v vs where tts[simp]: "tts = t # ts" and lts[simp]: "length ts = length xs"
and uus[simp]: "uus = u # us" and lus[simp]: "length us = length xs"
and vvs[simp]: "vvs = v # vs" and lvs[simp]: "length vs = length xs"
- using `length uus = length (x # xs)` `length vvs = length (x # xs)` `length tts = length (x # xs)`
+ using \<open>length uus = length (x # xs)\<close> \<open>length vvs = length (x # xs)\<close> \<open>length tts = length (x # xs)\<close>
apply(cases tts)
subgoal by auto
subgoal apply(cases uus)
subgoal by auto
subgoal by (cases vvs) auto . .
let ?rux = "substT r (Var u) x" let ?rvx = "substT r (Var v) x"
have 0: "rawpsubstT (rawpsubstT ?rux (zip (map Var us) xs)) (zip ts us) =
rawpsubstT (rawpsubstT ?rux (zip (map Var vs) xs)) (zip ts vs)"
using Cons.prems by (intro Cons.IH) (auto intro!: rawpsubstT dest!: set_zip_D simp: FvarsT_substT)
have 1: "rawpsubstT ?rux (zip (map Var vs) xs) =
substT (rawpsubstT r (zip (map Var vs) xs)) (Var u) x"
using Cons.prems
by (intro substT_rawpsubstT_fresh_switch[simplified,symmetric])
(auto intro!: rawpsubstT dest!: set_zip_D simp: subset_eq)
have 11: "rawpsubstT ?rvx (zip (map Var vs) xs) =
substT (rawpsubstT r (zip (map Var vs) xs)) (Var v) x"
using Cons.prems
by (intro substT_rawpsubstT_fresh_switch[simplified,symmetric])
(auto intro!: rawpsubstT dest!: set_zip_D simp: subset_eq)
have "substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var u) x) t u =
substT (rawpsubstT r (zip (map Var vs) xs)) t x"
using Cons.prems
by (intro substT_substT)
(auto 0 3 intro!: rawpsubstT_atrm[of r]
dest!: set_zip_D in_FvarsT_rawpsubstT_imp FvarsT_VarD simp: FvarsT_rawpsubstT)
also have "\<dots> = substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var v) x) t v"
using Cons.prems
by (intro substT_substT[symmetric])
(auto 0 3 intro!: rawpsubstT_atrm[of r] dest!: set_zip_D in_FvarsT_rawpsubstT_imp FvarsT_VarD
simp: FvarsT_rawpsubstT)
finally have
2: "substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var u) x) t u =
substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var v) x) t v" .
have "rawpsubstT (substT (rawpsubstT ?rux (zip (map Var us) xs)) t u) (zip ts us) =
substT (rawpsubstT (rawpsubstT ?rux (zip (map Var us) xs)) (zip ts us)) t u"
using Cons.prems
by (intro substT_rawpsubstT_fresh_switch[simplified,symmetric])
(auto 0 3 intro!: rawpsubstT_atrm[of ?rux] substT_atrm dest!: set_zip_D)
also have "\<dots> = substT (rawpsubstT (rawpsubstT ?rux (zip (map Var vs) xs)) (zip ts vs)) t u"
unfolding 0 ..
also have "\<dots> = rawpsubstT (substT (rawpsubstT ?rux (zip (map Var vs) xs)) t u) (zip ts vs)"
using Cons.prems
by (intro substT_rawpsubstT_fresh_switch[simplified])
(auto 0 3 intro!: rawpsubstT_atrm[of ?rux] dest!: set_zip_D)
also have "\<dots> = rawpsubstT (substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var u) x) t u) (zip ts vs)"
unfolding 1 ..
also have "\<dots> = rawpsubstT (substT (substT (rawpsubstT r (zip (map Var vs) xs)) (Var v) x) t v) (zip ts vs)"
unfolding 2 ..
also have "\<dots> = rawpsubstT (substT (rawpsubstT ?rvx (zip (map Var vs) xs)) t v) (zip ts vs)"
unfolding 11 ..
finally have "rawpsubstT (substT (rawpsubstT ?rux (zip (map Var us) xs)) t u) (zip ts us) =
rawpsubstT (substT (rawpsubstT ?rvx (zip (map Var vs) xs)) t v) (zip ts vs)" .
thus ?case by simp
qed auto
(* ... now getting rid of the disjointness hypothesis: *)
lemma rawpsubstT_compose_freshVar2:
assumes r[simp]: "r \<in> atrm"
and ts: "set ts \<subseteq> atrm"
and xs: "set xs \<subseteq> var" "distinct xs"
and us_facts: "set us \<subseteq> var" "distinct us"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set us \<inter> set xs = {}"
and vs_facts: "set vs \<subseteq> var" "distinct vs"
"set vs \<inter> FvarsT r = {}"
"set vs \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set vs \<inter> set xs = {}"
and l: "length us = length xs" "length vs = length xs" "length ts = length xs"
shows "rawpsubstT (rawpsubstT r (zip (map Var us) xs)) (zip ts us) =
rawpsubstT (rawpsubstT r (zip (map Var vs) xs)) (zip ts vs)" (is "?L = ?R")
proof-
have ts_trm: "set ts \<subseteq> trm" using ts by auto
define ws where "ws = getFrN (xs @ us @ vs) (r # ts) [] (length xs)"
have ws_facts: "set ws \<subseteq> var" "distinct ws"
"set ws \<inter> FvarsT r = {}"
"set ws \<inter> \<Union> (FvarsT ` (set ts)) = {}"
"set ws \<inter> set xs = {}" "set ws \<inter> set us = {}" "set ws \<inter> set vs = {}"
"length ws = length xs" using assms(1) ts_trm assms(3-17) unfolding ws_def
using getFrN_Fvars[of "xs @ us @ vs" "r # ts" "[]" _ "length xs"]
getFrN_FvarsT[of "xs @ us @ vs" "r # ts" "[]" _ "length xs"]
getFrN_var[of "xs @ us @ vs" "r # ts" "[]" _ "length xs"]
getFrN_length[of "xs @ us @ vs" "r # ts" "[]" "length xs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
subgoal by force
subgoal by force
subgoal by force
by auto
have "?L = rawpsubstT (rawpsubstT r (zip (map Var ws) xs)) (zip ts ws)"
using assms ws_facts by (intro rawpsubstT_compose_freshVar2_aux) auto
also have "\<dots> = ?R"
using assms ws_facts by (intro rawpsubstT_compose_freshVar2_aux) auto
finally show ?thesis .
qed
lemma in_fst_image: "a \<in> fst ` AB \<longleftrightarrow> (\<exists>b. (a,b) \<in> AB)" by force
(* For many cases, the simpler rawpsubstT can replace psubst: *)
lemma psubstT_eq_rawpsubstT:
assumes "r \<in> atrm" "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
(* ... namely, when the substituted variables do not belong to trms substituted for previous variables: *)
and "\<And> i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> snd (txs!j) \<notin> FvarsT (fst (txs!i))"
shows "psubstT r txs = rawpsubstT r txs"
proof-
have txs_trm: "r \<in> trm" "fst ` (set txs) \<subseteq> trm" using assms by auto
note frt = getFrN_FvarsT[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
and fr = getFrN_Fvars[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
and var = getFrN_var[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
and l = getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
define us where us: "us \<equiv> getFrN (map snd txs) (r # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(2,4,5) txs_trm unfolding us
apply -
subgoal by auto
subgoal using frt by auto
subgoal using frt by (simp add: in_fst_image Int_def) (metis prod.collapse)
subgoal using var by (simp add: in_fst_image Int_def) (metis)
subgoal using l by auto
subgoal by auto .
show ?thesis
using rawpsubstT_compose_freshVar assms us_facts
by (simp add: psubstT_def Let_def us[symmetric])
qed
(* Some particular cases: *)
lemma psubstT_eq_substT:
assumes "r \<in> atrm" "x \<in> var" and "t \<in> atrm"
shows "psubstT r [(t,x)] = substT r t x"
proof-
have "psubstT r [(t,x)] = rawpsubstT r [(t,x)]"
using assms by (intro psubstT_eq_rawpsubstT) auto
thus ?thesis by auto
qed
lemma psubstT_eq_rawpsubst2:
assumes "r \<in> atrm" "x1 \<in> var" "x2 \<in> var" "t1 \<in> atrm" "t2 \<in> atrm"
and "x1 \<noteq> x2" "x2 \<notin> FvarsT t1"
shows "psubstT r [(t1,x1),(t2,x2)] = rawpsubstT r [(t1,x1),(t2,x2)]"
using assms using less_SucE by (intro psubstT_eq_rawpsubstT) force+
lemma psubstT_eq_rawpsubst3:
assumes "r \<in> atrm" "x1 \<in> var" "x2 \<in> var" "x3 \<in> var" "t1 \<in> atrm" "t2 \<in> atrm" "t3 \<in> atrm"
and "x1 \<noteq> x2" "x1 \<noteq> x3" "x2 \<noteq> x3"
"x2 \<notin> FvarsT t1" "x3 \<notin> FvarsT t1" "x3 \<notin> FvarsT t2"
shows "psubstT r [(t1,x1),(t2,x2),(t3,x3)] = rawpsubstT r [(t1,x1),(t2,x2),(t3,x3)]"
using assms less_SucE less_Suc_eq_0_disj
by (intro psubstT_eq_rawpsubstT) auto
lemma psubstT_eq_rawpsubst4:
assumes "r \<in> atrm" "x1 \<in> var" "x2 \<in> var" "x3 \<in> var" "x4 \<in> var"
"t1 \<in> atrm" "t2 \<in> atrm" "t3 \<in> atrm" "t4 \<in> atrm"
and "x1 \<noteq> x2" "x1 \<noteq> x3" "x2 \<noteq> x3" "x1 \<noteq> x4" "x2 \<noteq> x4" "x3 \<noteq> x4"
"x2 \<notin> FvarsT t1" "x3 \<notin> FvarsT t1" "x3 \<notin> FvarsT t2" "x4 \<notin> FvarsT t1" "x4 \<notin> FvarsT t2" "x4 \<notin> FvarsT t3"
shows "psubstT r [(t1,x1),(t2,x2),(t3,x3),(t4,x4)] = rawpsubstT r [(t1,x1),(t2,x2),(t3,x3),(t4,x4)]"
using assms less_SucE less_Suc_eq_0_disj
by (intro psubstT_eq_rawpsubstT) auto
lemma rawpsubstT_same_Var[simp]:
assumes "r \<in> atrm" "set xs \<subseteq> var"
shows "rawpsubstT r (map (\<lambda>x. (Var x,x)) xs) = r"
using assms by (induct xs) auto
lemma psubstT_same_Var[simp]:
assumes "r \<in> atrm" "set xs \<subseteq> var" and "distinct xs"
shows "psubstT r (map (\<lambda>x. (Var x,x)) xs) = r"
proof-
have "psubstT r (map (\<lambda>x. (Var x,x)) xs) = rawpsubstT r (map (\<lambda>x. (Var x,x)) xs)"
using assms FvarsT_Var[of "xs ! _"] nth_mem[of _ xs]
by (intro psubstT_eq_rawpsubstT)
(auto simp: o_def distinct_conv_nth dest!: FvarsT_VarD)
thus ?thesis using assms by auto
qed
(* The following holds for all trms, so no need to prove it for a trms: *)
thm psubstT_notIn
(***)
(* Behavior of psubst w.r.t. equality formulas: *)
lemma rawpsubst_eql:
assumes "t1 \<in> trm" "t2 \<in> trm"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
shows "rawpsubst (eql t1 t2) txs = eql (rawpsubstT t1 txs) (rawpsubstT t2 txs)"
using assms apply (induct txs arbitrary: t1 t2)
subgoal by auto
subgoal for tx txs t1 t2 by (cases tx) auto .
lemma psubst_eql[simp]:
assumes "t1 \<in> atrm" "t2 \<in> atrm"
and "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "psubst (eql t1 t2) txs = eql (psubstT t1 txs) (psubstT t2 txs)"
proof-
have t12: "fst ` (set txs) \<subseteq> trm" using assms by auto
define us where us: "us \<equiv> getFrN (map snd txs) (map fst txs) [eql t1 t2] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT t1 = {}"
"set us \<inter> FvarsT t2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(1-3) t12 unfolding us
using getFrN_Fvars[of "map snd txs" "map fst txs" "[eql t1 t2]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "map fst txs" "[eql t1 t2]" _ "length txs"]
getFrN_var[of "map snd txs" "map fst txs" "[eql t1 t2]" _ "length txs"]
getFrN_length[of "map snd txs" "map fst txs" "[eql t1 t2]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by force
subgoal by fastforce
subgoal by (fastforce simp: image_iff)
by auto
define vs1 where vs1: "vs1 \<equiv> getFrN (map snd txs) (t1 # map fst txs) [] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> FvarsT t1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms(1-3) t12 unfolding vs1
using getFrN_Fvars[of "map snd txs" "t1 # map fst txs" "[]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "t1 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "t1 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "t1 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by auto
subgoal by force
subgoal by (fastforce simp: image_iff)
by auto
define vs2 where vs2: "vs2 \<equiv> getFrN (map snd txs) (t2 # map fst txs) [] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> FvarsT t2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms(1-3) t12 unfolding vs2
using getFrN_Fvars[of "map snd txs" "t2 # map fst txs" "[]" _ "length txs"]
getFrN_FvarsT[of "map snd txs" "t2 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "t2 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "t2 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by auto
subgoal by force
subgoal by (fastforce simp: image_iff)
by auto
let ?tus = "zip (map fst txs) us"
let ?uxs = "zip (map Var us) (map snd txs)"
let ?e = "rawpsubst (eql t1 t2) ?uxs"
have e: "?e = eql (rawpsubstT t1 ?uxs) (rawpsubstT t2 ?uxs)"
apply(rule rawpsubst_eql) using assms us_facts apply auto
apply(drule set_zip_rightD) apply simp apply blast
apply(drule set_zip_leftD) apply simp apply blast .
have 0: "rawpsubst ?e ?tus =
eql (rawpsubstT (rawpsubstT t1 ?uxs) ?tus) (rawpsubstT (rawpsubstT t2 ?uxs) ?tus)"
unfolding e using assms us_facts apply(intro rawpsubst_eql)
subgoal by (auto intro!: rawpsubstT dest!: set_zip_D)
subgoal by (auto intro!: rawpsubstT dest!: set_zip_D)
subgoal by (auto intro!: rawpsubstT dest!: set_zip_D)
subgoal by (fastforce intro!: rawpsubstT dest!: set_zip_D) .
have 1: "rawpsubstT (rawpsubstT t1 ?uxs) ?tus =
rawpsubstT (rawpsubstT t1 (zip (map Var vs1) (map snd txs))) (zip (map fst txs) vs1)"
using assms us_facts vs1_facts
by (intro rawpsubstT_compose_freshVar2) auto
have 2: "rawpsubstT (rawpsubstT t2 ?uxs) ?tus =
rawpsubstT (rawpsubstT t2 (zip (map Var vs2) (map snd txs))) (zip (map fst txs) vs2)"
using assms us_facts vs2_facts
by (intro rawpsubstT_compose_freshVar2) auto
show ?thesis unfolding psubstT_def psubst_def
by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 0 1 2)
qed
(* psubst versus the exists-unique quantifier: *)
lemma psubst_exu[simp]:
assumes "\<phi> \<in> fmla" "x \<in> var" "snd ` set txs \<subseteq> var" "fst ` set txs \<subseteq> atrm"
"x \<notin> snd ` set txs" "x \<notin> (\<Union>t \<in> fst ` set txs. FvarsT t)" "distinct (map snd txs)"
shows "psubst (exu x \<phi>) txs = exu x (psubst \<phi> txs)"
proof-
have f: "fst ` set txs \<subseteq> trm" using assms by (meson atrm_trm subset_trans)
note assms1 = assms(1-3) assms(5-7) f
define u where u: "u \<equiv> getFr (x # map snd txs) (map fst txs) [\<phi>]"
have u_facts: "u \<in> var" "u \<noteq> x"
"u \<notin> snd ` set txs" "u \<notin> (\<Union>t \<in> fst ` set txs. FvarsT t)" "u \<notin> Fvars \<phi>"
unfolding u using f getFr_FvarsT_Fvars[of "x # map snd txs" "map fst txs" "[\<phi>]"] by (auto simp: assms)
hence [simp]: "psubst (subst \<phi> (Var u) x) txs = subst (psubst \<phi> txs) (Var u) x"
using assms apply(intro psubst_subst_fresh_switch f) by auto
show ?thesis using f assms u_facts
by (subst exu_def_var[of _ u "psubst \<phi> txs"])
(auto dest!: in_Fvars_psubstD split: if_splits simp: exu_def_var[of _ u] )
qed
(* psubst versus the arithmetic trm constructors: *)
(* We already have: *)
thm psubstT_Var_not[no_vars]
lemma rawpsubstT_Var_in:
assumes "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)" and "(s,y) \<in> set txs"
and "\<And> i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> snd (txs!j) \<notin> FvarsT (fst (txs!i))"
shows "rawpsubstT (Var y) txs = s"
using assms proof(induction txs)
case (Cons tx txs)
obtain t x where tx[simp]: "tx = (t,x)" by (cases tx) auto
have 00: "FvarsT t \<inter> snd ` set txs = {}"
using Cons.prems(5)[of 0 "Suc _"] by (auto simp: set_conv_nth)
have "rawpsubstT (substT (Var y) t x) txs = s"
proof(cases "y = x")
- case [simp]:True hence [simp]: "s = t" using `distinct (map snd (tx # txs))`
- `(s, y) \<in> set (tx # txs)` using image_iff by fastforce
+ case [simp]:True hence [simp]: "s = t" using \<open>distinct (map snd (tx # txs))\<close>
+ \<open>(s, y) \<in> set (tx # txs)\<close> using image_iff by fastforce
show ?thesis using Cons.prems 00 by auto
next
case False
hence [simp]: "substT (Var y) t x = Var y"
using Cons.prems by (intro substT_notIn) auto
have "rawpsubstT (Var y) txs = s"
using Cons.prems apply(intro Cons.IH)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal using False by auto
subgoal by (metis length_Cons less_Suc_eq_0_disj nth_Cons_Suc) .
thus ?thesis by simp
qed
thus ?case by simp
qed auto
lemma psubstT_Var_in:
assumes "y \<in> var" "snd ` (set txs) \<subseteq> var" "fst ` (set txs) \<subseteq> trm"
and "distinct (map snd txs)" and "(s,y) \<in> set txs"
shows "psubstT (Var y) txs = s"
proof-
define us where us: "us \<equiv> getFrN (map snd txs) (Var y # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"y \<notin> set us"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms unfolding us
using getFrN_FvarsT[of "map snd txs" "Var y # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "Var y # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "Var y # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by force
subgoal by force
by auto
- obtain i where i[simp]: "i < length txs" "txs!i = (s,y)" using `(s,y) \<in> set txs`
+ obtain i where i[simp]: "i < length txs" "txs!i = (s,y)" using \<open>(s,y) \<in> set txs\<close>
by (metis in_set_conv_nth)
hence 00[simp]: "\<And> j. j < length txs \<Longrightarrow> txs ! j = txs ! i \<Longrightarrow> j = i"
- using `distinct (map snd txs)` distinct_Ex1 nth_mem by fastforce
+ using \<open>distinct (map snd txs)\<close> distinct_Ex1 nth_mem by fastforce
have 000[simp]: "\<And> j ia. j < length txs \<Longrightarrow> ia < length txs \<Longrightarrow> snd (txs ! j) \<noteq> us ! ia"
using assms us_facts
by (metis IntI empty_iff length_map list.set_map nth_map nth_mem)
have [simp]: "\<And>ii jj. ii < jj \<Longrightarrow> jj < length txs \<Longrightarrow> us ! ii \<in> var"
using nth_mem us_facts(1) us_facts(5) by auto
have [simp]: "\<And>i j. i < j \<Longrightarrow> j < length txs \<Longrightarrow> us ! j \<notin> FvarsT (fst (txs ! i))"
using us_facts(2,5) by (auto simp: Int_def)
have 0: "rawpsubstT (Var y) (zip (map Var us) (map snd txs)) = Var (us!i)"
using assms us_facts
by (intro rawpsubstT_Var_in)
(auto dest!: set_zip_D simp: in_set_conv_nth intro!: exI[of _ i])
have "rawpsubstT (rawpsubstT (Var y) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) = s"
unfolding 0 using assms us_facts
by (intro rawpsubstT_Var_in)
(auto dest!: set_zip_D simp: in_set_conv_nth intro!: exI[of _ i])
thus ?thesis unfolding psubstT_def by (simp add: Let_def us[symmetric])
qed
lemma psubstT_Var_Cons_aux:
assumes "y \<in> var" "x \<in> var" "t \<in> atrm"
"snd ` set txs \<subseteq> var" "fst ` set txs \<subseteq> atrm" "x \<notin> snd ` set txs"
"distinct (map snd txs)" "y \<noteq> x"
shows "psubstT (Var y) ((t, x) # txs) = psubstT (Var y) txs"
proof-
have txs_trm: "t \<in> trm" "fst ` set txs \<subseteq> trm" using assms by auto
note assms1 = assms(1,2) assms(4) assms(6-8) txs_trm
note fvt = getFrN_FvarsT[of "x # map snd txs" "Var y # t # map fst txs" "[]" _ "Suc (length txs)"]
and var = getFrN_var[of "x # map snd txs" "Var y # t # map fst txs" "[]" _ "Suc (length txs)"]
and l = getFrN_length[of "x # map snd txs" "Var y # t # map fst txs" "[]" "Suc (length txs)"]
define uus where uus:
"uus \<equiv> getFrN (x # map snd txs) (Var y # t # map fst txs) [] (Suc (length txs))"
have uus_facts: "set uus \<subseteq> var"
"set uus \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set uus \<inter> snd ` (set txs) = {}"
"set uus \<inter> FvarsT t = {}"
"x \<notin> set uus"
"y \<notin> set uus"
"length uus = Suc (length txs)"
"distinct uus"
using assms1 unfolding uus
apply -
subgoal by auto
subgoal using fvt by (simp add: in_fst_image Int_def) (metis prod.collapse)
subgoal using var by (force simp add: in_fst_image Int_def)
subgoal using fvt by auto
subgoal using var by (fastforce simp: in_fst_image Int_def)
subgoal using fvt by (force simp: in_fst_image Int_def)
subgoal using l by auto
subgoal by auto .
obtain u us where uus_us[simp]: "uus = u # us" using uus_facts by (cases uus) auto
have us_facts: "set us \<subseteq> var"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"set us \<inter> FvarsT t = {}"
"x \<notin> set us"
"y \<notin> set us"
"length us = length txs"
"distinct us"
and u_facts: "u \<in> var"
"u \<notin> \<Union> (FvarsT ` (fst ` (set txs)))"
"u \<notin> snd ` (set txs)"
"u \<notin>FvarsT t"
"u \<noteq> x"
"u \<noteq> y"
"u \<notin> set us"
using uus_facts by auto
note fvt = getFrN_FvarsT[of "map snd txs" "Var y # map fst txs" "[]" _ "length txs"]
and var = getFrN_var[of "map snd txs" "Var y # map fst txs" "[]" _ "length txs"]
and l = getFrN_length[of "map snd txs" "Var y # map fst txs" "[]" "length txs"]
define vs where vs: "vs \<equiv> getFrN (map snd txs) (Var y # map fst txs) [] (length txs)"
have vs_facts: "set vs \<subseteq> var"
"set vs \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"y \<notin> set vs"
"set vs \<inter> snd ` (set txs) = {}"
"length vs = length txs"
"distinct vs"
using assms1 unfolding vs
apply -
subgoal by auto
subgoal using fvt by (simp add: in_fst_image Int_def) (metis prod.collapse)
subgoal using fvt l by fastforce
subgoal using var by (force simp: Int_def in_fst_image)
subgoal using l by auto
subgoal by auto .
have 0: "substT (Var y) (Var u) x = Var y"
using assms u_facts by auto
have 1: "substT (rawpsubstT (Var y) (zip (map Var us) (map snd txs))) t u =
rawpsubstT (Var y) (zip (map Var us) (map snd txs))"
using assms u_facts us_facts
by (intro substT_notIn)
(auto 0 3 intro!: rawpsubstT dest!: set_zip_D in_FvarsT_rawpsubstT_imp FvarsT_VarD)
have "rawpsubstT (rawpsubstT (Var y) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT (Var y) (zip (map Var vs) (map snd txs))) (zip (map fst txs) vs)"
using assms vs_facts us_facts by (intro rawpsubstT_compose_freshVar2) auto
thus ?thesis unfolding psubstT_def
by (simp add: Let_def uus[symmetric] vs[symmetric] 0 1)
qed
text \<open>Simplification rules for parallel substitution:\<close>
lemma psubstT_Var_Cons[simp]:
"y \<in> var \<Longrightarrow> x \<in> var \<Longrightarrow> t \<in> atrm \<Longrightarrow>
snd ` set txs \<subseteq> var \<Longrightarrow> fst ` set txs \<subseteq> atrm \<Longrightarrow> distinct (map snd txs) \<Longrightarrow> x \<notin> snd ` set txs \<Longrightarrow>
psubstT (Var y) ((t,x) # txs) = (if y = x then t else psubstT (Var y) txs)"
apply(cases "y = x")
subgoal by (rule psubstT_Var_in) auto
subgoal by (auto intro!: psubstT_Var_Cons_aux) .
lemma psubstT_zer[simp]:
assumes "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "psubstT zer txs = zer"
using assms by (intro psubstT_num) auto
lemma rawpsubstT_suc:
assumes "r \<in> trm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "rawpsubstT (suc r) txs = suc (rawpsubstT r txs)"
using assms apply(induct txs arbitrary: r)
subgoal by simp
subgoal for tx txs r by (cases tx) auto .
lemma psubstT_suc[simp]:
assumes "r \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "psubstT (suc r) txs = suc (psubstT r txs)"
proof-
have 000: "r \<in> trm" "fst ` (set txs) \<subseteq> trm" using assms by auto
define us where us: "us \<equiv> getFrN (map snd txs) (suc r # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(2) 000 unfolding us
using getFrN_FvarsT[of "map snd txs" "suc r # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "suc r # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "suc r # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "suc r # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "suc r # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by auto
subgoal by force
by auto
define vs where vs: "vs \<equiv> getFrN (map snd txs) (r # map fst txs) [] (length txs)"
have vs_facts: "set vs \<subseteq> var"
"set vs \<inter> FvarsT r = {}"
"set vs \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs \<inter> snd ` (set txs) = {}"
"length vs = length txs"
"distinct vs"
using assms(2) 000 unfolding vs
using getFrN_FvarsT[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
by auto
have 0: "rawpsubstT (suc r) (zip (map Var vs) (map snd txs)) =
suc (rawpsubstT r (zip (map Var vs) (map snd txs)))"
using assms vs_facts by (intro rawpsubstT_suc) (auto dest!: set_zip_D)
have "rawpsubstT (rawpsubstT (suc r) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT (suc r) (zip (map Var vs) (map snd txs))) (zip (map fst txs) vs)"
using assms us_facts vs_facts by (intro rawpsubstT_compose_freshVar2) auto
also have "\<dots> = suc (rawpsubstT (rawpsubstT r (zip (map Var vs) (map snd txs))) (zip (map fst txs) vs))"
unfolding 0 using assms vs_facts apply(intro rawpsubstT_suc)
subgoal by (auto dest!: set_zip_D intro!: rawpsubstT)
subgoal by (auto dest!: set_zip_D)
subgoal by (fastforce dest!: set_zip_D simp: Int_def) .
finally show ?thesis
by (simp add: Let_def us[symmetric] vs[symmetric] psubstT_def)
qed
lemma rawpsubstT_pls:
assumes "r1 \<in> trm" "r2 \<in> trm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "rawpsubstT (pls r1 r2) txs = pls (rawpsubstT r1 txs) (rawpsubstT r2 txs)"
using assms apply(induct txs arbitrary: r1 r2)
subgoal by simp
subgoal for tx txs r by (cases tx) auto .
lemma psubstT_pls[simp]:
assumes "r1 \<in> atrm" "r2 \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "psubstT (pls r1 r2) txs = pls (psubstT r1 txs) (psubstT r2 txs)"
proof-
have 000: "fst ` (set txs) \<subseteq> trm" using assms by auto
define us where us: "us \<equiv> getFrN (map snd txs) (pls r1 r2 # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r1 = {}"
"set us \<inter> FvarsT r2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(1-3) 000 unfolding us
using getFrN_FvarsT[of "map snd txs" "pls r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "pls r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "pls r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "pls r1 r2 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "pls r1 r2 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by force
subgoal by auto
subgoal by force
by auto
define vs1 where vs1: "vs1 \<equiv> getFrN (map snd txs) (r1 # map fst txs) [] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> FvarsT r1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms(1-3) 000 unfolding vs1
using getFrN_FvarsT[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r1 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r1 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
by auto
define vs2 where vs2: "vs2 \<equiv> getFrN (map snd txs) (r2 # map fst txs) [] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> FvarsT r2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms(1-3) 000 unfolding vs2
using getFrN_FvarsT[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r2 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r2 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by force
by auto
have 0: "rawpsubstT (pls r1 r2) (zip (map Var us) (map snd txs)) =
pls (rawpsubstT r1 (zip (map Var us) (map snd txs)))
(rawpsubstT r2 (zip (map Var us) (map snd txs)))"
using assms us_facts by (intro rawpsubstT_pls) (auto dest!: set_zip_D)
have 1: "rawpsubstT (rawpsubstT r1 (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT r1 (zip (map Var vs1) (map snd txs))) (zip (map fst txs) vs1)"
using assms us_facts vs1_facts by (intro rawpsubstT_compose_freshVar2) auto
have 2: "rawpsubstT (rawpsubstT r2 (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT r2 (zip (map Var vs2) (map snd txs))) (zip (map fst txs) vs2)"
using assms us_facts vs2_facts by (intro rawpsubstT_compose_freshVar2) auto
have 3: "rawpsubstT (rawpsubstT (pls r1 r2) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
pls (rawpsubstT (rawpsubstT r1 (zip (map Var us) (map snd txs))) (zip (map fst txs) us))
(rawpsubstT (rawpsubstT r2 (zip (map Var us) (map snd txs))) (zip (map fst txs) us))"
unfolding 0 using assms us_facts apply(intro rawpsubstT_pls)
subgoal by (auto dest!: set_zip_D intro!: rawpsubstT)
subgoal by (force dest!: set_zip_D intro!: rawpsubstT simp: Int_def)
subgoal by (auto dest!: set_zip_D intro!: rawpsubstT)
subgoal by (fastforce dest!: set_zip_D intro!: rawpsubstT simp: Int_def) .
show ?thesis unfolding psubstT_def
by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 1 2 3)
qed
lemma rawpsubstT_tms:
assumes "r1 \<in> trm" "r2 \<in> trm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> trm"
shows "rawpsubstT (tms r1 r2) txs = tms (rawpsubstT r1 txs) (rawpsubstT r2 txs)"
using assms apply(induct txs arbitrary: r1 r2)
subgoal by simp
subgoal for tx txs r by (cases tx) auto .
lemma psubstT_tms[simp]:
assumes "r1 \<in> atrm" "r2 \<in> atrm" and "snd ` (set txs) \<subseteq> var" and "fst ` (set txs) \<subseteq> atrm"
and "distinct (map snd txs)"
shows "psubstT (tms r1 r2) txs = tms (psubstT r1 txs) (psubstT r2 txs)"
proof-
have 000: "fst ` (set txs) \<subseteq> trm" using assms by auto
define us where us: "us \<equiv> getFrN (map snd txs) (tms r1 r2 # map fst txs) [] (length txs)"
have us_facts: "set us \<subseteq> var"
"set us \<inter> FvarsT r1 = {}"
"set us \<inter> FvarsT r2 = {}"
"set us \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set us \<inter> snd ` (set txs) = {}"
"length us = length txs"
"distinct us"
using assms(1-3) 000 unfolding us
using getFrN_FvarsT[of "map snd txs" "tms r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "tms r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "tms r1 r2 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "tms r1 r2 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "tms r1 r2 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by force
subgoal by auto
subgoal by force
by auto
define vs1 where vs1: "vs1 \<equiv> getFrN (map snd txs) (r1 # map fst txs) [] (length txs)"
have vs1_facts: "set vs1 \<subseteq> var"
"set vs1 \<inter> FvarsT r1 = {}"
"set vs1 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs1 \<inter> snd ` (set txs) = {}"
"length vs1 = length txs"
"distinct vs1"
using assms(1-3) 000 unfolding vs1
using getFrN_FvarsT[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r1 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r1 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r1 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by auto
subgoal by force
subgoal by force
by auto
define vs2 where vs2: "vs2 \<equiv> getFrN (map snd txs) (r2 # map fst txs) [] (length txs)"
have vs2_facts: "set vs2 \<subseteq> var"
"set vs2 \<inter> FvarsT r2 = {}"
"set vs2 \<inter> \<Union> (FvarsT ` (fst ` (set txs))) = {}"
"set vs2 \<inter> snd ` (set txs) = {}"
"length vs2 = length txs"
"distinct vs2"
using assms(1-3) 000 unfolding vs2
using getFrN_FvarsT[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_Fvars[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_var[of "map snd txs" "r2 # map fst txs" "[]" _ "length txs"]
getFrN_length[of "map snd txs" "r2 # map fst txs" "[]" "length txs"]
getFrN_length[of "map snd txs" "r2 # map fst txs" "[]" "length txs"]
apply -
subgoal by auto
subgoal by force
subgoal by auto
subgoal by force
subgoal by force
by auto
have 0: "rawpsubstT (tms r1 r2) (zip (map Var us) (map snd txs)) =
tms (rawpsubstT r1 (zip (map Var us) (map snd txs)))
(rawpsubstT r2 (zip (map Var us) (map snd txs)))"
using assms us_facts by (intro rawpsubstT_tms) (auto dest!: set_zip_D)
have 1: "rawpsubstT (rawpsubstT r1 (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT r1 (zip (map Var vs1) (map snd txs))) (zip (map fst txs) vs1)"
using assms us_facts vs1_facts by (intro rawpsubstT_compose_freshVar2) auto
have 2: "rawpsubstT (rawpsubstT r2 (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
rawpsubstT (rawpsubstT r2 (zip (map Var vs2) (map snd txs))) (zip (map fst txs) vs2)"
using assms us_facts vs2_facts by (intro rawpsubstT_compose_freshVar2) auto
have 3: "rawpsubstT (rawpsubstT (tms r1 r2) (zip (map Var us) (map snd txs))) (zip (map fst txs) us) =
tms (rawpsubstT (rawpsubstT r1 (zip (map Var us) (map snd txs))) (zip (map fst txs) us))
(rawpsubstT (rawpsubstT r2 (zip (map Var us) (map snd txs))) (zip (map fst txs) us))"
unfolding 0 using assms us_facts apply(intro rawpsubstT_tms)
subgoal by (auto dest!: set_zip_D intro!: rawpsubstT)
subgoal by (force dest!: set_zip_D intro!: rawpsubstT simp: Int_def)
subgoal by (auto dest!: set_zip_D intro!: rawpsubstT)
subgoal by (fastforce dest!: set_zip_D intro!: rawpsubstT simp: Int_def) .
show ?thesis unfolding psubstT_def
by (simp add: Let_def us[symmetric] vs1[symmetric] vs2[symmetric] 1 2 3)
qed
section \<open>The (Nonstrict and Strict) Order Relations\<close>
text \<open>Lq (less than or equal to) is a formula with free vars xx and yy.
NB: Out of the two possible ways, adding zz to the left or to the right,
we choose the former, since this seems to enable Q (Robinson arithmetic)
to prove as many useful properties as possible.\<close>
definition Lq :: "'fmla" where
"Lq \<equiv> exi zz (eql (Var yy) (pls (Var zz) (Var xx)))"
text \<open>Alternative, more flexible definition , for any non-capturing bound variable:\<close>
lemma Lq_def2: "z \<in> var \<Longrightarrow> z \<noteq> yy \<Longrightarrow> z \<noteq> xx \<Longrightarrow> Lq = exi z (eql (Var yy) (pls (Var z) (Var xx)))"
unfolding Lq_def using exi_rename[of "eql (Var yy) (pls (Var zz) (Var xx))" zz z] by auto
lemma Lq[simp,intro!]: "Lq \<in> fmla"
unfolding Lq_def by auto
lemma Fvars_Lq[simp]: "Fvars Lq = {xx,yy}"
unfolding Lq_def by auto
text \<open>As usual, we also define a predicate version:\<close>
definition LLq where "LLq \<equiv> \<lambda> t1 t2. psubst Lq [(t1,xx), (t2,yy)]"
lemma LLq[simp,intro]:
assumes "t1 \<in> trm" "t2 \<in> trm"
shows "LLq t1 t2 \<in> fmla"
using assms unfolding LLq_def by auto
lemma LLq2[simp,intro!]:
"n \<in> num \<Longrightarrow> LLq n (Var yy') \<in> fmla"
by auto
lemma Fvars_LLq[simp]: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow>
Fvars (LLq t1 t2) = FvarsT t1 \<union> FvarsT t2"
unfolding LLq_def apply(subst Fvars_psubst)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal apply safe
subgoal by auto
subgoal by auto
subgoal by force
subgoal by force
subgoal by force
subgoal by force . .
text \<open>This lemma will be the working definition of LLq:\<close>
lemma LLq_pls:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm" "z \<in> var" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
shows "LLq t1 t2 = exi z (eql t2 (pls (Var z) t1))"
proof-
define z' where "z' \<equiv> getFr [xx,yy,z] [t1,t2] []"
have z_facts[simp]: "z' \<in> var" "z' \<noteq> yy" "z' \<noteq> xx" "z' \<noteq> z" "z \<noteq> z'" "z' \<notin> FvarsT t1" "z' \<notin> FvarsT t2"
using getFr_FvarsT_Fvars[of "[xx,yy,z]" "[t1,t2]" "[]"] unfolding z'_def by auto
have "LLq t1 t2 = exi z' (eql t2 (pls (Var z') t1))"
by (simp add: LLq_def Lq_def2[of z'])
also have "\<dots> = exi z (eql t2 (pls (Var z) t1))"
using exi_rename[of "eql t2 (pls (Var z') t1)" z' z, simplified] .
finally show ?thesis .
qed
lemma LLq_pls_zz:
assumes "t1 \<in> atrm" "t2 \<in> atrm" "zz \<notin> FvarsT t1" "zz \<notin> FvarsT t2"
shows "LLq t1 t2 = exi zz (eql t2 (pls (Var zz) t1))"
using assms by (intro LLq_pls) auto
text \<open>If we restrict attention to arithmetic terms, we can prove a uniform
substitution property for LLq:\<close>
lemma subst_LLq[simp]:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm" "s \<in> atrm" "x \<in> var"
shows "subst (LLq t1 t2) s x = LLq (substT t1 s x) (substT t2 s x)"
proof-
define z where "z \<equiv> getFr [xx,yy,x] [t1,t2,s] []"
have z_facts[simp]: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy" "z \<noteq> x" "z \<notin> FvarsT t1" "z \<notin> FvarsT t2" "z \<notin> FvarsT s"
using getFr_FvarsT_Fvars[of "[xx,yy,x]" "[t1,t2,s]" "[]"] unfolding z_def by auto
show ?thesis
by(simp add: FvarsT_substT LLq_pls[of _ _ z] subst2_fresh_switch Lq_def)
qed
lemma psubst_LLq[simp]:
assumes 1: "t1 \<in> atrm" "t2 \<in> atrm" "fst ` set txs \<subseteq> atrm"
and 2: "snd ` set txs \<subseteq> var"
and 3: "distinct (map snd txs)"
shows "psubst (LLq t1 t2) txs = LLq (psubstT t1 txs) (psubstT t2 txs)"
proof-
have 0: "t1 \<in> trm" "t2 \<in> trm" "fst ` set txs \<subseteq> trm" using 1 by auto
define z where z: "z \<equiv> getFr ([xx,yy] @ map snd txs) ([t1,t2] @ map fst txs) []"
have us_facts: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy"
"z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
"z \<notin> \<Union> (FvarsT ` (fst ` (set txs)))"
"z \<notin> snd ` (set txs)"
using 0 2 unfolding z
using getFr_FvarsT[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]" ]
getFr_Fvars[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]" ]
getFr_var[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]"]
apply -
subgoal by auto
subgoal by force
subgoal by force
subgoal by force
subgoal by force
subgoal by auto
subgoal by (force simp: image_iff) .
note in_FvarsT_psubstTD[dest!]
note if_splits[split]
show ?thesis
using assms 0 us_facts apply(subst LLq_pls[of _ _ z])
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal apply(subst LLq_pls[of _ _ z]) by auto .
qed
text \<open>Lq less than) is the strict version of the order relation.
We prove similar facts as for Lq\<close>
definition Ls :: "'fmla" where
"Ls \<equiv> cnj Lq (neg (eql (Var xx) (Var yy)))"
lemma Ls[simp,intro!]: "Ls \<in> fmla"
unfolding Ls_def by auto
lemma Fvars_Ls[simp]: "Fvars Ls = {xx,yy}"
unfolding Ls_def by auto
definition LLs where "LLs \<equiv> \<lambda> t1 t2. psubst Ls [(t1,xx), (t2,yy)]"
lemma LLs[simp,intro]:
assumes "t1 \<in> trm" "t2 \<in> trm"
shows "LLs t1 t2 \<in> fmla"
using assms unfolding LLs_def by auto
lemma LLs2[simp,intro!]:
"n \<in> num \<Longrightarrow> LLs n (Var yy') \<in> fmla"
by auto
lemma Fvars_LLs[simp]: "t1 \<in> trm \<Longrightarrow> t2 \<in> trm \<Longrightarrow>
Fvars (LLs t1 t2) = FvarsT t1 \<union> FvarsT t2"
unfolding LLs_def apply(subst Fvars_psubst)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal apply safe
subgoal by auto
subgoal by auto
subgoal by force
subgoal by force
subgoal by force
subgoal by force . .
text \<open>The working definition of LLs:\<close>
lemma LLs_LLq:
"t1 \<in> atrm \<Longrightarrow> t2 \<in> atrm \<Longrightarrow>
LLs t1 t2 = cnj (LLq t1 t2) (neg (eql t1 t2))"
by (simp add: LLs_def Ls_def LLq_def)
lemma subst_LLs[simp]:
assumes [simp]: "t1 \<in> atrm" "t2 \<in> atrm" "s \<in> atrm" "x \<in> var"
shows "subst (LLs t1 t2) s x = LLs (substT t1 s x) (substT t2 s x)"
by(simp add: LLs_LLq subst2_fresh_switch Ls_def)
lemma psubst_LLs[simp]:
assumes 1: "t1 \<in> atrm" "t2 \<in> atrm" "fst ` set txs \<subseteq> atrm"
and 2: "snd ` set txs \<subseteq> var"
and 3: "distinct (map snd txs)"
shows "psubst (LLs t1 t2) txs = LLs (psubstT t1 txs) (psubstT t2 txs)"
proof-
have 0: "t1 \<in> trm" "t2 \<in> trm" "fst ` set txs \<subseteq> trm" using 1 by auto
define z where z: "z \<equiv> getFr ([xx,yy] @ map snd txs) ([t1,t2] @ map fst txs) []"
have us_facts: "z \<in> var" "z \<noteq> xx" "z \<noteq> yy"
"z \<notin> FvarsT t1" "z \<notin> FvarsT t2"
"z \<notin> \<Union> (FvarsT ` (fst ` (set txs)))"
"z \<notin> snd ` (set txs)"
using 0 2 unfolding z
using getFr_FvarsT[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]" ]
getFr_Fvars[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]" ]
getFr_var[of "[xx,yy] @ map snd txs" "[t1,t2] @ map fst txs" "[]"]
apply -
subgoal by auto
subgoal by force
subgoal by force
subgoal by force
subgoal by force
subgoal by auto
subgoal by (force simp: image_iff) .
show ?thesis
using assms 0 us_facts by (simp add: LLs_LLq)
qed
section \<open>Bounded Quantification\<close>
text \<open>Bounded forall\<close>
definition ball :: "'var \<Rightarrow> 'trm \<Rightarrow> 'fmla \<Rightarrow> 'fmla" where
"ball x t \<phi> \<equiv> all x (imp (LLq (Var x) t) \<phi>)"
lemma ball[simp, intro]: "x \<in> var \<Longrightarrow> t \<in> trm \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> ball x t \<phi> \<in> fmla"
unfolding ball_def by auto
lemma Fvars_ball[simp]:
"x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> Fvars (ball x t \<phi>) = (Fvars \<phi> \<union> FvarsT t) - {x}"
unfolding ball_def by auto
lemma subst_ball:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> x \<noteq> y \<Longrightarrow> x \<notin> FvarsT t1 \<Longrightarrow>
subst (ball x t \<phi>) t1 y = ball x (substT t t1 y) (subst \<phi> t1 y)"
unfolding ball_def by simp
lemma psubst_ball:
"\<phi> \<in> fmla \<Longrightarrow> y \<in> var \<Longrightarrow> snd ` set txs \<subseteq> var \<Longrightarrow> t \<in> atrm \<Longrightarrow>
fst ` set txs \<subseteq> trm \<Longrightarrow> fst ` set txs \<subseteq> atrm \<Longrightarrow> y \<notin> snd ` set txs \<Longrightarrow> y \<notin> (\<Union>t \<in> fst ` set txs. FvarsT t) \<Longrightarrow>
distinct (map snd txs) \<Longrightarrow>
psubst (ball y t \<phi>) txs = ball y (psubstT t txs) (psubst \<phi> txs)"
unfolding ball_def by simp
text \<open>Bounded exists\<close>
definition bexi :: "'var \<Rightarrow> 'trm \<Rightarrow> 'fmla \<Rightarrow> 'fmla" where
"bexi x t \<phi> \<equiv> exi x (cnj (LLq (Var x) t) \<phi>)"
lemma bexi[simp, intro]: "x \<in> var \<Longrightarrow> t \<in> trm \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> bexi x t \<phi> \<in> fmla"
unfolding bexi_def by auto
lemma Fvars_bexi[simp]:
"x \<in> var \<Longrightarrow> \<phi> \<in> fmla \<Longrightarrow> t \<in> trm \<Longrightarrow> Fvars (bexi x t \<phi>) = (Fvars \<phi> \<union> FvarsT t) - {x}"
unfolding bexi_def by auto
lemma subst_bexi:
"\<phi> \<in> fmla \<Longrightarrow> t \<in> atrm \<Longrightarrow> t1 \<in> atrm \<Longrightarrow> x \<in> var \<Longrightarrow> y \<in> var \<Longrightarrow> x \<noteq> y \<Longrightarrow> x \<notin> FvarsT t1 \<Longrightarrow>
subst (bexi x t \<phi>) t1 y = bexi x (substT t t1 y) (subst \<phi> t1 y)"
unfolding bexi_def by simp
lemma psubst_bexi:
"\<phi> \<in> fmla \<Longrightarrow> y \<in> var \<Longrightarrow> snd ` set txs \<subseteq> var \<Longrightarrow> t \<in> atrm \<Longrightarrow>
fst ` set txs \<subseteq> trm \<Longrightarrow> fst ` set txs \<subseteq> atrm \<Longrightarrow> y \<notin> snd ` set txs \<Longrightarrow> y \<notin> (\<Union>t \<in> fst ` set txs. FvarsT t) \<Longrightarrow>
distinct (map snd txs) \<Longrightarrow>
psubst (bexi y t \<phi>) txs = bexi y (psubstT t txs) (psubst \<phi> txs)"
unfolding bexi_def by simp
end \<comment> \<open>context @{locale Syntax_Arith}\<close>
(*<*)
end
(*>*)
diff --git a/thys/Szpilrajn/Szpilrajn.thy b/thys/Szpilrajn/Szpilrajn.thy
--- a/thys/Szpilrajn/Szpilrajn.thy
+++ b/thys/Szpilrajn/Szpilrajn.thy
@@ -1,270 +1,270 @@
(*<*)
theory Szpilrajn
imports Main
begin
(*>*)
section "Introduction"
text \<open>
We formalize the Szpilrajn extension theorem~\cite{Szpilrajn:1930}, also known
as order-extension principal:
Every strict partial order can be extended to strict linear order.
This is a formalization of the proof presented in the Wikipedia article~\cite{wiki}.
A strict partial order is a transitive and irreflexive relation:\<close>
definition "strict_partial_order r \<equiv> trans r \<and> irrefl r"
lemma show_strict_partial_order[intro]:
assumes "trans r" and "irrefl r"
shows "strict_partial_order r"
by (simp add: assms strict_partial_order_def)
lemma strict_partial_order_acyclic:
assumes "strict_partial_order r"
shows "acyclic r"
by (metis acyclic_irrefl assms strict_partial_order_def trancl_id)
text \<open>A typical example is \<^term>\<open>(\<subset>)\<close> on sets:\<close>
lemma strict_partial_order_subset:
"strict_partial_order {(x,y). x \<subset> y}"
proof
show "trans {(x,y). x \<subset> y}"
by (auto simp add: trans_def)
show "irrefl {(x, y). x \<subset> y}"
by (simp add: irrefl_def)
qed
text \<open>A strict linear order has all the properties of a strict partial order, but is also total: \<close>
lemma strict_linear_order_def:
"strict_linear_order r \<longleftrightarrow> strict_partial_order r \<and> total r"
by (simp add: strict_linear_order_on_def strict_partial_order_def)
section "The Proof"
text \<open>A relation \<^term>\<open>r\<close> is a strict extension of a base relation \<^term>\<open>base_r\<close>
if \<^term>\<open>r\<close> is a strict partial order and \<^term>\<open>r\<close> includes \<^term>\<open>base_r\<close>:\<close>
definition "strict_ext base_r r \<equiv> strict_partial_order r \<and> base_r \<subseteq> r"
text \<open>We start by proving that a strict partial order with two incomparable elements
\<^term>\<open>x\<close> and \<^term>\<open>y\<close> can be extended to a strict partial order where \<^term>\<open>x < y\<close>. \<close>
lemma can_extend_partial_order:
assumes spo: "strict_partial_order r"
and no1: "(x,y) \<notin> r"
and no2: "(y,x) \<notin> r"
and neq: "x\<noteq>y"
shows "strict_ext r ((r \<union> {(x,y)})\<^sup>+)"
unfolding strict_ext_def proof (intro conjI show_strict_partial_order)
show "trans ((r \<union> {(x, y)})\<^sup>+)"
by simp
show "r \<subseteq> (r \<union> {(x, y)})\<^sup>+" by auto
from spo have "trans r" and "irrefl r"
by (auto simp add: strict_partial_order_def)
show "irrefl ((r \<union> {(x, y)})\<^sup>+)"
proof (clarsimp simp add: acyclic_irrefl[symmetric], intro conjI)
show "acyclic r"
by (simp add: spo strict_partial_order_acyclic)
show "(y, x) \<notin> r\<^sup>*"
using \<open>trans r\<close> neq no2 rtranclD by fastforce
qed
qed
text \<open>With this, we can start the proof of the Szpilrajn extension theorem.
For this we will use a variant of Zorns Lemma, which only considers nonempty chains:\<close>
lemma Zorns_po_lemma_nonempty:
assumes po: "Partial_order r"
and u: "\<And>C. \<lbrakk>C \<in> Chains r; C\<noteq>{}\<rbrakk> \<Longrightarrow> \<exists>u\<in>Field r. \<forall>a\<in>C. (a, u) \<in> r"
and ne: "r \<noteq> {}"
shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
proof -
- from `r\<noteq>{}` obtain x where "x\<in>Field r"
+ from \<open>r\<noteq>{}\<close> obtain x where "x\<in>Field r"
using FieldI2 by fastforce
with assms show ?thesis
using Zorns_po_lemma by (metis empty_iff)
qed
theorem Szpilrajn:
assumes "strict_partial_order base_r"
shows "\<exists>r. strict_linear_order r \<and> base_r \<subseteq> r"
proof -
text \<open>We define an order on the set of strict extensions of the base relation \<^term>\<open>base_r\<close>,
where \<^term>\<open>r \<le> s\<close> iff \<^term>\<open>r \<subseteq> s\<close>:\<close>
define order_of_orders :: "('a rel) rel" where order_of_orders_def:
"order_of_orders = {(r,s). r\<subseteq>s \<and> strict_ext base_r r \<and> strict_ext base_r s }"
have ord_Field: "Field order_of_orders = {r. strict_ext base_r r}"
by (auto simp add: Field_def order_of_orders_def)
text \<open>We now show that this set has a maximum and that any maximum of this set is
a strict linear order and as thus is one of the extensions we are looking for.\<close>
text \<open>We begin by showing the existence of a maximal element \<^term>\<open>m\<close> using Zorns Lemma:\<close>
have "\<exists>m\<in>Field order_of_orders.
\<forall>a\<in>Field order_of_orders. (m, a) \<in> order_of_orders \<longrightarrow> a = m"
proof (rule Zorns_po_lemma_nonempty)
text \<open>Zorns Lemma requires us to prove that our \<^term>\<open>order_of_orders\<close> is a nonempty partial order
and that every nonempty chain has an upper bound.
The partial order property is trivial, since we used \<^term>\<open>(\<subseteq>)\<close> for the relation:\<close>
show "Partial_order order_of_orders"
by (auto simp add: order_of_orders_def order_on_defs refl_on_def Field_def trans_def antisym_def)
text \<open>Also, our order is obviously not empty since it contains \<^term>\<open>(base_r, base_r)\<close>:\<close>
have "(base_r, base_r)\<in>order_of_orders"
using assms strict_ext_def by (auto simp add: order_of_orders_def)
thus "order_of_orders \<noteq> {}" by force
text \<open>Next we show that each chain has an upper bound.
For the upper bound we take the union of all relations in the chain. \<close>
show "\<exists>u\<in>Field order_of_orders. \<forall>a\<in>C. (a, u) \<in> order_of_orders"
if C_def: "C \<in> Chains order_of_orders" and C_nonemtpy: "C \<noteq> {}"
for C
proof (rule bexI[where x="\<Union>C"])
text \<open>Obviously each element in the chain is a strict extension of \<^term>\<open>base_r\<close> by definition
and as such it is transitive, irreflexive and extends the base relation.\<close>
have r_se: "strict_ext base_r r" if "r \<in> C" for r
- using `r \<in> C` C_def by (auto simp add: Chains_def order_of_orders_def)
+ using \<open>r \<in> C\<close> C_def by (auto simp add: Chains_def order_of_orders_def)
hence r_trans: "trans r"
and r_irrefl: "irrefl r"
and r_extends_base: "base_r \<subseteq> r"
if "r \<in> C" for r
using that by (auto simp add: strict_ext_def strict_partial_order_def)
text \<open>Because a chain is ordered, the union of the chain is also transitive:\<close>
have C_ordered: "r\<subseteq>s \<or> s\<subseteq>r" if "r \<in> C" and "s \<in> C" for r s
using C_def that by (auto simp add: Chains_def order_of_orders_def)
hence "trans (\<Union>C)"
by (simp add: chain_subset_def chain_subset_trans_Union r_trans)
text \<open>The other properties also can be transferred from the single relations
to the union of the chain.
Therefore the union is also a strict extension of \<^term>\<open>base_r\<close>: \<close>
moreover have "irrefl (\<Union>C)"
using irrefl_def r_irrefl by auto
moreover have "base_r \<subseteq> (\<Union>C)"
by (simp add: less_eq_Sup r_extends_base that)
ultimately have "strict_ext base_r (\<Union>C)"
by (simp add: show_strict_partial_order strict_ext_def that)
show "(\<Union>C) \<in> Field order_of_orders"
by (simp add: \<open>strict_ext base_r (\<Union> C)\<close> ord_Field)
text \<open>The union is obviously an upper bound for the chain: \<close>
show "\<forall>a\<in>C. (a, \<Union> C) \<in> order_of_orders"
by (simp add: Sup_upper \<open>strict_ext base_r (\<Union> C)\<close> order_of_orders_def r_se)
qed
qed
text \<open>Let our maximal element be named \<^term>\<open>max\<close>:\<close>
from this obtain max
where max_field: "max\<in>Field order_of_orders"
and is_max:
"\<forall>a\<in>Field order_of_orders. (max, a) \<in> order_of_orders \<longrightarrow> a = max"
by auto
from max_field have max_se: "strict_ext base_r max"
using ord_Field by auto
hence max_spo: "strict_partial_order max"
and "base_r \<subseteq> max"
using strict_ext_def by auto
text \<open>We still have to show, that \<^term>\<open>max\<close> is a strict linear order,
meaning that it is also a total order: \<close>
have "total max"
proof
fix x y :: 'a
assume "x\<noteq>y"
show "(x, y) \<in> max \<or> (y, x) \<in> max"
proof (rule ccontr, auto)
text \<open>Assume that \<^term>\<open>max\<close> is not total, and \<^term>\<open>x\<close> and \<^term>\<open>y\<close> are incomparable.
Then we can extend \<^term>\<open>max\<close> by setting $x < y$:\<close>
assume "(x, y) \<notin> max" and "(y, x) \<notin> max"
let ?max' = "((max \<union> {(x, y)})\<^sup>+)"
- from max_spo `(x, y) \<notin> max` `(y, x) \<notin> max` `x\<noteq>y`
+ from max_spo \<open>(x, y) \<notin> max\<close> \<open>(y, x) \<notin> max\<close> \<open>x\<noteq>y\<close>
have max'_se_max: "strict_ext max ?max'" by (rule can_extend_partial_order)
hence max'_se: "strict_ext base_r ?max'"
by (meson \<open>base_r \<subseteq> max\<close> strict_ext_def subset_trans)
text \<open>The extended relation is greater than \<^term>\<open>max\<close>, which is a contradiction.\<close>
have "(max, ?max')\<in>order_of_orders"
using max'_se by (auto simp add: order_of_orders_def max_se)
thus False
using FieldI2 \<open>(x, y) \<notin> max\<close> is_max by fastforce
qed
qed
with max_spo have "strict_linear_order max"
by (auto simp add: strict_linear_order_def)
with \<open>base_r \<subseteq> max\<close>
show "\<exists>r. strict_linear_order r \<and> base_r \<subseteq> r" by auto
qed
text \<open>As a corollary, we can also show that we can extend any \<^term>\<open>acyclic\<close> relation
to a strict linear order: \<close>
corollary can_extend_acyclic_order_to_strict_linear:
assumes "acyclic base_r"
shows "\<exists>r. strict_linear_order r \<and> base_r \<subseteq> r"
proof -
have "strict_partial_order (base_r\<^sup>+)"
using acyclic_irrefl assms trans_trancl by blast
thus ?thesis
by (meson Szpilrajn r_into_trancl' subset_iff)
qed
text \<open>Let us conclude with an example, showing that there exists a strict linear
order on sets, which includes the subset relation:\<close>
lemma exists_strict_partial_order_on_sets:
shows "\<exists>r. strict_linear_order r \<and> {(x,y). x \<subset> y} \<subseteq> r"
using strict_partial_order_subset by (rule Szpilrajn)
(*<*)
end
(*>*)

File Metadata

Mime Type
application/octet-stream
Expires
Sun, Apr 21, 8:13 PM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
MOfC2Mm98FhT
Default Alt Text
(5 MB)

Event Timeline